# HG changeset patch # User zhang # Date 1348922292 0 # Node ID 1ce04eb1c8adf9f67135bfe74f792a298195e8f7 # Parent cbb4ac6c80815ff84c9ad91e0badd329ed3d08c1 Initial upload of the formal construction of Universal Turing Machine. diff -r cbb4ac6c8081 -r 1ce04eb1c8ad utm/IsaMakefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/utm/IsaMakefile Sat Sep 29 12:38:12 2012 +0000 @@ -0,0 +1,31 @@ + +## targets + +default: utm +images: utm +test: + +all: images test + + +## global settings + +SRC = $(ISABELLE_HOME)/src +OUT = $(ISABELLE_OUTPUT) +LOG = $(OUT)/log + +USEDIR = $(ISABELLE_TOOL) usedir -v true -i true -d pdf ## -D generated + + +## utm + +utm: $(OUT)/utm + +$(OUT)/utm: ## ROOT.ML document/root.tex *.thy + @$(USEDIR) -b HOL utm + + +## clean + +clean: + @rm -f $(OUT)/utm diff -r cbb4ac6c8081 -r 1ce04eb1c8ad utm/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/utm/ROOT.ML Sat Sep 29 12:38:12 2012 +0000 @@ -0,0 +1,12 @@ +(* + turing_basic.thy : The basic definitions of Turing Machine. + uncomputable.thy : The existence of Turing uncomputable functions. + abacus.thy : The basic definitions of Abacus machine (An intermediate langauge underneath recursive functions) and + the compilation of Abacus machines into Turing Machines. + recursive.thy : The basic defintions of Recursive Functions and the compilation of Recursive Functions into + Abacus machines. + UF.thy : The construction of Universal Function, named "rec_F" and the proof of its correctness. + UTM.thy: Obtaining Uinversal Turing Machine by scarfolding the Turing Machine compiled from "rec_F" with some + initialization and termination processing Turing Machines. +*) + no_document use_thys ["turing_basic", "uncomputable", "abacus", "recursive", "UF", "UTM"] diff -r cbb4ac6c8081 -r 1ce04eb1c8ad utm/UF.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/utm/UF.thy Sat Sep 29 12:38:12 2012 +0000 @@ -0,0 +1,4914 @@ +theory UF +imports Main rec_def turing_basic GCD abacus +begin + +text {* + This theory file constructs the Universal Function @{text "rec_F"}, which is the UTM defined + in terms of recursive functions. This @{text "rec_F"} is essentially an + interpreter of Turing Machines. Once the correctness of @{text "rec_F"} is established, + UTM can easil be obtained by compling @{text "rec_F"} into the corresponding Turing Machine. +*} + + +section {* The construction of component functions *} + +text {* + This section constructs a set of component functions used to construct @{text "rec_F"}. + *} + +text {* + The recursive function used to do arithmatic addition. +*} +definition rec_add :: "recf" + where + "rec_add \ Pr 1 (id 1 0) (Cn 3 s [id 3 2])" + +text {* + The recursive function used to do arithmatic multiplication. +*} +definition rec_mult :: "recf" + where + "rec_mult = Pr 1 z (Cn 3 rec_add [id 3 0, id 3 2])" + +text {* + The recursive function used to do arithmatic precede. +*} +definition rec_pred :: "recf" + where + "rec_pred = Cn 1 (Pr 1 z (id 3 1)) [id 1 0, id 1 0]" + +text {* + The recursive function used to do arithmatic subtraction. +*} +definition rec_minus :: "recf" + where + "rec_minus = Pr 1 (id 1 0) (Cn 3 rec_pred [id 3 2])" + +text {* + @{text "constn n"} is the recursive function which computes + nature number @{text "n"}. +*} +fun constn :: "nat \ recf" + where + "constn 0 = z" | + "constn (Suc n) = Cn 1 s [constn n]" + + +text {* + Signal function, which returns 1 when the input argument is greater than @{text "0"}. +*} +definition rec_sg :: "recf" + where + "rec_sg = Cn 1 rec_minus [constn 1, + Cn 1 rec_minus [constn 1, id 1 0]]" + +text {* + @{text "rec_less"} compares its two arguments, returns @{text "1"} if + the first is less than the second; otherwise returns @{text "0"}. + *} +definition rec_less :: "recf" + where + "rec_less = Cn 2 rec_sg [Cn 2 rec_minus [id 2 1, id 2 0]]" + +text {* + @{text "rec_not"} inverse its argument: returns @{text "1"} when the + argument is @{text "0"}; returns @{text "0"} otherwise. + *} +definition rec_not :: "recf" + where + "rec_not = Cn 1 rec_minus [constn 1, id 1 0]" + +text {* + @{text "rec_eq"} compares its two arguments: returns @{text "1"} + if they are equal; return @{text "0"} otherwise. + *} +definition rec_eq :: "recf" + where + "rec_eq = Cn 2 rec_minus [Cn 2 (constn 1) [id 2 0], + Cn 2 rec_add [Cn 2 rec_minus [id 2 0, id 2 1], + Cn 2 rec_minus [id 2 1, id 2 0]]]" + +text {* + @{text "rec_conj"} computes the conjunction of its two arguments, + returns @{text "1"} if both of them are non-zero; returns @{text "0"} + otherwise. + *} +definition rec_conj :: "recf" + where + "rec_conj = Cn 2 rec_sg [Cn 2 rec_mult [id 2 0, id 2 1]] " + +text {* + @{text "rec_disj"} computes the disjunction of its two arguments, + returns @{text "0"} if both of them are zero; returns @{text "0"} + otherwise. + *} +definition rec_disj :: "recf" + where + "rec_disj = Cn 2 rec_sg [Cn 2 rec_add [id 2 0, id 2 1]]" + + +text {* + Computes the arity of recursive function. + *} + +fun arity :: "recf \ nat" + where + "arity z = 1" +| "arity s = 1" +| "arity (id m n) = m" +| "arity (Cn n f gs) = n" +| "arity (Pr n f g) = Suc n" +| "arity (Mn n f) = n" + +text {* + @{text "get_fstn_args n (Suc k)"} returns + @{text "[id n 0, id n 1, id n 2, \, id n k]"}, + the effect of which is to take out the first @{text "Suc k"} + arguments out of the @{text "n"} input arguments. + *} +(* get_fstn_args *) +fun get_fstn_args :: "nat \ nat \ recf list" + where + "get_fstn_args n 0 = []" +| "get_fstn_args n (Suc y) = get_fstn_args n y @ [id n y]" + +text {* + @{text "rec_sigma f"} returns the recursive functions which + sums up the results of @{text "f"}: + \[ + (rec\_sigma f)(x, y) = f(x, 0) + f(x, 1) + \cdots + f(x, y) + \] +*} +fun rec_sigma :: "recf \ recf" + where + "rec_sigma rf = + (let vl = arity rf in + Pr (vl - 1) (Cn (vl - 1) rf (get_fstn_args (vl - 1) (vl - 1) @ + [Cn (vl - 1) (constn 0) [id (vl - 1) 0]])) + (Cn (Suc vl) rec_add [id (Suc vl) vl, + Cn (Suc vl) rf (get_fstn_args (Suc vl) (vl - 1) + @ [Cn (Suc vl) s [id (Suc vl) (vl - 1)]])]))" + +text {* + @{text "rec_exec"} is the interpreter function for + reursive functions. The function is defined such that + it always returns meaningful results for primitive recursive + functions. + *} +function rec_exec :: "recf \ nat list \ nat" + where + "rec_exec z xs = 0" | + "rec_exec s xs = (Suc (xs ! 0))" | + "rec_exec (id m n) xs = (xs ! n)" | + "rec_exec (Cn n f gs) xs = + (let ys = (map (\ a. rec_exec a xs) gs) in + rec_exec f ys)" | + "rec_exec (Pr n f g) xs = + (if last xs = 0 then + rec_exec f (butlast xs) + else rec_exec g (butlast xs @ [last xs - 1] @ + [rec_exec (Pr n f g) (butlast xs @ [last xs - 1])]))" | + "rec_exec (Mn n f) xs = (LEAST x. rec_exec f (xs @ [x]) = 0)" +by pat_completeness auto +termination +proof + show "wf (measures [\ (r, xs). size r, (\ (r, xs). last xs)])" + by auto +next + fix n f gs xs x + assume "(x::recf) \ set gs" + thus "((x, xs), Cn n f gs, xs) \ + measures [\(r, xs). size r, \(r, xs). last xs]" + by(induct gs, auto) +next + fix n f gs xs x + assume "x = map (\a. rec_exec a xs) gs" + "\x. x \ set gs \ rec_exec_dom (x, xs)" + thus "((f, x), Cn n f gs, xs) \ + measures [\(r, xs). size r, \(r, xs). last xs]" + by(auto) +next + fix n f g xs + show "((f, butlast xs), Pr n f g, xs) \ + measures [\(r, xs). size r, \(r, xs). last xs]" + by auto +next + fix n f g xs + assume "last xs \ (0::nat)" thus + "((Pr n f g, butlast xs @ [last xs - 1]), Pr n f g, xs) + \ measures [\(r, xs). size r, \(r, xs). last xs]" + by auto +next + fix n f g xs + show "((g, butlast xs @ [last xs - 1] @ [rec_exec (Pr n f g) (butlast xs @ [last xs - 1])]), + Pr n f g, xs) \ measures [\(r, xs). size r, \(r, xs). last xs]" + by auto +next + fix n f xs x + show "((f, xs @ [x]), Mn n f, xs) \ + measures [\(r, xs). size r, \(r, xs). last xs]" + by auto +qed + +declare rec_exec.simps[simp del] constn.simps[simp del] + +text {* + Correctness of @{text "rec_add"}. + *} +lemma add_lemma: "\ x y. rec_exec rec_add [x, y] = x + y" +by(induct_tac y, auto simp: rec_add_def rec_exec.simps) + +text {* + Correctness of @{text "rec_mult"}. + *} +lemma mult_lemma: "\ x y. rec_exec rec_mult [x, y] = x * y" +by(induct_tac y, auto simp: rec_mult_def rec_exec.simps add_lemma) + +text {* + Correctness of @{text "rec_pred"}. + *} +lemma pred_lemma: "\ x. rec_exec rec_pred [x] = x - 1" +by(induct_tac x, auto simp: rec_pred_def rec_exec.simps) + +text {* + Correctness of @{text "rec_minus"}. + *} +lemma minus_lemma: "\ x y. rec_exec rec_minus [x, y] = x - y" +by(induct_tac y, auto simp: rec_exec.simps rec_minus_def pred_lemma) + +text {* + Correctness of @{text "rec_sg"}. + *} +lemma sg_lemma: "\ x. rec_exec rec_sg [x] = (if x = 0 then 0 else 1)" +by(auto simp: rec_sg_def minus_lemma rec_exec.simps constn.simps) + +text {* + Correctness of @{text "constn"}. + *} +lemma constn_lemma: "rec_exec (constn n) [x] = n" +by(induct n, auto simp: rec_exec.simps constn.simps) + +text {* + Correctness of @{text "rec_less"}. + *} +lemma less_lemma: "\ x y. rec_exec rec_less [x, y] = + (if x < y then 1 else 0)" +by(induct_tac y, auto simp: rec_exec.simps + rec_less_def minus_lemma sg_lemma) + +text {* + Correctness of @{text "rec_not"}. + *} +lemma not_lemma: + "\ x. rec_exec rec_not [x] = (if x = 0 then 1 else 0)" +by(induct_tac x, auto simp: rec_exec.simps rec_not_def + constn_lemma minus_lemma) + +text {* + Correctness of @{text "rec_eq"}. + *} +lemma eq_lemma: "\ x y. rec_exec rec_eq [x, y] = (if x = y then 1 else 0)" +by(induct_tac y, auto simp: rec_exec.simps rec_eq_def constn_lemma add_lemma minus_lemma) + +text {* + Correctness of @{text "rec_conj"}. + *} +lemma conj_lemma: "\ x y. rec_exec rec_conj [x, y] = (if x = 0 \ y = 0 then 0 + else 1)" +by(induct_tac y, auto simp: rec_exec.simps sg_lemma rec_conj_def mult_lemma) + + +text {* + Correctness of @{text "rec_disj"}. + *} +lemma disj_lemma: "\ x y. rec_exec rec_disj [x, y] = (if x = 0 \ y = 0 then 0 + else 1)" +by(induct_tac y, auto simp: rec_disj_def sg_lemma add_lemma rec_exec.simps) + + +text {* + @{text "primrec recf n"} is true iff + @{text "recf"} is a primitive recursive function + with arity @{text "n"}. + *} +inductive primerec :: "recf \ nat \ bool" + where +prime_z[intro]: "primerec z (Suc 0)" | +prime_s[intro]: "primerec s (Suc 0)" | +prime_id[intro!]: "\n < m\ \ primerec (id m n) m" | +prime_cn[intro!]: "\primerec f k; length gs = k; + \ i < length gs. primerec (gs ! i) m; m = n\ + \ primerec (Cn n f gs) m" | +prime_pr[intro!]: "\primerec f n; + primerec g (Suc (Suc n)); m = Suc n\ + \ primerec (Pr n f g) m" + +inductive_cases prime_cn_reverse'[elim]: "primerec (Cn n f gs) n" +inductive_cases prime_mn_reverse: "primerec (Mn n f) m" +inductive_cases prime_z_reverse[elim]: "primerec z n" +inductive_cases prime_s_reverse[elim]: "primerec s n" +inductive_cases prime_id_reverse[elim]: "primerec (id m n) k" +inductive_cases prime_cn_reverse[elim]: "primerec (Cn n f gs) m" +inductive_cases prime_pr_reverse[elim]: "primerec (Pr n f g) m" + +declare mult_lemma[simp] add_lemma[simp] pred_lemma[simp] + minus_lemma[simp] sg_lemma[simp] constn_lemma[simp] + less_lemma[simp] not_lemma[simp] eq_lemma[simp] + conj_lemma[simp] disj_lemma[simp] + +text {* + @{text "Sigma"} is the logical specification of + the recursive function @{text "rec_sigma"}. + *} +function Sigma :: "(nat list \ nat) \ nat list \ nat" + where + "Sigma g xs = (if last xs = 0 then g xs + else (Sigma g (butlast xs @ [last xs - 1]) + + g xs)) " +by pat_completeness auto +termination +proof + show "wf (measure (\ (f, xs). last xs))" by auto +next + fix g xs + assume "last (xs::nat list) \ 0" + thus "((g, butlast xs @ [last xs - 1]), g, xs) + \ measure (\(f, xs). last xs)" + by auto +qed + +declare rec_exec.simps[simp del] get_fstn_args.simps[simp del] + arity.simps[simp del] Sigma.simps[simp del] + rec_sigma.simps[simp del] + + +section {* Properties of @{text rec_sigma} *} + +lemma [simp]: "arity z = 1" + by(simp add: arity.simps) + +lemma rec_pr_0_simp_rewrite: " + rec_exec (Pr n f g) (xs @ [0]) = rec_exec f xs" +by(simp add: rec_exec.simps) + +lemma rec_pr_0_simp_rewrite_single_param: " + rec_exec (Pr n f g) [0] = rec_exec f []" +by(simp add: rec_exec.simps) + +lemma rec_pr_Suc_simp_rewrite: + "rec_exec (Pr n f g) (xs @ [Suc x]) = + rec_exec g (xs @ [x] @ + [rec_exec (Pr n f g) (xs @ [x])])" +by(simp add: rec_exec.simps) + +lemma rec_pr_Suc_simp_rewrite_single_param: + "rec_exec (Pr n f g) ([Suc x]) = + rec_exec g ([x] @ [rec_exec (Pr n f g) ([x])])" +by(simp add: rec_exec.simps) + +thm Sigma.simps + +lemma Sigma_0_simp_rewrite_single_param: + "Sigma f [0] = f [0]" +by(simp add: Sigma.simps) + +lemma Sigma_0_simp_rewrite: + "Sigma f (xs @ [0]) = f (xs @ [0])" +by(simp add: Sigma.simps) + +lemma Sigma_Suc_simp_rewrite: + "Sigma f (xs @ [Suc x]) = Sigma f (xs @ [x]) + f (xs @ [Suc x])" +by(simp add: Sigma.simps) + +lemma Sigma_Suc_simp_rewrite_single: + "Sigma f ([Suc x]) = Sigma f ([x]) + f ([Suc x])" +by(simp add: Sigma.simps) + +lemma [simp]: "(xs @ ys) ! (Suc (length xs)) = ys ! 1" +by(simp add: nth_append) + +lemma get_fstn_args_take: "\length xs = m; n \ m\ \ + map (\ f. rec_exec f xs) (get_fstn_args m n)= take n xs" +proof(induct n) + case 0 thus "?case" + by(simp add: get_fstn_args.simps) +next + case (Suc n) thus "?case" + by(simp add: get_fstn_args.simps rec_exec.simps + take_Suc_conv_app_nth) +qed + +lemma [simp]: "primerec f n \ arity f = n" + apply(case_tac f) + apply(auto simp: arity.simps ) + apply(erule_tac prime_mn_reverse) + done + +lemma rec_sigma_Suc_simp_rewrite: + "primerec f (Suc (length xs)) + \ rec_exec (rec_sigma f) (xs @ [Suc x]) = + rec_exec (rec_sigma f) (xs @ [x]) + rec_exec f (xs @ [Suc x])" + apply(induct x) + apply(auto simp: rec_sigma.simps Let_def rec_pr_Suc_simp_rewrite + rec_exec.simps get_fstn_args_take) + done + +text {* + The correctness of @{text "rec_sigma"} with respect to its specification. + *} +lemma sigma_lemma: + "primerec rg (Suc (length xs)) + \ rec_exec (rec_sigma rg) (xs @ [x]) = Sigma (rec_exec rg) (xs @ [x])" +apply(induct x) +apply(auto simp: rec_exec.simps rec_sigma.simps Let_def + get_fstn_args_take Sigma_0_simp_rewrite + Sigma_Suc_simp_rewrite) +done + +text {* + @{text "rec_accum f (x1, x2, \, xn, k) = + f(x1, x2, \, xn, 0) * + f(x1, x2, \, xn, 1) * + \ + f(x1, x2, \, xn, k)"} +*} +fun rec_accum :: "recf \ recf" + where + "rec_accum rf = + (let vl = arity rf in + Pr (vl - 1) (Cn (vl - 1) rf (get_fstn_args (vl - 1) (vl - 1) @ + [Cn (vl - 1) (constn 0) [id (vl - 1) 0]])) + (Cn (Suc vl) rec_mult [id (Suc vl) (vl), + Cn (Suc vl) rf (get_fstn_args (Suc vl) (vl - 1) + @ [Cn (Suc vl) s [id (Suc vl) (vl - 1)]])]))" + +text {* + @{text "Accum"} is the formal specification of @{text "rec_accum"}. + *} +function Accum :: "(nat list \ nat) \ nat list \ nat" + where + "Accum f xs = (if last xs = 0 then f xs + else (Accum f (butlast xs @ [last xs - 1]) * + f xs))" +by pat_completeness auto +termination +proof + show "wf (measure (\ (f, xs). last xs))" + by auto +next + fix f xs + assume "last xs \ (0::nat)" + thus "((f, butlast xs @ [last xs - 1]), f, xs) \ + measure (\(f, xs). last xs)" + by auto +qed + +lemma rec_accum_Suc_simp_rewrite: + "primerec f (Suc (length xs)) + \ rec_exec (rec_accum f) (xs @ [Suc x]) = + rec_exec (rec_accum f) (xs @ [x]) * rec_exec f (xs @ [Suc x])" + apply(induct x) + apply(auto simp: rec_sigma.simps Let_def rec_pr_Suc_simp_rewrite + rec_exec.simps get_fstn_args_take) + done + +text {* + The correctness of @{text "rec_accum"} with respect to its specification. +*} +lemma accum_lemma : + "primerec rg (Suc (length xs)) + \ rec_exec (rec_accum rg) (xs @ [x]) = Accum (rec_exec rg) (xs @ [x])" +apply(induct x) +apply(auto simp: rec_exec.simps rec_sigma.simps Let_def + get_fstn_args_take) +done + +declare rec_accum.simps [simp del] + +text {* + @{text "rec_all t f (x1, x2, \, xn)"} + computes the charactrization function of the following FOL formula: + @{text "(\ x \ t(x1, x2, \, xn). (f(x1, x2, \, xn, x) > 0))"} +*} +fun rec_all :: "recf \ recf \ recf" + where + "rec_all rt rf = + (let vl = arity rf in + Cn (vl - 1) rec_sg [Cn (vl - 1) (rec_accum rf) + (get_fstn_args (vl - 1) (vl - 1) @ [rt])])" + +lemma rec_accum_ex: "primerec rf (Suc (length xs)) \ + (rec_exec (rec_accum rf) (xs @ [x]) = 0) = + (\ t \ x. rec_exec rf (xs @ [t]) = 0)" +apply(induct x, simp_all add: rec_accum_Suc_simp_rewrite) +apply(simp add: rec_exec.simps rec_accum.simps get_fstn_args_take, + auto) +apply(rule_tac x = ta in exI, simp) +apply(case_tac "t = Suc x", simp_all) +apply(rule_tac x = t in exI, simp) +done + +text {* + The correctness of @{text "rec_all"}. + *} +lemma all_lemma: + "\primerec rf (Suc (length xs)); + primerec rt (length xs)\ + \ rec_exec (rec_all rt rf) xs = (if (\ x \ (rec_exec rt xs). 0 < rec_exec rf (xs @ [x])) then 1 + else 0)" +apply(auto simp: rec_all.simps) +apply(simp add: rec_exec.simps map_append get_fstn_args_take split: if_splits) +apply(drule_tac x = "rec_exec rt xs" in rec_accum_ex) +apply(case_tac "rec_exec (rec_accum rf) (xs @ [rec_exec rt xs]) = 0", simp_all) +apply(erule_tac exE, erule_tac x = t in allE, simp) +apply(simp add: rec_exec.simps map_append get_fstn_args_take) +apply(drule_tac x = "rec_exec rt xs" in rec_accum_ex) +apply(case_tac "rec_exec (rec_accum rf) (xs @ [rec_exec rt xs]) = 0", simp, simp) +apply(erule_tac x = x in allE, simp) +done + +text {* + @{text "rec_ex t f (x1, x2, \, xn)"} + computes the charactrization function of the following FOL formula: + @{text "(\ x \ t(x1, x2, \, xn). (f(x1, x2, \, xn, x) > 0))"} +*} +fun rec_ex :: "recf \ recf \ recf" + where + "rec_ex rt rf = + (let vl = arity rf in + Cn (vl - 1) rec_sg [Cn (vl - 1) (rec_sigma rf) + (get_fstn_args (vl - 1) (vl - 1) @ [rt])])" + +lemma rec_sigma_ex: "primerec rf (Suc (length xs)) + \ (rec_exec (rec_sigma rf) (xs @ [x]) = 0) = + (\ t \ x. rec_exec rf (xs @ [t]) = 0)" +apply(induct x, simp_all add: rec_sigma_Suc_simp_rewrite) +apply(simp add: rec_exec.simps rec_sigma.simps + get_fstn_args_take, auto) +apply(case_tac "t = Suc x", simp_all) +done + +text {* + The correctness of @{text "ex_lemma"}. + *} +lemma ex_lemma:" + \primerec rf (Suc (length xs)); + primerec rt (length xs)\ +\ (rec_exec (rec_ex rt rf) xs = + (if (\ x \ (rec_exec rt xs). 0 bool) \ nat list \ nat \ nat" + where "Minr Rr xs w = (let setx = {y | y. (y \ w) \ Rr (xs @ [y])} in + if (setx = {}) then (Suc w) + else (Min setx))" + +declare Minr.simps[simp del] rec_all.simps[simp del] + +text {* + The following is a set of auxilliary lemmas about @{text "Minr"}. +*} +lemma Minr_range: "Minr Rr xs w \ w \ Minr Rr xs w = Suc w" +apply(auto simp: Minr.simps) +apply(subgoal_tac "Min {x. x \ w \ Rr (xs @ [x])} \ x") +apply(erule_tac order_trans, simp) +apply(rule_tac Min_le, auto) +done + +lemma [simp]: "{x. x \ Suc w \ Rr (xs @ [x])} + = (if Rr (xs @ [Suc w]) then insert (Suc w) + {x. x \ w \ Rr (xs @ [x])} + else {x. x \ w \ Rr (xs @ [x])})" +by(auto, case_tac "x = Suc w", auto) + +lemma [simp]: "Minr Rr xs w \ w \ Minr Rr xs (Suc w) = Minr Rr xs w" +apply(simp add: Minr.simps, auto) +apply(case_tac "\x\w. \ Rr (xs @ [x])", auto) +done + +lemma [simp]: "\x\w. \ Rr (xs @ [x]) \ + {x. x \ w \ Rr (xs @ [x])} = {} " +by auto + +lemma [simp]: "\Minr Rr xs w = Suc w; Rr (xs @ [Suc w])\ \ + Minr Rr xs (Suc w) = Suc w" +apply(simp add: Minr.simps) +apply(case_tac "\x\w. \ Rr (xs @ [x])", auto) +done + +lemma [simp]: "\Minr Rr xs w = Suc w; \ Rr (xs @ [Suc w])\ \ + Minr Rr xs (Suc w) = Suc (Suc w)" +apply(simp add: Minr.simps) +apply(case_tac "\x\w. \ Rr (xs @ [x])", auto) +apply(subgoal_tac "Min {x. x \ w \ Rr (xs @ [x])} \ + {x. x \ w \ Rr (xs @ [x])}", simp) +apply(rule_tac Min_in, auto) +done + +lemma Minr_Suc_simp: + "Minr Rr xs (Suc w) = + (if Minr Rr xs w \ w then Minr Rr xs w + else if (Rr (xs @ [Suc w])) then (Suc w) + else Suc (Suc w))" +by(insert Minr_range[of Rr xs w], auto) + +text {* + @{text "rec_Minr"} is the recursive function + used to implement @{text "Minr"}: + if @{text "Rr"} is implemented by a recursive function @{text "recf"}, + then @{text "rec_Minr recf"} is the recursive function used to + implement @{text "Minr Rr"} + *} +fun rec_Minr :: "recf \ recf" + where + "rec_Minr rf = + (let vl = arity rf + in let rq = rec_all (id vl (vl - 1)) (Cn (Suc vl) + rec_not [Cn (Suc vl) rf + (get_fstn_args (Suc vl) (vl - 1) @ + [id (Suc vl) (vl)])]) + in rec_sigma rq)" + +lemma length_getpren_params[simp]: "length (get_fstn_args m n) = n" +by(induct n, auto simp: get_fstn_args.simps) + +lemma length_app: + "(length (get_fstn_args (arity rf - Suc 0) + (arity rf - Suc 0) + @ [Cn (arity rf - Suc 0) (constn 0) + [recf.id (arity rf - Suc 0) 0]])) + = (Suc (arity rf - Suc 0))" + apply(simp) +done + +lemma primerec_accum: "primerec (rec_accum rf) n \ primerec rf n" +apply(auto simp: rec_accum.simps Let_def) +apply(erule_tac prime_pr_reverse, simp) +apply(erule_tac prime_cn_reverse, simp only: length_app) +done + +lemma primerec_all: "primerec (rec_all rt rf) n \ + primerec rt n \ primerec rf (Suc n)" +apply(simp add: rec_all.simps Let_def) +apply(erule_tac prime_cn_reverse, simp) +apply(erule_tac prime_cn_reverse, simp) +apply(erule_tac x = n in allE, simp add: nth_append primerec_accum) +done + +lemma min_Suc_Suc[simp]: "min (Suc (Suc x)) x = x" + by auto + +declare numeral_3_eq_3[simp] + +lemma [intro]: "primerec rec_pred (Suc 0)" +apply(simp add: rec_pred_def) +apply(rule_tac prime_cn, auto) +apply(case_tac i, auto intro: prime_id) +done + +lemma [intro]: "primerec rec_minus (Suc (Suc 0))" + apply(auto simp: rec_minus_def) + done + +lemma [intro]: "primerec (constn n) (Suc 0)" + apply(induct n) + apply(auto simp: constn.simps intro: prime_z prime_cn prime_s) + done + +lemma [intro]: "primerec rec_sg (Suc 0)" + apply(simp add: rec_sg_def) + apply(rule_tac k = "Suc (Suc 0)" in prime_cn, auto) + apply(case_tac i, auto) + apply(case_tac ia, auto intro: prime_id) + done + +lemma [simp]: "length (get_fstn_args m n) = n" + apply(induct n) + apply(auto simp: get_fstn_args.simps) + done + +lemma primerec_getpren[elim]: "\i < n; n \ m\ \ primerec (get_fstn_args m n ! i) m" +apply(induct n, auto simp: get_fstn_args.simps) +apply(case_tac "i = n", auto simp: nth_append intro: prime_id) +done + +lemma [intro]: "primerec rec_add (Suc (Suc 0))" +apply(simp add: rec_add_def) +apply(rule_tac prime_pr, auto) +done + +lemma [intro]:"primerec rec_mult (Suc (Suc 0))" +apply(simp add: rec_mult_def ) +apply(rule_tac prime_pr, auto intro: prime_z) +apply(case_tac i, auto intro: prime_id) +done + +lemma [elim]: "\primerec rf n; n \ Suc (Suc 0)\ \ + primerec (rec_accum rf) n" +apply(auto simp: rec_accum.simps) +apply(simp add: nth_append, auto) +apply(case_tac i, auto intro: prime_id) +apply(auto simp: nth_append) +done + +lemma primerec_all_iff: + "\primerec rt n; primerec rf (Suc n); n > 0\ \ + primerec (rec_all rt rf) n" + apply(simp add: rec_all.simps, auto) + apply(auto, simp add: nth_append, auto) + done + +lemma [simp]: "Rr (xs @ [0]) \ + Min {x. x = (0::nat) \ Rr (xs @ [x])} = 0" +by(rule_tac Min_eqI, simp, simp, simp) + +lemma [intro]: "primerec rec_not (Suc 0)" +apply(simp add: rec_not_def) +apply(rule prime_cn, auto) +apply(case_tac i, auto intro: prime_id) +done + +lemma Min_false1[simp]: "\\ Min {uu. uu \ w \ 0 < rec_exec rf (xs @ [uu])} \ w; + x \ w; 0 < rec_exec rf (xs @ [x])\ + \ False" +apply(subgoal_tac "finite {uu. uu \ w \ 0 < rec_exec rf (xs @ [uu])}") +apply(subgoal_tac "{uu. uu \ w \ 0 < rec_exec rf (xs @ [uu])} \ {}") +apply(simp add: Min_le_iff, simp) +apply(rule_tac x = x in exI, simp) +apply(simp) +done + +lemma sigma_minr_lemma: + assumes prrf: "primerec rf (Suc (length xs))" + shows "UF.Sigma (rec_exec (rec_all (recf.id (Suc (length xs)) (length xs)) + (Cn (Suc (Suc (length xs))) rec_not + [Cn (Suc (Suc (length xs))) rf (get_fstn_args (Suc (Suc (length xs))) + (length xs) @ [recf.id (Suc (Suc (length xs))) (Suc (length xs))])]))) + (xs @ [w]) = + Minr (\args. 0 < rec_exec rf args) xs w" +proof(induct w) + let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" + let ?rf = "(Cn (Suc (Suc (length xs))) + rec_not [Cn (Suc (Suc (length xs))) rf + (get_fstn_args (Suc (Suc (length xs))) (length xs) @ + [recf.id (Suc (Suc (length xs))) + (Suc ((length xs)))])])" + let ?rq = "(rec_all ?rt ?rf)" + have prrf: "primerec ?rf (Suc (length (xs @ [0]))) \ + primerec ?rt (length (xs @ [0]))" + apply(auto simp: prrf nth_append)+ + done + show "Sigma (rec_exec (rec_all ?rt ?rf)) (xs @ [0]) + = Minr (\args. 0 < rec_exec rf args) xs 0" + apply(simp add: Sigma.simps) + apply(simp only: prrf all_lemma, + auto simp: rec_exec.simps get_fstn_args_take Minr.simps) + apply(rule_tac Min_eqI, auto) + done +next + fix w + let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" + let ?rf = "(Cn (Suc (Suc (length xs))) + rec_not [Cn (Suc (Suc (length xs))) rf + (get_fstn_args (Suc (Suc (length xs))) (length xs) @ + [recf.id (Suc (Suc (length xs))) + (Suc ((length xs)))])])" + let ?rq = "(rec_all ?rt ?rf)" + assume ind: + "Sigma (rec_exec (rec_all ?rt ?rf)) (xs @ [w]) = Minr (\args. 0 < rec_exec rf args) xs w" + have prrf: "primerec ?rf (Suc (length (xs @ [Suc w]))) \ + primerec ?rt (length (xs @ [Suc w]))" + apply(auto simp: prrf nth_append)+ + done + show "UF.Sigma (rec_exec (rec_all ?rt ?rf)) + (xs @ [Suc w]) = + Minr (\args. 0 < rec_exec rf args) xs (Suc w)" + apply(auto simp: Sigma_Suc_simp_rewrite ind Minr_Suc_simp) + apply(simp_all only: prrf all_lemma) + apply(auto simp: rec_exec.simps get_fstn_args_take Let_def Minr.simps split: if_splits) + apply(drule_tac Min_false1, simp, simp, simp) + apply(case_tac "x = Suc w", simp, simp) + apply(drule_tac Min_false1, simp, simp, simp) + apply(drule_tac Min_false1, simp, simp, simp) + done +qed + +text {* + The correctness of @{text "rec_Minr"}. + *} +lemma Minr_lemma: " + \primerec rf (Suc (length xs))\ + \ rec_exec (rec_Minr rf) (xs @ [w]) = + Minr (\ args. (0 < rec_exec rf args)) xs w" +proof - + let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" + let ?rf = "(Cn (Suc (Suc (length xs))) + rec_not [Cn (Suc (Suc (length xs))) rf + (get_fstn_args (Suc (Suc (length xs))) (length xs) @ + [recf.id (Suc (Suc (length xs))) + (Suc ((length xs)))])])" + let ?rq = "(rec_all ?rt ?rf)" + assume h: "primerec rf (Suc (length xs))" + have h1: "primerec ?rq (Suc (length xs))" + apply(rule_tac primerec_all_iff) + apply(auto simp: h nth_append)+ + done + moreover have "arity rf = Suc (length xs)" + using h by auto + ultimately show "rec_exec (rec_Minr rf) (xs @ [w]) = + Minr (\ args. (0 < rec_exec rf args)) xs w" + apply(simp add: rec_exec.simps rec_Minr.simps arity.simps Let_def + sigma_lemma all_lemma) + apply(rule_tac sigma_minr_lemma) + apply(simp add: h) + done +qed + +text {* + @{text "rec_le"} is the comparasion function + which compares its two arguments, testing whether the + first is less or equal to the second. + *} +definition rec_le :: "recf" + where + "rec_le = Cn (Suc (Suc 0)) rec_disj [rec_less, rec_eq]" + +text {* + The correctness of @{text "rec_le"}. + *} +lemma le_lemma: + "\x y. rec_exec rec_le [x, y] = (if (x \ y) then 1 else 0)" +by(auto simp: rec_le_def rec_exec.simps) + +text {* + Defintiion of @{text "Max[Rr]"} on page 77 of Boolos's book. +*} + +fun Maxr :: "(nat list \ bool) \ nat list \ nat \ nat" + where + "Maxr Rr xs w = (let setx = {y. y \ w \ Rr (xs @[y])} in + if setx = {} then 0 + else Max setx)" + +text {* + @{text "rec_maxr"} is the recursive function + used to implementation @{text "Maxr"}. + *} +fun rec_maxr :: "recf \ recf" + where + "rec_maxr rr = (let vl = arity rr in + let rt = id (Suc vl) (vl - 1) in + let rf1 = Cn (Suc (Suc vl)) rec_le + [id (Suc (Suc vl)) + ((Suc vl)), id (Suc (Suc vl)) (vl)] in + let rf2 = Cn (Suc (Suc vl)) rec_not + [Cn (Suc (Suc vl)) + rr (get_fstn_args (Suc (Suc vl)) + (vl - 1) @ + [id (Suc (Suc vl)) ((Suc vl))])] in + let rf = Cn (Suc (Suc vl)) rec_disj [rf1, rf2] in + let rq = rec_all rt rf in + let Qf = Cn (Suc vl) rec_not [rec_all rt rf] + in Cn vl (rec_sigma Qf) (get_fstn_args vl vl @ + [id vl (vl - 1)]))" + +declare rec_maxr.simps[simp del] Maxr.simps[simp del] +declare le_lemma[simp] +lemma [simp]: "(min (Suc (Suc (Suc (x)))) (x)) = x" +by simp + +declare numeral_2_eq_2[simp] + +lemma [intro]: "primerec rec_disj (Suc (Suc 0))" + apply(simp add: rec_disj_def, auto) + apply(auto) + apply(case_tac ia, auto intro: prime_id) + done + +lemma [intro]: "primerec rec_less (Suc (Suc 0))" + apply(simp add: rec_less_def, auto) + apply(auto) + apply(case_tac ia , auto intro: prime_id) + done + +lemma [intro]: "primerec rec_eq (Suc (Suc 0))" + apply(simp add: rec_eq_def) + apply(rule_tac prime_cn, auto) + apply(case_tac i, auto) + apply(case_tac ia, auto) + apply(case_tac [!] i, auto intro: prime_id) + done + +lemma [intro]: "primerec rec_le (Suc (Suc 0))" + apply(simp add: rec_le_def) + apply(rule_tac prime_cn, auto) + apply(case_tac i, auto) + done + +lemma [simp]: + "length ys = Suc n \ (take n ys @ [ys ! n, ys ! n]) = + ys @ [ys ! n]" +apply(simp) +apply(subgoal_tac "\ xs y. ys = xs @ [y]", auto) +apply(rule_tac x = "butlast ys" in exI, rule_tac x = "last ys" in exI) +apply(case_tac "ys = []", simp_all) +done + +lemma Maxr_Suc_simp: + "Maxr Rr xs (Suc w) =(if Rr (xs @ [Suc w]) then Suc w + else Maxr Rr xs w)" +apply(auto simp: Maxr.simps) +apply(rule_tac max_absorb1) +apply(subgoal_tac "(Max {y. y \ w \ Rr (xs @ [y])} \ (Suc w)) = + (\a\{y. y \ w \ Rr (xs @ [y])}. a \ (Suc w))", simp) +apply(rule_tac Max_le_iff, auto) +done + + +lemma [simp]: "min (Suc n) n = n" by simp + +lemma Sigma_0: "\ i \ n. (f (xs @ [i]) = 0) \ + Sigma f (xs @ [n]) = 0" +apply(induct n, simp add: Sigma.simps) +apply(simp add: Sigma_Suc_simp_rewrite) +done + +lemma [elim]: "\k Sigma f (xs @ [w]) = Suc w" +apply(induct w) +apply(simp add: Sigma.simps, simp) +apply(simp add: Sigma.simps) +done + +lemma Sigma_max_point: "\\ k < ma. f (xs @ [k]) = 1; + \ k \ ma. f (xs @ [k]) = 0; ma \ w\ + \ Sigma f (xs @ [w]) = ma" +apply(induct w, auto) +apply(rule_tac Sigma_0, simp) +apply(simp add: Sigma_Suc_simp_rewrite) +apply(case_tac "ma = Suc w", auto) +done + +lemma Sigma_Max_lemma: + assumes prrf: "primerec rf (Suc (length xs))" + shows "UF.Sigma (rec_exec (Cn (Suc (Suc (length xs))) rec_not + [rec_all (recf.id (Suc (Suc (length xs))) (length xs)) + (Cn (Suc (Suc (Suc (length xs)))) rec_disj + [Cn (Suc (Suc (Suc (length xs)))) rec_le + [recf.id (Suc (Suc (Suc (length xs)))) (Suc (Suc (length xs))), + recf.id (Suc (Suc (Suc (length xs)))) (Suc (length xs))], + Cn (Suc (Suc (Suc (length xs)))) rec_not + [Cn (Suc (Suc (Suc (length xs)))) rf + (get_fstn_args (Suc (Suc (Suc (length xs)))) (length xs) @ + [recf.id (Suc (Suc (Suc (length xs)))) (Suc (Suc (length xs)))])]])])) + ((xs @ [w]) @ [w]) = + Maxr (\args. 0 < rec_exec rf args) xs w" +proof - + let ?rt = "(recf.id (Suc (Suc (length xs))) ((length xs)))" + let ?rf1 = "Cn (Suc (Suc (Suc (length xs)))) + rec_le [recf.id (Suc (Suc (Suc (length xs)))) + ((Suc (Suc (length xs)))), recf.id + (Suc (Suc (Suc (length xs)))) ((Suc (length xs)))]" + let ?rf2 = "Cn (Suc (Suc (Suc (length xs)))) rf + (get_fstn_args (Suc (Suc (Suc (length xs)))) + (length xs) @ + [recf.id (Suc (Suc (Suc (length xs)))) + ((Suc (Suc (length xs))))])" + let ?rf3 = "Cn (Suc (Suc (Suc (length xs)))) rec_not [?rf2]" + let ?rf = "Cn (Suc (Suc (Suc (length xs)))) rec_disj [?rf1, ?rf3]" + let ?rq = "rec_all ?rt ?rf" + let ?notrq = "Cn (Suc (Suc (length xs))) rec_not [?rq]" + show "?thesis" + proof(auto simp: Maxr.simps) + assume h: "\x\w. rec_exec rf (xs @ [x]) = 0" + have "primerec ?rf (Suc (length (xs @ [w, i]))) \ + primerec ?rt (length (xs @ [w, i]))" + using prrf + apply(auto) + apply(case_tac i, auto) + apply(case_tac ia, auto simp: h nth_append) + done + hence "Sigma (rec_exec ?notrq) ((xs@[w])@[w]) = 0" + apply(rule_tac Sigma_0) + apply(auto simp: rec_exec.simps all_lemma + get_fstn_args_take nth_append h) + done + thus "UF.Sigma (rec_exec ?notrq) + (xs @ [w, w]) = 0" + by simp + next + fix x + assume h: "x \ w" "0 < rec_exec rf (xs @ [x])" + hence "\ ma. Max {y. y \ w \ 0 < rec_exec rf (xs @ [y])} = ma" + by auto + from this obtain ma where k1: + "Max {y. y \ w \ 0 < rec_exec rf (xs @ [y])} = ma" .. + hence k2: "ma \ w \ 0 < rec_exec rf (xs @ [ma])" + using h + apply(subgoal_tac + "Max {y. y \ w \ 0 < rec_exec rf (xs @ [y])} \ {y. y \ w \ 0 < rec_exec rf (xs @ [y])}") + apply(erule_tac CollectE, simp) + apply(rule_tac Max_in, auto) + done + hence k3: "\ k < ma. (rec_exec ?notrq (xs @ [w, k]) = 1)" + apply(auto simp: nth_append) + apply(subgoal_tac "primerec ?rf (Suc (length (xs @ [w, k]))) \ + primerec ?rt (length (xs @ [w, k]))") + apply(auto simp: rec_exec.simps all_lemma get_fstn_args_take nth_append) + using prrf + apply(case_tac i, auto) + apply(case_tac ia, auto simp: h nth_append) + done + have k4: "\ k \ ma. (rec_exec ?notrq (xs @ [w, k]) = 0)" + apply(auto) + apply(subgoal_tac "primerec ?rf (Suc (length (xs @ [w, k]))) \ + primerec ?rt (length (xs @ [w, k]))") + apply(auto simp: rec_exec.simps all_lemma get_fstn_args_take nth_append) + apply(subgoal_tac "x \ Max {y. y \ w \ 0 < rec_exec rf (xs @ [y])}", + simp add: k1) + apply(rule_tac Max_ge, auto) + using prrf + apply(case_tac i, auto) + apply(case_tac ia, auto simp: h nth_append) + done + from k3 k4 k1 have "Sigma (rec_exec ?notrq) ((xs @ [w]) @ [w]) = ma" + apply(rule_tac Sigma_max_point, simp, simp, simp add: k2) + done + from k1 and this show "Sigma (rec_exec ?notrq) (xs @ [w, w]) = + Max {y. y \ w \ 0 < rec_exec rf (xs @ [y])}" + by simp + qed +qed + +text {* + The correctness of @{text "rec_maxr"}. + *} +lemma Maxr_lemma: + assumes h: "primerec rf (Suc (length xs))" + shows "rec_exec (rec_maxr rf) (xs @ [w]) = + Maxr (\ args. 0 < rec_exec rf args) xs w" +proof - + from h have "arity rf = Suc (length xs)" + by auto + thus "?thesis" + proof(simp add: rec_exec.simps rec_maxr.simps nth_append get_fstn_args_take) + let ?rt = "(recf.id (Suc (Suc (length xs))) ((length xs)))" + let ?rf1 = "Cn (Suc (Suc (Suc (length xs)))) + rec_le [recf.id (Suc (Suc (Suc (length xs)))) + ((Suc (Suc (length xs)))), recf.id + (Suc (Suc (Suc (length xs)))) ((Suc (length xs)))]" + let ?rf2 = "Cn (Suc (Suc (Suc (length xs)))) rf + (get_fstn_args (Suc (Suc (Suc (length xs)))) + (length xs) @ + [recf.id (Suc (Suc (Suc (length xs)))) + ((Suc (Suc (length xs))))])" + let ?rf3 = "Cn (Suc (Suc (Suc (length xs)))) rec_not [?rf2]" + let ?rf = "Cn (Suc (Suc (Suc (length xs)))) rec_disj [?rf1, ?rf3]" + let ?rq = "rec_all ?rt ?rf" + let ?notrq = "Cn (Suc (Suc (length xs))) rec_not [?rq]" + have prt: "primerec ?rt (Suc (Suc (length xs)))" + by(auto intro: prime_id) + have prrf: "primerec ?rf (Suc (Suc (Suc (length xs))))" + apply(auto) + apply(case_tac i, auto) + apply(case_tac ia, auto intro: prime_id) + apply(simp add: h) + apply(simp add: nth_append, auto intro: prime_id) + done + from prt and prrf have prrq: "primerec ?rq + (Suc (Suc (length xs)))" + by(erule_tac primerec_all_iff, auto) + hence prnotrp: "primerec ?notrq (Suc (length ((xs @ [w]))))" + by(rule_tac prime_cn, auto) + have g1: "rec_exec (rec_sigma ?notrq) ((xs @ [w]) @ [w]) + = Maxr (\args. 0 < rec_exec rf args) xs w" + using prnotrp + using sigma_lemma + apply(simp only: sigma_lemma) + apply(rule_tac Sigma_Max_lemma) + apply(simp add: h) + done + thus "rec_exec (rec_sigma ?notrq) + (xs @ [w, w]) = + Maxr (\args. 0 < rec_exec rf args) xs w" + apply(simp) + done + qed +qed + +text {* + @text "quo"} is the formal specification of division. + *} +fun quo :: "nat list \ nat" + where + "quo [x, y] = (let Rr = + (\ zs. ((zs ! (Suc 0) * zs ! (Suc (Suc 0)) + \ zs ! 0) \ zs ! Suc 0 \ (0::nat))) + in Maxr Rr [x, y] x)" + +declare quo.simps[simp del] + +text {* + The following lemmas shows more directly the menaing of @{text "quo"}: + *} +lemma [elim!]: "y > 0 \ quo [x, y] = x div y" +proof(simp add: quo.simps Maxr.simps, auto, + rule_tac Max_eqI, simp, auto) + fix xa ya + assume h: "y * ya \ x" "y > 0" + hence "(y * ya) div y \ x div y" + by(insert div_le_mono[of "y * ya" x y], simp) + from this and h show "ya \ x div y" by simp +next + fix xa + show "y * (x div y) \ x" + apply(subgoal_tac "y * (x div y) + x mod y = x") + apply(rule_tac k = "x mod y" in add_leD1, simp) + apply(simp) + done +qed + +lemma [intro]: "quo [x, 0] = 0" +by(simp add: quo.simps Maxr.simps) + +lemma quo_div: "quo [x, y] = x div y" +by(case_tac "y=0", auto) + +text {* + @{text "rec_noteq"} is the recursive function testing whether its + two arguments are not equal. + *} +definition rec_noteq:: "recf" + where + "rec_noteq = Cn (Suc (Suc 0)) rec_not [Cn (Suc (Suc 0)) + rec_eq [id (Suc (Suc 0)) (0), id (Suc (Suc 0)) + ((Suc 0))]]" + +text {* + The correctness of @{text "rec_noteq"}. + *} +lemma noteq_lemma: + "\ x y. rec_exec rec_noteq [x, y] = + (if x \ y then 1 else 0)" +by(simp add: rec_exec.simps rec_noteq_def) + +declare noteq_lemma[simp] + +text {* + @{text "rec_quo"} is the recursive function used to implement @{text "quo"} + *} +definition rec_quo :: "recf" + where + "rec_quo = (let rR = Cn (Suc (Suc (Suc 0))) rec_conj + [Cn (Suc (Suc (Suc 0))) rec_le + [Cn (Suc (Suc (Suc 0))) rec_mult + [id (Suc (Suc (Suc 0))) (Suc 0), + id (Suc (Suc (Suc 0))) ((Suc (Suc 0)))], + id (Suc (Suc (Suc 0))) (0)], + Cn (Suc (Suc (Suc 0))) rec_noteq + [id (Suc (Suc (Suc 0))) (Suc (0)), + Cn (Suc (Suc (Suc 0))) (constn 0) + [id (Suc (Suc (Suc 0))) (0)]]] + in Cn (Suc (Suc 0)) (rec_maxr rR)) [id (Suc (Suc 0)) + (0),id (Suc (Suc 0)) (Suc (0)), + id (Suc (Suc 0)) (0)]" + +lemma [intro]: "primerec rec_conj (Suc (Suc 0))" + apply(simp add: rec_conj_def) + apply(rule_tac prime_cn, auto)+ + apply(case_tac i, auto intro: prime_id) + done + +lemma [intro]: "primerec rec_noteq (Suc (Suc 0))" +apply(simp add: rec_noteq_def) +apply(rule_tac prime_cn, auto)+ +apply(case_tac i, auto intro: prime_id) +done + + +lemma quo_lemma1: "rec_exec rec_quo [x, y] = quo [x, y]" +proof(simp add: rec_exec.simps rec_quo_def) + let ?rR = "(Cn (Suc (Suc (Suc 0))) rec_conj + [Cn (Suc (Suc (Suc 0))) rec_le + [Cn (Suc (Suc (Suc 0))) rec_mult + [recf.id (Suc (Suc (Suc 0))) (Suc (0)), + recf.id (Suc (Suc (Suc 0))) (Suc (Suc (0)))], + recf.id (Suc (Suc (Suc 0))) (0)], + Cn (Suc (Suc (Suc 0))) rec_noteq + [recf.id (Suc (Suc (Suc 0))) + (Suc (0)), Cn (Suc (Suc (Suc 0))) (constn 0) + [recf.id (Suc (Suc (Suc 0))) (0)]]])" + have "rec_exec (rec_maxr ?rR) ([x, y]@ [ x]) = Maxr (\ args. 0 < rec_exec ?rR args) [x, y] x" + proof(rule_tac Maxr_lemma, simp) + show "primerec ?rR (Suc (Suc (Suc 0)))" + apply(auto) + apply(case_tac i, auto) + apply(case_tac [!] ia, auto) + apply(case_tac i, auto) + done + qed + hence g1: "rec_exec (rec_maxr ?rR) ([x, y, x]) = + Maxr (\ args. if rec_exec ?rR args = 0 then False + else True) [x, y] x" + by simp + have g2: "Maxr (\ args. if rec_exec ?rR args = 0 then False + else True) [x, y] x = quo [x, y]" + apply(simp add: rec_exec.simps) + apply(simp add: Maxr.simps quo.simps, auto) + done + from g1 and g2 show + "rec_exec (rec_maxr ?rR) ([x, y, x]) = quo [x, y]" + by simp +qed + +text {* + The correctness of @{text "quo"}. + *} +lemma quo_lemma2: "rec_exec rec_quo [x, y] = x div y" + using quo_lemma1[of x y] quo_div[of x y] + by simp + +text {* + @{text "rec_mod"} is the recursive function used to implement + the reminder function. + *} +definition rec_mod :: "recf" + where + "rec_mod = Cn (Suc (Suc 0)) rec_minus [id (Suc (Suc 0)) (0), + Cn (Suc (Suc 0)) rec_mult [rec_quo, id (Suc (Suc 0)) + (Suc (0))]]" + +text {* + The correctness of @{text "rec_mod"}: + *} +lemma mod_lemma: "\ x y. rec_exec rec_mod [x, y] = (x mod y)" +proof(simp add: rec_exec.simps rec_mod_def quo_lemma2) + fix x y + show "x - x div y * y = x mod (y::nat)" + using mod_div_equality2[of y x] + apply(subgoal_tac "y * (x div y) = (x div y ) * y", arith, simp) + done +qed + +text{* lemmas for embranch function*} +type_synonym ftype = "nat list \ nat" +type_synonym rtype = "nat list \ bool" + +text {* + The specifation of the mutli-way branching statement on + page 79 of Boolos's book. + *} +fun Embranch :: "(ftype * rtype) list \ nat list \ nat" + where + "Embranch [] xs = 0" | + "Embranch (gc # gcs) xs = ( + let (g, c) = gc in + if c xs then g xs else Embranch gcs xs)" + +fun rec_embranch' :: "(recf * recf) list \ nat \ recf" + where + "rec_embranch' [] vl = Cn vl z [id vl (vl - 1)]" | + "rec_embranch' ((rg, rc) # rgcs) vl = Cn vl rec_add + [Cn vl rec_mult [rg, rc], rec_embranch' rgcs vl]" + +text {* + @{text "rec_embrach"} is the recursive function used to implement + @{text "Embranch"}. + *} +fun rec_embranch :: "(recf * recf) list \ recf" + where + "rec_embranch ((rg, rc) # rgcs) = + (let vl = arity rg in + rec_embranch' ((rg, rc) # rgcs) vl)" + +declare Embranch.simps[simp del] rec_embranch.simps[simp del] + +lemma embranch_all0: + "\\ j < length rcs. rec_exec (rcs ! j) xs = 0; + length rgs = length rcs; + rcs \ []; + list_all (\ rf. primerec rf (length xs)) (rgs @ rcs)\ \ + rec_exec (rec_embranch (zip rgs rcs)) xs = 0" +proof(induct rcs arbitrary: rgs, simp, case_tac rgs, simp) + fix a rcs rgs aa list + assume ind: + "\rgs. \\j []; + list_all (\rf. primerec rf (length xs)) (rgs @ rcs)\ \ + rec_exec (rec_embranch (zip rgs rcs)) xs = 0" + and h: "\j []" + "list_all (\rf. primerec rf (length xs)) (rgs @ a # rcs)" + "rgs = aa # list" + have g: "rcs \ [] \ rec_exec (rec_embranch (zip list rcs)) xs = 0" + using h + by(rule_tac ind, auto) + show "rec_exec (rec_embranch (zip rgs (a # rcs))) xs = 0" + proof(case_tac "rcs = []", simp) + show "rec_exec (rec_embranch (zip rgs [a])) xs = 0" + using h + apply(simp add: rec_embranch.simps rec_exec.simps) + apply(erule_tac x = 0 in allE, simp) + done + next + assume "rcs \ []" + hence "rec_exec (rec_embranch (zip list rcs)) xs = 0" + using g by simp + thus "rec_exec (rec_embranch (zip rgs (a # rcs))) xs = 0" + using h + apply(simp add: rec_embranch.simps rec_exec.simps) + apply(case_tac rcs, + auto simp: rec_exec.simps rec_embranch.simps) + apply(case_tac list, + auto simp: rec_exec.simps rec_embranch.simps) + done + qed +qed + + +lemma embranch_exec_0: "\rec_exec aa xs = 0; zip rgs list \ []; + list_all (\ rf. primerec rf (length xs)) ([a, aa] @ rgs @ list)\ + \ rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs + = rec_exec (rec_embranch (zip rgs list)) xs" +apply(simp add: rec_exec.simps rec_embranch.simps) +apply(case_tac "zip rgs list", simp, case_tac ab, + simp add: rec_embranch.simps rec_exec.simps) +apply(subgoal_tac "arity a = length xs", auto) +apply(subgoal_tac "arity aaa = length xs", auto) +apply(case_tac rgs, simp, case_tac list, simp, simp) +done + +lemma zip_null_iff: "\length xs = k; length ys = k; zip xs ys = []\ \ xs = [] \ ys = []" +apply(case_tac xs, simp, simp) +apply(case_tac ys, simp, simp) +done + +lemma zip_null_gr: "\length xs = k; length ys = k; zip xs ys \ []\ \ 0 < k" +apply(case_tac xs, simp, simp) +done + +lemma Embranch_0: + "\length rgs = k; length rcs = k; k > 0; + \ j < k. rec_exec (rcs ! j) xs = 0\ \ + Embranch (zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) rcs)) xs = 0" +proof(induct rgs arbitrary: rcs k, simp, simp) + fix a rgs rcs k + assume ind: + "\rcs k. \length rgs = k; length rcs = k; 0 < k; \j + \ Embranch (zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) rcs)) xs = 0" + and h: "Suc (length rgs) = k" "length rcs = k" + "\jr args. 0 < rec_exec r args) rcs)) xs = 0" + apply(case_tac rcs, simp, case_tac "rgs = []", simp) + apply(simp add: Embranch.simps) + apply(erule_tac x = 0 in allE, simp) + apply(simp add: Embranch.simps) + apply(erule_tac x = 0 in all_dupE, simp) + apply(rule_tac ind, simp, simp, simp, auto) + apply(erule_tac x = "Suc j" in allE, simp) + done +qed + +text {* + The correctness of @{text "rec_embranch"}. + *} +lemma embranch_lemma: + assumes branch_num: + "length rgs = n" "length rcs = n" "n > 0" + and partition: + "(\ i < n. (rec_exec (rcs ! i) xs = 1 \ (\ j < n. j \ i \ + rec_exec (rcs ! j) xs = 0)))" + and prime_all: "list_all (\ rf. primerec rf (length xs)) (rgs @ rcs)" + shows "rec_exec (rec_embranch (zip rgs rcs)) xs = + Embranch (zip (map rec_exec rgs) + (map (\ r args. 0 < rec_exec r args) rcs)) xs" + using branch_num partition prime_all +proof(induct rgs arbitrary: rcs n, simp) + fix a rgs rcs n + assume ind: + "\rcs n. \length rgs = n; length rcs = n; 0 < n; + \i (\j i \ rec_exec (rcs ! j) xs = 0); + list_all (\rf. primerec rf (length xs)) (rgs @ rcs)\ + \ rec_exec (rec_embranch (zip rgs rcs)) xs = + Embranch (zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) rcs)) xs" + and h: "length (a # rgs) = n" "length (rcs::recf list) = n" "0 < n" + " \i + (\j i \ rec_exec (rcs ! j) xs = 0)" + "list_all (\rf. primerec rf (length xs)) ((a # rgs) @ rcs)" + from h show "rec_exec (rec_embranch (zip (a # rgs) rcs)) xs = + Embranch (zip (map rec_exec (a # rgs)) (map (\r args. + 0 < rec_exec r args) rcs)) xs" + apply(case_tac rcs, simp, simp) + apply(case_tac "rec_exec aa xs = 0") + apply(case_tac [!] "zip rgs list = []", simp) + apply(subgoal_tac "rgs = [] \ list = []", simp add: Embranch.simps rec_exec.simps rec_embranch.simps) + apply(rule_tac zip_null_iff, simp, simp, simp) +thm Embranch.simps + proof - + fix aa list + assume g: + "Suc (length rgs) = n" "Suc (length list) = n" + "\i + (\j i \ rec_exec ((aa # list) ! j) xs = 0)" + "primerec a (length xs) \ + list_all (\rf. primerec rf (length xs)) rgs \ + primerec aa (length xs) \ + list_all (\rf. primerec rf (length xs)) list" + "rec_exec aa xs = 0" "rcs = aa # list" "zip rgs list \ []" + have "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs + = rec_exec (rec_embranch (zip rgs list)) xs" + apply(rule embranch_exec_0, simp_all add: g) + done + from g and this show "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = + Embranch ((rec_exec a, \args. 0 < rec_exec aa args) # + zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) list)) xs" + apply(simp add: Embranch.simps) + apply(rule_tac n = "n - Suc 0" in ind) + apply(case_tac n, simp, simp) + apply(case_tac n, simp, simp) + apply(case_tac n, simp, simp add: zip_null_gr ) + apply(auto) + apply(case_tac i, simp, simp) + apply(rule_tac x = nat in exI, simp) + apply(rule_tac allI, erule_tac x = "Suc j" in allE, simp) + done + next + fix aa list + assume g: "Suc (length rgs) = n" "Suc (length list) = n" + "\i + (\j i \ rec_exec ((aa # list) ! j) xs = 0)" + "primerec a (length xs) \ list_all (\rf. primerec rf (length xs)) rgs \ + primerec aa (length xs) \ list_all (\rf. primerec rf (length xs)) list" + "rcs = aa # list" "rec_exec aa xs \ 0" "zip rgs list = []" + thus "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = + Embranch ((rec_exec a, \args. 0 < rec_exec aa args) # + zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) list)) xs" + apply(subgoal_tac "rgs = [] \ list = []", simp) + prefer 2 + apply(rule_tac zip_null_iff, simp, simp, simp) + apply(simp add: rec_exec.simps rec_embranch.simps Embranch.simps, auto) + done + next + fix aa list + assume g: "Suc (length rgs) = n" "Suc (length list) = n" + "\i + (\j i \ rec_exec ((aa # list) ! j) xs = 0)" + "primerec a (length xs) \ list_all (\rf. primerec rf (length xs)) rgs + \ primerec aa (length xs) \ list_all (\rf. primerec rf (length xs)) list" + "rcs = aa # list" "rec_exec aa xs \ 0" "zip rgs list \ []" + have "rec_exec aa xs = Suc 0" + using g + apply(case_tac "rec_exec aa xs", simp, auto) + done + moreover have "rec_exec (rec_embranch' (zip rgs list) (length xs)) xs = 0" + proof - + have "rec_embranch' (zip rgs list) (length xs) = rec_embranch (zip rgs list)" + using g + apply(case_tac "zip rgs list", simp, case_tac ab) + apply(simp add: rec_embranch.simps) + apply(subgoal_tac "arity aaa = length xs", simp, auto) + apply(case_tac rgs, simp, simp, case_tac list, simp, simp) + done + moreover have "rec_exec (rec_embranch (zip rgs list)) xs = 0" + proof(rule embranch_all0) + show " \j []" + using g + apply(case_tac list, simp, simp) + done + next + show "list_all (\rf. primerec rf (length xs)) (rgs @ list)" + using g + apply auto + done + qed + ultimately show "rec_exec (rec_embranch' (zip rgs list) (length xs)) xs = 0" + by simp + qed + moreover have + "Embranch (zip (map rec_exec rgs) + (map (\r args. 0 < rec_exec r args) list)) xs = 0" + using g + apply(rule_tac k = "length rgs" in Embranch_0) + apply(simp, case_tac n, simp, simp) + apply(case_tac rgs, simp, simp) + apply(auto) + apply(case_tac i, simp) + apply(erule_tac x = "Suc j" in allE, simp) + apply(simp) + apply(rule_tac x = 0 in allE, auto) + done + moreover have "arity a = length xs" + using g + apply(auto) + done + ultimately show "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = + Embranch ((rec_exec a, \args. 0 < rec_exec aa args) # + zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) list)) xs" + apply(simp add: rec_exec.simps rec_embranch.simps Embranch.simps) + done + qed +qed + +text{* + @{text "prime n"} means @{text "n"} is a prime number. +*} +fun Prime :: "nat \ bool" + where + "Prime x = (1 < x \ (\ u < x. (\ v < x. u * v \ x)))" + +declare Prime.simps [simp del] + +lemma primerec_all1: + "primerec (rec_all rt rf) n \ primerec rt n" + by (simp add: primerec_all) + +lemma primerec_all2: "primerec (rec_all rt rf) n \ + primerec rf (Suc n)" +by(insert primerec_all[of rt rf n], simp) + +text {* + @{text "rec_prime"} is the recursive function used to implement + @{text "Prime"}. + *} +definition rec_prime :: "recf" + where + "rec_prime = Cn (Suc 0) rec_conj + [Cn (Suc 0) rec_less [constn 1, id (Suc 0) (0)], + rec_all (Cn 1 rec_minus [id 1 0, constn 1]) + (rec_all (Cn 2 rec_minus [id 2 0, Cn 2 (constn 1) + [id 2 0]]) (Cn 3 rec_noteq + [Cn 3 rec_mult [id 3 1, id 3 2], id 3 0]))]" + +(* +lemma prime_lemma1: + "(rec_exec rec_prime [x] = Suc 0) \ + (rec_exec rec_prime [x] = 0)" +apply(auto simp: rec_exec.simps rec_prime_def) +done +*) +declare numeral_2_eq_2[simp del] numeral_3_eq_3[simp del] + +lemma exec_tmp: + "rec_exec (rec_all (Cn 2 rec_minus [recf.id 2 0, Cn 2 (constn (Suc 0)) [recf.id 2 0]]) + (Cn 3 rec_noteq [Cn 3 rec_mult [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0])) [x, k] = + ((if (\w\rec_exec (Cn 2 rec_minus [recf.id 2 0, Cn 2 (constn (Suc 0)) [recf.id 2 0]]) ([x, k]). + 0 < rec_exec (Cn 3 rec_noteq [Cn 3 rec_mult [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0]) + ([x, k] @ [w])) then 1 else 0))" +apply(rule_tac all_lemma) +apply(auto) +apply(case_tac [!] i, auto) +apply(case_tac ia, auto simp: numeral_3_eq_3 numeral_2_eq_2) +done + +text {* + The correctness of @{text "Prime"}. + *} +lemma prime_lemma: "rec_exec rec_prime [x] = (if Prime x then 1 else 0)" +proof(simp add: rec_exec.simps rec_prime_def) + let ?rt1 = "(Cn 2 rec_minus [recf.id 2 0, + Cn 2 (constn (Suc 0)) [recf.id 2 0]])" + let ?rf1 = "(Cn 3 rec_noteq [Cn 3 rec_mult + [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 (0)])" + let ?rt2 = "(Cn (Suc 0) rec_minus + [recf.id (Suc 0) 0, constn (Suc 0)])" + let ?rf2 = "rec_all ?rt1 ?rf1" + have h1: "rec_exec (rec_all ?rt2 ?rf2) ([x]) = + (if (\k\rec_exec ?rt2 ([x]). 0 < rec_exec ?rf2 ([x] @ [k])) then 1 else 0)" + proof(rule_tac all_lemma, simp_all) + show "primerec ?rf2 (Suc (Suc 0))" + apply(rule_tac primerec_all_iff) + apply(auto) + apply(case_tac [!] i, auto simp: numeral_2_eq_2) + apply(case_tac ia, auto simp: numeral_3_eq_3) + done + next + show "primerec (Cn (Suc 0) rec_minus + [recf.id (Suc 0) 0, constn (Suc 0)]) (Suc 0)" + apply(auto) + apply(case_tac i, auto) + done + qed + from h1 show + "(Suc 0 < x \ (rec_exec (rec_all ?rt2 ?rf2) [x] = 0 \ + \ Prime x) \ + (0 < rec_exec (rec_all ?rt2 ?rf2) [x] \ Prime x)) \ + (\ Suc 0 < x \ \ Prime x \ (rec_exec (rec_all ?rt2 ?rf2) [x] = 0 + \ \ Prime x))" + apply(auto simp:rec_exec.simps) + apply(simp add: exec_tmp rec_exec.simps) + proof - + assume "\k\x - Suc 0. (0::nat) < (if \w\x - Suc 0. + 0 < (if k * w \ x then 1 else (0 :: nat)) then 1 else 0)" "Suc 0 < x" + thus "Prime x" + apply(simp add: rec_exec.simps split: if_splits) + apply(simp add: Prime.simps, auto) + apply(erule_tac x = u in allE, auto) + apply(case_tac u, simp, case_tac nat, simp, simp) + apply(case_tac v, simp, case_tac nat, simp, simp) + done + next + assume "\ Suc 0 < x" "Prime x" + thus "False" + apply(simp add: Prime.simps) + done + next + fix k + assume "rec_exec (rec_all ?rt1 ?rf1) + [x, k] = 0" "k \ x - Suc 0" "Prime x" + thus "False" + apply(simp add: exec_tmp rec_exec.simps Prime.simps split: if_splits) + done + next + fix k + assume "rec_exec (rec_all ?rt1 ?rf1) + [x, k] = 0" "k \ x - Suc 0" "Prime x" + thus "False" + apply(simp add: exec_tmp rec_exec.simps Prime.simps split: if_splits) + done + qed +qed + +definition rec_dummyfac :: "recf" + where + "rec_dummyfac = Pr 1 (constn 1) + (Cn 3 rec_mult [id 3 2, Cn 3 s [id 3 1]])" + +text {* + The recursive function used to implment factorization. + *} +definition rec_fac :: "recf" + where + "rec_fac = Cn 1 rec_dummyfac [id 1 0, id 1 0]" + +text {* + Formal specification of factorization. + *} +fun fac :: "nat \ nat" ("_!" [100] 99) + where + "fac 0 = 1" | + "fac (Suc x) = (Suc x) * fac x" + +lemma [simp]: "rec_exec rec_dummyfac [0, 0] = Suc 0" +by(simp add: rec_dummyfac_def rec_exec.simps) + +lemma rec_cn_simp: "rec_exec (Cn n f gs) xs = + (let rgs = map (\ g. rec_exec g xs) gs in + rec_exec f rgs)" +by(simp add: rec_exec.simps) + +lemma rec_id_simp: "rec_exec (id m n) xs = xs ! n" + by(simp add: rec_exec.simps) + +lemma fac_dummy: "rec_exec rec_dummyfac [x, y] = y !" +apply(induct y) +apply(auto simp: rec_dummyfac_def rec_exec.simps) +done + +text {* + The correctness of @{text "rec_fac"}. + *} +lemma fac_lemma: "rec_exec rec_fac [x] = x!" +apply(simp add: rec_fac_def rec_exec.simps fac_dummy) +done + +declare fac.simps[simp del] + +text {* + @{text "Np x"} returns the first prime number after @{text "x"}. + *} +fun Np ::"nat \ nat" + where + "Np x = Min {y. y \ Suc (x!) \ x < y \ Prime y}" + +declare Np.simps[simp del] rec_Minr.simps[simp del] + +text {* + @{text "rec_np"} is the recursive function used to implement + @{text "Np"}. + *} +definition rec_np :: "recf" + where + "rec_np = (let Rr = Cn 2 rec_conj [Cn 2 rec_less [id 2 0, id 2 1], + Cn 2 rec_prime [id 2 1]] + in Cn 1 (rec_Minr Rr) [id 1 0, Cn 1 s [rec_fac]])" + +lemma [simp]: "n < Suc (n!)" +apply(induct n, simp) +apply(simp add: fac.simps) +apply(case_tac n, auto simp: fac.simps) +done + +lemma divsor_ex: +"\\ Prime x; x > Suc 0\ \ (\ u > Suc 0. (\ v > Suc 0. u * v = x))" + by(auto simp: Prime.simps) + +lemma divsor_prime_ex: "\\ Prime x; x > Suc 0\ \ + \ p. Prime p \ p dvd x" +apply(induct x rule: wf_induct[where r = "measure (\ y. y)"], simp) +apply(drule_tac divsor_ex, simp, auto) +apply(erule_tac x = u in allE, simp) +apply(case_tac "Prime u", simp) +apply(rule_tac x = u in exI, simp, auto) +done + +lemma [intro]: "0 < n!" +apply(induct n) +apply(auto simp: fac.simps) +done + +lemma fac_Suc: "Suc n! = (Suc n) * (n!)" by(simp add: fac.simps) + +lemma fac_dvd: "\0 < q; q \ n\ \ q dvd n!" +apply(induct n, simp) +apply(case_tac "q \ n", simp add: fac_Suc) +apply(subgoal_tac "q = Suc n", simp only: fac_Suc) +apply(rule_tac dvd_mult2, simp, simp) +done + +lemma fac_dvd2: "\Suc 0 < q; q dvd n!; q \ n\ \ \ q dvd Suc (n!)" +proof(auto simp: dvd_def) + fix k ka + assume h1: "Suc 0 < q" "q \ n" + and h2: "Suc (q * k) = q * ka" + have "k < ka" + proof - + have "q * k < q * ka" + using h2 by arith + thus "k < ka" + using h1 + by(auto) + qed + hence "\d. d > 0 \ ka = d + k" + by(rule_tac x = "ka - k" in exI, simp) + from this obtain d where "d > 0 \ ka = d + k" .. + from h2 and this and h1 show "False" + by(simp add: add_mult_distrib2) +qed + +lemma prime_ex: "\ p. n < p \ p \ Suc (n!) \ Prime p" +proof(cases "Prime (n! + 1)") + case True thus "?thesis" + by(rule_tac x = "Suc (n!)" in exI, simp) +next + assume h: "\ Prime (n! + 1)" + hence "\ p. Prime p \ p dvd (n! + 1)" + by(erule_tac divsor_prime_ex, auto) + from this obtain q where k: "Prime q \ q dvd (n! + 1)" .. + thus "?thesis" + proof(cases "q > n") + case True thus "?thesis" + using k + apply(rule_tac x = q in exI, auto) + apply(rule_tac dvd_imp_le, auto) + done + next + case False thus "?thesis" + proof - + assume g: "\ n < q" + have j: "q > Suc 0" + using k by(case_tac q, auto simp: Prime.simps) + hence "q dvd n!" + using g + apply(rule_tac fac_dvd, auto) + done + hence "\ q dvd Suc (n!)" + using g j + by(rule_tac fac_dvd2, auto) + thus "?thesis" + using k by simp + qed + qed +qed + +lemma Suc_Suc_induct[elim!]: "\i < Suc (Suc 0); + primerec (ys ! 0) n; primerec (ys ! 1) n\ \ primerec (ys ! i) n" +by(case_tac i, auto) + +lemma [intro]: "primerec rec_prime (Suc 0)" +apply(auto simp: rec_prime_def, auto) +apply(rule_tac primerec_all_iff, auto, auto) +apply(rule_tac primerec_all_iff, auto, auto simp: + numeral_2_eq_2 numeral_3_eq_3) +done + +text {* + The correctness of @{text "rec_np"}. + *} +lemma np_lemma: "rec_exec rec_np [x] = Np x" +proof(auto simp: rec_np_def rec_exec.simps Let_def fac_lemma) + let ?rr = "(Cn 2 rec_conj [Cn 2 rec_less [recf.id 2 0, + recf.id 2 (Suc 0)], Cn 2 rec_prime [recf.id 2 (Suc 0)]])" + let ?R = "\ zs. zs ! 0 < zs ! 1 \ Prime (zs ! 1)" + have g1: "rec_exec (rec_Minr ?rr) ([x] @ [Suc (x!)]) = + Minr (\ args. 0 < rec_exec ?rr args) [x] (Suc (x!))" + by(rule_tac Minr_lemma, auto simp: rec_exec.simps + prime_lemma, auto simp: numeral_2_eq_2 numeral_3_eq_3) + have g2: "Minr (\ args. 0 < rec_exec ?rr args) [x] (Suc (x!)) = Np x" + using prime_ex[of x] + apply(auto simp: Minr.simps Np.simps rec_exec.simps) + apply(erule_tac x = p in allE, simp add: prime_lemma) + apply(simp add: prime_lemma split: if_splits) + apply(subgoal_tac + "{uu. (Prime uu \ (x < uu \ uu \ Suc (x!)) \ x < uu) \ Prime uu} + = {y. y \ Suc (x!) \ x < y \ Prime y}", auto) + done + from g1 and g2 show "rec_exec (rec_Minr ?rr) ([x, Suc (x!)]) = Np x" + by simp +qed + +text {*lemmas for power*} +text {* + @{text "rec_power"} is the recursive function used to implement + power function. + *} +definition rec_power :: "recf" + where + "rec_power = Pr 1 (constn 1) (Cn 3 rec_mult [id 3 0, id 3 2])" + +text {* + The correctness of @{text "rec_power"}. + *} +lemma power_lemma: "rec_exec rec_power [x, y] = x^y" + by(induct y, auto simp: rec_exec.simps rec_power_def) + +text{* + @{text "Pi k"} returns the @{text "k"}-th prime number. + *} +fun Pi :: "nat \ nat" + where + "Pi 0 = 2" | + "Pi (Suc x) = Np (Pi x)" + +definition rec_dummy_pi :: "recf" + where + "rec_dummy_pi = Pr 1 (constn 2) (Cn 3 rec_np [id 3 2])" + +text {* + @{text "rec_pi"} is the recursive function used to implement + @{text "Pi"}. + *} +definition rec_pi :: "recf" + where + "rec_pi = Cn 1 rec_dummy_pi [id 1 0, id 1 0]" + +lemma pi_dummy_lemma: "rec_exec rec_dummy_pi [x, y] = Pi y" +apply(induct y) +by(auto simp: rec_exec.simps rec_dummy_pi_def Pi.simps np_lemma) + +text {* + The correctness of @{text "rec_pi"}. + *} +lemma pi_lemma: "rec_exec rec_pi [x] = Pi x" +apply(simp add: rec_pi_def rec_exec.simps pi_dummy_lemma) +done + +text{*follows: lemmas for lo*} +fun loR :: "nat list \ bool" + where + "loR [x, y, u] = (x mod (y^u) = 0)" + +declare loR.simps[simp del] + +text {* + @{text "Lo"} specifies the @{text "lo"} function given on page 79 of + Boolos's book. It is one of the two notions of integeral logarithmatic + operation on that page. The other is @{text "lg"}. + *} +fun lo :: " nat \ nat \ nat" + where + "lo x y = (if x > 1 \ y > 1 \ {u. loR [x, y, u]} \ {} then Max {u. loR [x, y, u]} + else 0)" + +declare lo.simps[simp del] + +lemma [elim]: "primerec rf n \ n > 0" +apply(induct rule: primerec.induct, auto) +done + +lemma primerec_sigma[intro!]: + "\n > Suc 0; primerec rf n\ \ + primerec (rec_sigma rf) n" +apply(simp add: rec_sigma.simps) +apply(auto, auto simp: nth_append) +done + +lemma [intro!]: "\primerec rf n; n > 0\ \ primerec (rec_maxr rf) n" +apply(simp add: rec_maxr.simps) +apply(rule_tac prime_cn, auto) +apply(rule_tac primerec_all_iff, auto, auto simp: nth_append) +done + +lemma Suc_Suc_Suc_induct[elim!]: + "\i < Suc (Suc (Suc (0::nat))); primerec (ys ! 0) n; + primerec (ys ! 1) n; + primerec (ys ! 2) n\ \ primerec (ys ! i) n" +apply(case_tac i, auto, case_tac nat, simp, simp add: numeral_2_eq_2) +done + +lemma [intro]: "primerec rec_quo (Suc (Suc 0))" +apply(simp add: rec_quo_def) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}] 1*}, auto+)+ +done + +lemma [intro]: "primerec rec_mod (Suc (Suc 0))" +apply(simp add: rec_mod_def) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}] 1*}, auto+)+ +done + +lemma [intro]: "primerec rec_power (Suc (Suc 0))" +apply(simp add: rec_power_def numeral_2_eq_2 numeral_3_eq_3) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ +done + +text {* + @{text "rec_lo"} is the recursive function used to implement @{text "Lo"}. +*} +definition rec_lo :: "recf" + where + "rec_lo = (let rR = Cn 3 rec_eq [Cn 3 rec_mod [id 3 0, + Cn 3 rec_power [id 3 1, id 3 2]], + Cn 3 (constn 0) [id 3 1]] in + let rb = Cn 2 (rec_maxr rR) [id 2 0, id 2 1, id 2 0] in + let rcond = Cn 2 rec_conj [Cn 2 rec_less [Cn 2 (constn 1) + [id 2 0], id 2 0], + Cn 2 rec_less [Cn 2 (constn 1) + [id 2 0], id 2 1]] in + let rcond2 = Cn 2 rec_minus + [Cn 2 (constn 1) [id 2 0], rcond] + in Cn 2 rec_add [Cn 2 rec_mult [rb, rcond], + Cn 2 rec_mult [Cn 2 (constn 0) [id 2 0], rcond2]])" + +lemma rec_lo_Maxr_lor: + "\Suc 0 < x; Suc 0 < y\ \ + rec_exec rec_lo [x, y] = Maxr loR [x, y] x" +proof(auto simp: rec_exec.simps rec_lo_def Let_def + numeral_2_eq_2 numeral_3_eq_3) + let ?rR = "(Cn (Suc (Suc (Suc 0))) rec_eq + [Cn (Suc (Suc (Suc 0))) rec_mod [recf.id (Suc (Suc (Suc 0))) 0, + Cn (Suc (Suc (Suc 0))) rec_power [recf.id (Suc (Suc (Suc 0))) + (Suc 0), recf.id (Suc (Suc (Suc 0))) (Suc (Suc 0))]], + Cn (Suc (Suc (Suc 0))) (constn 0) [recf.id (Suc (Suc (Suc 0))) (Suc 0)]])" + have h: "rec_exec (rec_maxr ?rR) ([x, y] @ [x]) = + Maxr (\ args. 0 < rec_exec ?rR args) [x, y] x" + by(rule_tac Maxr_lemma, auto simp: rec_exec.simps + mod_lemma power_lemma, auto simp: numeral_2_eq_2 numeral_3_eq_3) + have "Maxr loR [x, y] x = Maxr (\ args. 0 < rec_exec ?rR args) [x, y] x" + apply(simp add: rec_exec.simps mod_lemma power_lemma) + apply(simp add: Maxr.simps loR.simps) + done + from h and this show "rec_exec (rec_maxr ?rR) [x, y, x] = + Maxr loR [x, y] x" + apply(simp) + done +qed + +lemma [simp]: "Max {ya. ya = 0 \ loR [0, y, ya]} = 0" +apply(rule_tac Max_eqI, auto simp: loR.simps) +done + +lemma [simp]: "Suc 0 < y \ Suc (Suc 0) < y * y" +apply(induct y, simp) +apply(case_tac y, simp, simp) +done + +lemma less_mult: "\x > 0; y > Suc 0\ \ x < y * x" +apply(case_tac y, simp, simp) +done + +lemma x_less_exp: "\y > Suc 0\ \ x < y^x" +apply(induct x, simp, simp) +apply(case_tac x, simp, auto) +apply(rule_tac y = "y* y^nat" in le_less_trans, simp) +apply(rule_tac less_mult, auto) +done + +lemma le_mult: "y \ (0::nat) \ x \ x * y" + by(induct y, simp, simp) + +lemma uplimit_loR: "\Suc 0 < x; Suc 0 < y; loR [x, y, xa]\ \ + xa \ x" +apply(simp add: loR.simps) +apply(rule_tac classical, auto) +apply(subgoal_tac "xa < y^xa") +apply(subgoal_tac "y^xa \ y^xa * q", simp) +apply(rule_tac le_mult, case_tac q, simp, simp) +apply(rule_tac x_less_exp, simp) +done + +lemma [simp]: "\xa \ x; loR [x, y, xa]; Suc 0 < x; Suc 0 < y\ \ + {u. loR [x, y, u]} = {ya. ya \ x \ loR [x, y, ya]}" +apply(rule_tac Collect_cong, auto) +apply(erule_tac uplimit_loR, simp, simp) +done + +lemma Maxr_lo: "\Suc 0 < x; Suc 0 < y\ \ + Maxr loR [x, y] x = lo x y" +apply(simp add: Maxr.simps lo.simps, auto) +apply(erule_tac x = xa in allE, simp, simp add: uplimit_loR) +done + +lemma lo_lemma': "\Suc 0 < x; Suc 0 < y\ \ + rec_exec rec_lo [x, y] = lo x y" +by(simp add: Maxr_lo rec_lo_Maxr_lor) + +lemma lo_lemma'': "\\ Suc 0 < x\ \ rec_exec rec_lo [x, y] = lo x y" +apply(case_tac x, auto simp: rec_exec.simps rec_lo_def + Let_def lo.simps) +done + +lemma lo_lemma''': "\\ Suc 0 < y\ \ rec_exec rec_lo [x, y] = lo x y" +apply(case_tac y, auto simp: rec_exec.simps rec_lo_def + Let_def lo.simps) +done + +text {* + The correctness of @{text "rec_lo"}: +*} +lemma lo_lemma: "rec_exec rec_lo [x, y] = lo x y" +apply(case_tac "Suc 0 < x \ Suc 0 < y") +apply(auto simp: lo_lemma' lo_lemma'' lo_lemma''') +done + +fun lgR :: "nat list \ bool" + where + "lgR [x, y, u] = (y^u \ x)" + +text {* + @{text "lg"} specifies the @{text "lg"} function given on page 79 of + Boolos's book. It is one of the two notions of integeral logarithmatic + operation on that page. The other is @{text "lo"}. + *} +fun lg :: "nat \ nat \ nat" + where + "lg x y = (if x > 1 \ y > 1 \ {u. lgR [x, y, u]} \ {} then + Max {u. lgR [x, y, u]} + else 0)" + +declare lg.simps[simp del] lgR.simps[simp del] + +text {* + @{text "rec_lg"} is the recursive function used to implement @{text "lg"}. + *} +definition rec_lg :: "recf" + where + "rec_lg = (let rec_lgR = Cn 3 rec_le + [Cn 3 rec_power [id 3 1, id 3 2], id 3 0] in + let conR1 = Cn 2 rec_conj [Cn 2 rec_less + [Cn 2 (constn 1) [id 2 0], id 2 0], + Cn 2 rec_less [Cn 2 (constn 1) + [id 2 0], id 2 1]] in + let conR2 = Cn 2 rec_not [conR1] in + Cn 2 rec_add [Cn 2 rec_mult + [conR1, Cn 2 (rec_maxr rec_lgR) + [id 2 0, id 2 1, id 2 0]], + Cn 2 rec_mult [conR2, Cn 2 (constn 0) + [id 2 0]]])" + +lemma lg_maxr: "\Suc 0 < x; Suc 0 < y\ \ + rec_exec rec_lg [x, y] = Maxr lgR [x, y] x" +proof(simp add: rec_exec.simps rec_lg_def Let_def) + assume h: "Suc 0 < x" "Suc 0 < y" + let ?rR = "(Cn 3 rec_le [Cn 3 rec_power + [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0])" + have "rec_exec (rec_maxr ?rR) ([x, y] @ [x]) + = Maxr ((\ args. 0 < rec_exec ?rR args)) [x, y] x" + proof(rule Maxr_lemma) + show "primerec (Cn 3 rec_le [Cn 3 rec_power + [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0]) (Suc (length [x, y]))" + apply(auto simp: numeral_3_eq_3)+ + done + qed + moreover have "Maxr lgR [x, y] x = Maxr ((\ args. 0 < rec_exec ?rR args)) [x, y] x" + apply(simp add: rec_exec.simps power_lemma) + apply(simp add: Maxr.simps lgR.simps) + done + ultimately show "rec_exec (rec_maxr ?rR) [x, y, x] = Maxr lgR [x, y] x" + by simp +qed + +lemma [simp]: "\Suc 0 < y; lgR [x, y, xa]\ \ xa \ x" +apply(simp add: lgR.simps) +apply(subgoal_tac "y^xa > xa", simp) +apply(erule x_less_exp) +done + +lemma [simp]: "\Suc 0 < x; Suc 0 < y; lgR [x, y, xa]\ \ + {u. lgR [x, y, u]} = {ya. ya \ x \ lgR [x, y, ya]}" +apply(rule_tac Collect_cong, auto) +done + +lemma maxr_lg: "\Suc 0 < x; Suc 0 < y\ \ Maxr lgR [x, y] x = lg x y" +apply(simp add: lg.simps Maxr.simps, auto) +apply(erule_tac x = xa in allE, simp) +done + +lemma lg_lemma': "\Suc 0 < x; Suc 0 < y\ \ rec_exec rec_lg [x, y] = lg x y" +apply(simp add: maxr_lg lg_maxr) +done + +lemma lg_lemma'': "\ Suc 0 < x \ rec_exec rec_lg [x, y] = lg x y" +apply(simp add: rec_exec.simps rec_lg_def Let_def lg.simps) +done + +lemma lg_lemma''': "\ Suc 0 < y \ rec_exec rec_lg [x, y] = lg x y" +apply(simp add: rec_exec.simps rec_lg_def Let_def lg.simps) +done + +text {* + The correctness of @{text "rec_lg"}. + *} +lemma lg_lemma: "rec_exec rec_lg [x, y] = lg x y" +apply(case_tac "Suc 0 < x \ Suc 0 < y", auto simp: + lg_lemma' lg_lemma'' lg_lemma''') +done + +text {* + @{text "Entry sr i"} returns the @{text "i"}-th entry of a list of natural + numbers encoded by number @{text "sr"} using Godel's coding. + *} +fun Entry :: "nat \ nat \ nat" + where + "Entry sr i = lo sr (Pi (Suc i))" + +text {* + @{text "rec_entry"} is the recursive function used to implement + @{text "Entry"}. + *} +definition rec_entry:: "recf" + where + "rec_entry = Cn 2 rec_lo [id 2 0, Cn 2 rec_pi [Cn 2 s [id 2 1]]]" + +declare Pi.simps[simp del] + +text {* + The correctness of @{text "rec_entry"}. + *} +lemma entry_lemma: "rec_exec rec_entry [str, i] = Entry str i" + by(simp add: rec_entry_def rec_exec.simps lo_lemma pi_lemma) + +section {* The construction of @{text "F"} *} + +text {* + Using the auxilliary functions obtained in last section, + we are going to contruct the function @{text "F"}, + which is an interpreter of Turing Machines. + *} + +fun listsum2 :: "nat list \ nat \ nat" + where + "listsum2 xs 0 = 0" +| "listsum2 xs (Suc n) = listsum2 xs n + xs ! n" + +fun rec_listsum2 :: "nat \ nat \ recf" + where + "rec_listsum2 vl 0 = Cn vl z [id vl 0]" +| "rec_listsum2 vl (Suc n) = Cn vl rec_add + [rec_listsum2 vl n, id vl (n)]" + +declare listsum2.simps[simp del] rec_listsum2.simps[simp del] + +lemma listsum2_lemma: "\length xs = vl; n \ vl\ \ + rec_exec (rec_listsum2 vl n) xs = listsum2 xs n" +apply(induct n, simp_all) +apply(simp_all add: rec_exec.simps rec_listsum2.simps listsum2.simps) +done + +fun strt' :: "nat list \ nat \ nat" + where + "strt' xs 0 = 0" +| "strt' xs (Suc n) = (let dbound = listsum2 xs n + n in + strt' xs n + (2^(xs ! n + dbound) - 2^dbound))" + +fun rec_strt' :: "nat \ nat \ recf" + where + "rec_strt' vl 0 = Cn vl z [id vl 0]" +| "rec_strt' vl (Suc n) = (let rec_dbound = + Cn vl rec_add [rec_listsum2 vl n, Cn vl (constn n) [id vl 0]] + in Cn vl rec_add [rec_strt' vl n, Cn vl rec_minus + [Cn vl rec_power [Cn vl (constn 2) [id vl 0], Cn vl rec_add + [id vl (n), rec_dbound]], + Cn vl rec_power [Cn vl (constn 2) [id vl 0], rec_dbound]]])" + +declare strt'.simps[simp del] rec_strt'.simps[simp del] + +lemma strt'_lemma: "\length xs = vl; n \ vl\ \ + rec_exec (rec_strt' vl n) xs = strt' xs n" +apply(induct n) +apply(simp_all add: rec_exec.simps rec_strt'.simps strt'.simps + Let_def power_lemma listsum2_lemma) +done + +text {* + @{text "strt"} corresponds to the @{text "strt"} function on page 90 of B book, but + this definition generalises the original one to deal with multiple input arguments. + *} +fun strt :: "nat list \ nat" + where + "strt xs = (let ys = map Suc xs in + strt' ys (length ys))" + +fun rec_map :: "recf \ nat \ recf list" + where + "rec_map rf vl = map (\ i. Cn vl rf [id vl (i)]) [0.. recf" + where + "rec_strt vl = Cn vl (rec_strt' vl vl) (rec_map s vl)" + +lemma map_s_lemma: "length xs = vl \ + map ((\a. rec_exec a xs) \ (\i. Cn vl s [recf.id vl i])) + [0.. ys y. xs = ys @ [y]", auto) +proof - + fix ys y + assume ind: "\xs. length xs = length (ys::nat list) \ + map ((\a. rec_exec a xs) \ (\i. Cn (length ys) s + [recf.id (length ys) (i)])) [0..a. rec_exec a (ys @ [y])) \ (\i. Cn (Suc (length ys)) s + [recf.id (Suc (length ys)) (i)])) [0..a. rec_exec a ys) \ (\i. Cn (length ys) s + [recf.id (length ys) (i)])) [0..a. rec_exec a (ys @ [y])) \ (\i. Cn (Suc (length ys)) s + [recf.id (Suc (length ys)) (i)])) [0..a. rec_exec a ys) \ (\i. Cn (length ys) s + [recf.id (length ys) (i)])) [0..ys y. xs = ys @ [y]" + apply(rule_tac x = "butlast xs" in exI, rule_tac x = "last xs" in exI) + apply(subgoal_tac "xs \ []", auto) + done +qed + +text {* + The correctness of @{text "rec_strt"}. + *} +lemma strt_lemma: "length xs = vl \ + rec_exec (rec_strt vl) xs = strt xs" +apply(simp add: strt.simps rec_exec.simps strt'_lemma) +apply(subgoal_tac "(map ((\a. rec_exec a xs) \ (\i. Cn vl s [recf.id vl (i)])) [0.. nat" + where + "scan r = r mod 2" + +text {* + @{text "rec_scan"} is the implemention of @{text "scan"}. + *} +definition rec_scan :: "recf" + where "rec_scan = Cn 1 rec_mod [id 1 0, constn 2]" + +text {* + The correctness of @{text "scan"}. + *} +lemma scan_lemma: "rec_exec rec_scan [r] = r mod 2" + by(simp add: rec_exec.simps rec_scan_def mod_lemma) + +fun newleft0 :: "nat list \ nat" + where + "newleft0 [p, r] = p" + +definition rec_newleft0 :: "recf" + where + "rec_newleft0 = id 2 0" + +fun newrgt0 :: "nat list \ nat" + where + "newrgt0 [p, r] = r - scan r" + +definition rec_newrgt0 :: "recf" + where + "rec_newrgt0 = Cn 2 rec_minus [id 2 1, Cn 2 rec_scan [id 2 1]]" + +(*newleft1, newrgt1: left rgt number after execute on step*) +fun newleft1 :: "nat list \ nat" + where + "newleft1 [p, r] = p" + +definition rec_newleft1 :: "recf" + where + "rec_newleft1 = id 2 0" + +fun newrgt1 :: "nat list \ nat" + where + "newrgt1 [p, r] = r + 1 - scan r" + +definition rec_newrgt1 :: "recf" + where + "rec_newrgt1 = + Cn 2 rec_minus [Cn 2 rec_add [id 2 1, Cn 2 (constn 1) [id 2 0]], + Cn 2 rec_scan [id 2 1]]" + +fun newleft2 :: "nat list \ nat" + where + "newleft2 [p, r] = p div 2" + +definition rec_newleft2 :: "recf" + where + "rec_newleft2 = Cn 2 rec_quo [id 2 0, Cn 2 (constn 2) [id 2 0]]" + +fun newrgt2 :: "nat list \ nat" + where + "newrgt2 [p, r] = 2 * r + p mod 2" + +definition rec_newrgt2 :: "recf" + where + "rec_newrgt2 = + Cn 2 rec_add [Cn 2 rec_mult [Cn 2 (constn 2) [id 2 0], id 2 1], + Cn 2 rec_mod [id 2 0, Cn 2 (constn 2) [id 2 0]]]" + +fun newleft3 :: "nat list \ nat" + where + "newleft3 [p, r] = 2 * p + r mod 2" + +definition rec_newleft3 :: "recf" + where + "rec_newleft3 = + Cn 2 rec_add [Cn 2 rec_mult [Cn 2 (constn 2) [id 2 0], id 2 0], + Cn 2 rec_mod [id 2 1, Cn 2 (constn 2) [id 2 0]]]" + +fun newrgt3 :: "nat list \ nat" + where + "newrgt3 [p, r] = r div 2" + +definition rec_newrgt3 :: "recf" + where + "rec_newrgt3 = Cn 2 rec_quo [id 2 1, Cn 2 (constn 2) [id 2 0]]" + +text {* + The @{text "new_left"} function on page 91 of B book. + *} +fun newleft :: "nat \ nat \ nat \ nat" + where + "newleft p r a = (if a = 0 \ a = 1 then newleft0 [p, r] + else if a = 2 then newleft2 [p, r] + else if a = 3 then newleft3 [p, r] + else p)" + +text {* + @{text "rec_newleft"} is the recursive function used to + implement @{text "newleft"}. + *} +definition rec_newleft :: "recf" + where + "rec_newleft = + (let g0 = + Cn 3 rec_newleft0 [id 3 0, id 3 1] in + let g1 = Cn 3 rec_newleft2 [id 3 0, id 3 1] in + let g2 = Cn 3 rec_newleft3 [id 3 0, id 3 1] in + let g3 = id 3 0 in + let r0 = Cn 3 rec_disj + [Cn 3 rec_eq [id 3 2, Cn 3 (constn 0) [id 3 0]], + Cn 3 rec_eq [id 3 2, Cn 3 (constn 1) [id 3 0]]] in + let r1 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 2) [id 3 0]] in + let r2 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 3) [id 3 0]] in + let r3 = Cn 3 rec_less [Cn 3 (constn 3) [id 3 0], id 3 2] in + let gs = [g0, g1, g2, g3] in + let rs = [r0, r1, r2, r3] in + rec_embranch (zip gs rs))" + +declare newleft.simps[simp del] + + +lemma Suc_Suc_Suc_Suc_induct: + "\i < Suc (Suc (Suc (Suc 0))); i = 0 \ P i; + i = 1 \ P i; i =2 \ P i; + i =3 \ P i\ \ P i" +apply(case_tac i, simp, case_tac nat, simp, + case_tac nata, simp, case_tac natb, simp, simp) +done + +declare quo_lemma2[simp] mod_lemma[simp] + +text {* + The correctness of @{text "rec_newleft"}. + *} +lemma newleft_lemma: + "rec_exec rec_newleft [p, r, a] = newleft p r a" +proof(simp only: rec_newleft_def Let_def) + let ?rgs = "[Cn 3 rec_newleft0 [recf.id 3 0, recf.id 3 1], Cn 3 rec_newleft2 + [recf.id 3 0, recf.id 3 1], Cn 3 rec_newleft3 [recf.id 3 0, recf.id 3 1], recf.id 3 0]" + let ?rrs = + "[Cn 3 rec_disj [Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 0) + [recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 1) [recf.id 3 0]]], + Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 2) [recf.id 3 0]], + Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 3) [recf.id 3 0]], + Cn 3 rec_less [Cn 3 (constn 3) [recf.id 3 0], recf.id 3 2]]" + thm embranch_lemma + have k1: "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] + = Embranch (zip (map rec_exec ?rgs) (map (\r args. 0 < rec_exec r args) ?rrs)) [p, r, a]" + apply(rule_tac embranch_lemma ) + apply(auto simp: numeral_3_eq_3 numeral_2_eq_2 rec_newleft0_def + rec_newleft1_def rec_newleft2_def rec_newleft3_def)+ + apply(case_tac "a = 0 \ a = 1", rule_tac x = 0 in exI) + prefer 2 + apply(case_tac "a = 2", rule_tac x = "Suc 0" in exI) + prefer 2 + apply(case_tac "a = 3", rule_tac x = "2" in exI) + prefer 2 + apply(case_tac "a > 3", rule_tac x = "3" in exI, auto) + apply(auto simp: rec_exec.simps) + apply(erule_tac [!] Suc_Suc_Suc_Suc_induct, auto simp: rec_exec.simps) + done(* + have "Embranch (zip (map rec_exec ?rgs) (map (\r args. 0 < rec_exec r args) ?rrs)) [p, r, a] + = Embranch (zip ?gs ?rs) [p, r, a]" + apply(simp add)*) + have k2: "Embranch (zip (map rec_exec ?rgs) (map (\r args. 0 < rec_exec r args) ?rrs)) [p, r, a] = newleft p r a" + apply(simp add: Embranch.simps) + apply(simp add: rec_exec.simps) + apply(auto simp: newleft.simps rec_newleft0_def rec_exec.simps + rec_newleft1_def rec_newleft2_def rec_newleft3_def) + done + from k1 and k2 show + "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] = newleft p r a" + by simp +qed + +text {* + The @{text "newrght"} function is one similar to @{text "newleft"}, but used to + compute the right number. + *} +fun newrght :: "nat \ nat \ nat \ nat" + where + "newrght p r a = (if a = 0 then newrgt0 [p, r] + else if a = 1 then newrgt1 [p, r] + else if a = 2 then newrgt2 [p, r] + else if a = 3 then newrgt3 [p, r] + else r)" + +text {* + @{text "rec_newrght"} is the recursive function used to implement + @{text "newrgth"}. + *} +definition rec_newrght :: "recf" + where + "rec_newrght = + (let g0 = Cn 3 rec_newrgt0 [id 3 0, id 3 1] in + let g1 = Cn 3 rec_newrgt1 [id 3 0, id 3 1] in + let g2 = Cn 3 rec_newrgt2 [id 3 0, id 3 1] in + let g3 = Cn 3 rec_newrgt3 [id 3 0, id 3 1] in + let g4 = id 3 1 in + let r0 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 0) [id 3 0]] in + let r1 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 1) [id 3 0]] in + let r2 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 2) [id 3 0]] in + let r3 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 3) [id 3 0]] in + let r4 = Cn 3 rec_less [Cn 3 (constn 3) [id 3 0], id 3 2] in + let gs = [g0, g1, g2, g3, g4] in + let rs = [r0, r1, r2, r3, r4] in + rec_embranch (zip gs rs))" +declare newrght.simps[simp del] + +lemma numeral_4_eq_4: "4 = Suc 3" +by auto + +lemma Suc_5_induct: + "\i < Suc (Suc (Suc (Suc (Suc 0)))); i = 0 \ P 0; + i = 1 \ P 1; i = 2 \ P 2; i = 3 \ P 3; i = 4 \ P 4\ \ P i" +apply(case_tac i, auto) +apply(case_tac nat, auto) +apply(case_tac nata, auto simp: numeral_2_eq_2) +apply(case_tac nat, auto simp: numeral_3_eq_3 numeral_4_eq_4) +done + +lemma [intro]: "primerec rec_scan (Suc 0)" +apply(auto simp: rec_scan_def, auto) +done + +text {* + The correctness of @{text "rec_newrght"}. + *} +lemma newrght_lemma: "rec_exec rec_newrght [p, r, a] = newrght p r a" +proof(simp only: rec_newrght_def Let_def) + let ?gs' = "[newrgt0, newrgt1, newrgt2, newrgt3, \ zs. zs ! 1]" + let ?r0 = "\ zs. zs ! 2 = 0" + let ?r1 = "\ zs. zs ! 2 = 1" + let ?r2 = "\ zs. zs ! 2 = 2" + let ?r3 = "\ zs. zs ! 2 = 3" + let ?r4 = "\ zs. zs ! 2 > 3" + let ?gs = "map (\ g. (\ zs. g [zs ! 0, zs ! 1])) ?gs'" + let ?rs = "[?r0, ?r1, ?r2, ?r3, ?r4]" + let ?rgs = + "[Cn 3 rec_newrgt0 [recf.id 3 0, recf.id 3 1], + Cn 3 rec_newrgt1 [recf.id 3 0, recf.id 3 1], + Cn 3 rec_newrgt2 [recf.id 3 0, recf.id 3 1], + Cn 3 rec_newrgt3 [recf.id 3 0, recf.id 3 1], recf.id 3 1]" + let ?rrs = + "[Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 0) [recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, + Cn 3 (constn 1) [recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 2) [recf.id 3 0]], + Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 3) [recf.id 3 0]], + Cn 3 rec_less [Cn 3 (constn 3) [recf.id 3 0], recf.id 3 2]]" + + have k1: "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] + = Embranch (zip (map rec_exec ?rgs) (map (\r args. 0 < rec_exec r args) ?rrs)) [p, r, a]" + apply(rule_tac embranch_lemma) + apply(auto simp: numeral_3_eq_3 numeral_2_eq_2 rec_newrgt0_def + rec_newrgt1_def rec_newrgt2_def rec_newrgt3_def)+ + apply(case_tac "a = 0", rule_tac x = 0 in exI) + prefer 2 + apply(case_tac "a = 1", rule_tac x = "Suc 0" in exI) + prefer 2 + apply(case_tac "a = 2", rule_tac x = "2" in exI) + prefer 2 + apply(case_tac "a = 3", rule_tac x = "3" in exI) + prefer 2 + apply(case_tac "a > 3", rule_tac x = "4" in exI, auto simp: rec_exec.simps) + apply(erule_tac [!] Suc_5_induct, auto simp: rec_exec.simps) + done + have k2: "Embranch (zip (map rec_exec ?rgs) + (map (\r args. 0 < rec_exec r args) ?rrs)) [p, r, a] = newrght p r a" + apply(auto simp:Embranch.simps rec_exec.simps) + apply(auto simp: newrght.simps rec_newrgt3_def rec_newrgt2_def + rec_newrgt1_def rec_newrgt0_def rec_exec.simps + scan_lemma) + done + from k1 and k2 show + "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] = + newrght p r a" by simp +qed + +declare Entry.simps[simp del] + +text {* + The @{text "actn"} function given on page 92 of B book, which is used to + fetch Turing Machine intructions. + In @{text "actn m q r"}, @{text "m"} is the Godel coding of a Turing Machine, + @{text "q"} is the current state of Turing Machine, @{text "r"} is the + right number of Turing Machine tape. + *} +fun actn :: "nat \ nat \ nat \ nat" + where + "actn m q r = (if q \ 0 then Entry m (4*(q - 1) + 2 * scan r) + else 4)" + +text {* + @{text "rec_actn"} is the recursive function used to implement @{text "actn"} + *} +definition rec_actn :: "recf" + where + "rec_actn = + Cn 3 rec_add [Cn 3 rec_mult + [Cn 3 rec_entry [id 3 0, Cn 3 rec_add [Cn 3 rec_mult + [Cn 3 (constn 4) [id 3 0], + Cn 3 rec_minus [id 3 1, Cn 3 (constn 1) [id 3 0]]], + Cn 3 rec_mult [Cn 3 (constn 2) [id 3 0], + Cn 3 rec_scan [id 3 2]]]], + Cn 3 rec_noteq [id 3 1, Cn 3 (constn 0) [id 3 0]]], + Cn 3 rec_mult [Cn 3 (constn 4) [id 3 0], + Cn 3 rec_eq [id 3 1, Cn 3 (constn 0) [id 3 0]]]] " + +text {* + The correctness of @{text "actn"}. + *} +lemma actn_lemma: "rec_exec rec_actn [m, q, r] = actn m q r" + by(auto simp: rec_actn_def rec_exec.simps entry_lemma scan_lemma) + +(* Stop point *) + +fun newstat :: "nat \ nat \ nat \ nat" + where + "newstat m q r = (if q \ 0 then Entry m (4*(q - 1) + 2*scan r + 1) + else 0)" + +definition rec_newstat :: "recf" + where + "rec_newstat = Cn 3 rec_add + [Cn 3 rec_mult [Cn 3 rec_entry [id 3 0, + Cn 3 rec_add [Cn 3 rec_mult [Cn 3 (constn 4) [id 3 0], + Cn 3 rec_minus [id 3 1, Cn 3 (constn 1) [id 3 0]]], + Cn 3 rec_add [Cn 3 rec_mult [Cn 3 (constn 2) [id 3 0], + Cn 3 rec_scan [id 3 2]], Cn 3 (constn 1) [id 3 0]]]], + Cn 3 rec_noteq [id 3 1, Cn 3 (constn 0) [id 3 0]]], + Cn 3 rec_mult [Cn 3 (constn 0) [id 3 0], + Cn 3 rec_eq [id 3 1, Cn 3 (constn 0) [id 3 0]]]] " + +lemma newstat_lemma: "rec_exec rec_newstat [m, q, r] = newstat m q r" +by(auto simp: rec_exec.simps entry_lemma scan_lemma rec_newstat_def) + +declare newstat.simps[simp del] actn.simps[simp del] + +text{*code the configuration*} + +fun trpl :: "nat \ nat \ nat \ nat" + where + "trpl p q r = (Pi 0)^p * (Pi 1)^q * (Pi 2)^r" + +definition rec_trpl :: "recf" + where + "rec_trpl = Cn 3 rec_mult [Cn 3 rec_mult + [Cn 3 rec_power [Cn 3 (constn (Pi 0)) [id 3 0], id 3 0], + Cn 3 rec_power [Cn 3 (constn (Pi 1)) [id 3 0], id 3 1]], + Cn 3 rec_power [Cn 3 (constn (Pi 2)) [id 3 0], id 3 2]]" +declare trpl.simps[simp del] +lemma trpl_lemma: "rec_exec rec_trpl [p, q, r] = trpl p q r" +by(auto simp: rec_trpl_def rec_exec.simps power_lemma trpl.simps) + +text{*left, stat, rght: decode func*} +fun left :: "nat \ nat" + where + "left c = lo c (Pi 0)" + +fun stat :: "nat \ nat" + where + "stat c = lo c (Pi 1)" + +fun rght :: "nat \ nat" + where + "rght c = lo c (Pi 2)" + +thm Prime.simps + +fun inpt :: "nat \ nat list \ nat" + where + "inpt m xs = trpl 0 1 (strt xs)" + +fun newconf :: "nat \ nat \ nat" + where + "newconf m c = trpl (newleft (left c) (rght c) + (actn m (stat c) (rght c))) + (newstat m (stat c) (rght c)) + (newrght (left c) (rght c) + (actn m (stat c) (rght c)))" + +declare left.simps[simp del] stat.simps[simp del] rght.simps[simp del] + inpt.simps[simp del] newconf.simps[simp del] + +definition rec_left :: "recf" + where + "rec_left = Cn 1 rec_lo [id 1 0, constn (Pi 0)]" + +definition rec_right :: "recf" + where + "rec_right = Cn 1 rec_lo [id 1 0, constn (Pi 2)]" + +definition rec_stat :: "recf" + where + "rec_stat = Cn 1 rec_lo [id 1 0, constn (Pi 1)]" + +definition rec_inpt :: "nat \ recf" + where + "rec_inpt vl = Cn vl rec_trpl + [Cn vl (constn 0) [id vl 0], + Cn vl (constn 1) [id vl 0], + Cn vl (rec_strt (vl - 1)) + (map (\ i. id vl (i)) [1..a. rec_exec a (m # xs)) \ + (\i. recf.id (Suc (length xs)) (i))) + [Suc 0.. i. xs ! (i - 1)) [Suc 0.. map (\ i. xs ! (i - 1)) + [Suc 0.. ys y. xs = ys @ [y]", auto) +proof - + fix ys y + assume ind: + "\xs. length (ys::nat list) = length (xs::nat list) \ + map (\i. xs ! (i - Suc 0)) [Suc 0.. length (ys::nat list)" + have "map (\i. ys ! (i - Suc 0)) [Suc 0..i. (ys @ [y]) ! (i - Suc 0)) [Suc 0..i. ys ! (i - Suc 0)) [Suc 0..i. (ys @ [y]) ! (i - Suc 0)) + [Suc 0..ys y. xs = ys @ [y]" + apply(rule_tac x = "butlast xs" in exI, + rule_tac x = "last xs" in exI) + apply(case_tac "xs \ []", auto) + done +qed + +lemma [elim]: + "Suc 0 \ length xs \ + (map ((\a. rec_exec a (m # xs)) \ + (\i. recf.id (Suc (length xs)) (i))) + [Suc 0..Suc (length xs) = vl\ \ + rec_exec (rec_inpt vl) (m # xs) = inpt m xs" +apply(auto simp: rec_exec.simps rec_inpt_def + trpl_lemma inpt.simps strt_lemma) +apply(subgoal_tac + "(map ((\a. rec_exec a (m # xs)) \ + (\i. recf.id (Suc (length xs)) (i))) + [Suc 0.. nat \ nat \ nat" + where + "conf m r 0 = trpl 0 (Suc 0) r" +| "conf m r (Suc t) = newconf m (conf m r t)" + +declare conf.simps[simp del] + +text {* + @{text "conf"} is implemented by the following recursive function @{text "rec_conf"}. + *} +definition rec_conf :: "recf" + where + "rec_conf = Pr 2 (Cn 2 rec_trpl [Cn 2 (constn 0) [id 2 0], Cn 2 (constn (Suc 0)) [id 2 0], id 2 1]) + (Cn 4 rec_newconf [id 4 0, id 4 3])" + +lemma conf_step: + "rec_exec rec_conf [m, r, Suc t] = + rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" +proof - + have "rec_exec rec_conf ([m, r] @ [Suc t]) = + rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" + by(simp only: rec_conf_def rec_pr_Suc_simp_rewrite, + simp add: rec_exec.simps) + thus "rec_exec rec_conf [m, r, Suc t] = + rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" + by simp +qed + +text {* + The correctness of @{text "rec_conf"}. + *} +lemma conf_lemma: + "rec_exec rec_conf [m, r, t] = conf m r t" +apply(induct t) +apply(simp add: rec_conf_def rec_exec.simps conf.simps inpt_lemma trpl_lemma) +apply(simp add: conf_step conf.simps) +done + +text {* + @{text "NSTD c"} returns true if the configureation coded by @{text "c"} is no a stardard + final configuration. + *} +fun NSTD :: "nat \ bool" + where + "NSTD c = (stat c \ 0 \ left c \ 0 \ + rght c \ 2^(lg (rght c + 1) 2) - 1 \ rght c = 0)" + +text {* + @{text "rec_NSTD"} is the recursive function implementing @{text "NSTD"}. + *} +definition rec_NSTD :: "recf" + where + "rec_NSTD = + Cn 1 rec_disj [ + Cn 1 rec_disj [ + Cn 1 rec_disj + [Cn 1 rec_noteq [rec_stat, constn 0], + Cn 1 rec_noteq [rec_left, constn 0]] , + Cn 1 rec_noteq [rec_right, + Cn 1 rec_minus [Cn 1 rec_power + [constn 2, Cn 1 rec_lg + [Cn 1 rec_add + [rec_right, constn 1], + constn 2]], constn 1]]], + Cn 1 rec_eq [rec_right, constn 0]]" + +lemma NSTD_lemma1: "rec_exec rec_NSTD [c] = Suc 0 \ + rec_exec rec_NSTD [c] = 0" +by(simp add: rec_exec.simps rec_NSTD_def) + +declare NSTD.simps[simp del] +lemma NSTD_lemma2': "(rec_exec rec_NSTD [c] = Suc 0) \ NSTD c" +apply(simp add: rec_exec.simps rec_NSTD_def stat_lemma left_lemma + lg_lemma right_lemma power_lemma NSTD.simps eq_lemma) +apply(auto) +apply(case_tac "0 < left c", simp, simp) +done + +lemma NSTD_lemma2'': + "NSTD c \ (rec_exec rec_NSTD [c] = Suc 0)" +apply(simp add: rec_exec.simps rec_NSTD_def stat_lemma + left_lemma lg_lemma right_lemma power_lemma NSTD.simps) +apply(auto split: if_splits) +done + +text {* + The correctness of @{text "NSTD"}. + *} +lemma NSTD_lemma2: "(rec_exec rec_NSTD [c] = Suc 0) = NSTD c" +using NSTD_lemma1 +apply(auto intro: NSTD_lemma2' NSTD_lemma2'') +done + +fun nstd :: "nat \ nat" + where + "nstd c = (if NSTD c then 1 else 0)" + +lemma nstd_lemma: "rec_exec rec_NSTD [c] = nstd c" +using NSTD_lemma1 +apply(simp add: NSTD_lemma2, auto) +done + +text {* GGGGGGGGGGGGGGGGGGGGGGG *} + +text{* + @{text "nonstep m r t"} means afer @{text "t"} steps of execution, the TM coded by @{text "m"} + is not at a stardard final configuration. + *} +fun nonstop :: "nat \ nat \ nat \ nat" + where + "nonstop m r t = nstd (conf m r t)" + +text {* + @{text "rec_nonstop"} is the recursive function implementing @{text "nonstop"}. + *} +definition rec_nonstop :: "recf" + where + "rec_nonstop = Cn 3 rec_NSTD [rec_conf]" + +text {* + The correctness of @{text "rec_nonstop"}. + *} +lemma nonstop_lemma: + "rec_exec rec_nonstop [m, r, t] = nonstop m r t" +apply(simp add: rec_exec.simps rec_nonstop_def nstd_lemma conf_lemma) +done + +text{* + @{text "rec_halt"} is the recursive function calculating the steps a TM needs to execute before + to reach a stardard final configuration. This recursive function is the only one + using @{text "Mn"} combinator. So it is the only non-primitive recursive function + needs to be used in the construction of the universal function @{text "F"}. + *} + +definition rec_halt :: "recf" + where + "rec_halt = Mn (Suc (Suc 0)) (rec_nonstop)" + +declare nonstop.simps[simp del] + +(* when mn, use rec_calc_rel instead of rec_exec*) + +lemma primerec_not0: "primerec f n \ n > 0" +by(induct f n rule: primerec.induct, auto) + +lemma [elim]: "primerec f 0 \ RR" +apply(drule_tac primerec_not0, simp) +done + +lemma [simp]: "length xs = Suc n \ length (butlast xs) = n" +apply(subgoal_tac "\ y ys. xs = ys @ [y]", auto) +apply(rule_tac x = "last xs" in exI) +apply(rule_tac x = "butlast xs" in exI) +apply(case_tac "xs = []", auto) +done + +text {* + The lemma relates the interpreter of primitive fucntions with + the calculation relation of general recursive functions. + *} +lemma prime_rel_exec_eq: "primerec r (length xs) + \ rec_calc_rel r xs rs = (rec_exec r xs = rs)" +proof(induct r xs arbitrary: rs rule: rec_exec.induct, simp_all) + fix xs rs + assume "primerec z (length (xs::nat list))" + hence "length xs = Suc 0" by(erule_tac prime_z_reverse, simp) + thus "rec_calc_rel z xs rs = (rec_exec z xs = rs)" + apply(case_tac xs, simp, auto) + apply(erule_tac calc_z_reverse, simp add: rec_exec.simps) + apply(simp add: rec_exec.simps, rule_tac calc_z) + done +next + fix xs rs + assume "primerec s (length (xs::nat list))" + hence "length xs = Suc 0" .. + thus "rec_calc_rel s xs rs = (rec_exec s xs = rs)" + by(case_tac xs, auto simp: rec_exec.simps intro: calc_s + elim: calc_s_reverse) +next + fix m n xs rs + assume "primerec (recf.id m n) (length (xs::nat list))" + thus + "rec_calc_rel (recf.id m n) xs rs = + (rec_exec (recf.id m n) xs = rs)" + apply(erule_tac prime_id_reverse) + apply(simp add: rec_exec.simps, auto) + apply(erule_tac calc_id_reverse, simp) + apply(rule_tac calc_id, auto) + done +next + fix n f gs xs rs + assume ind1: + "\x rs. \x \ set gs; primerec x (length xs)\ \ + rec_calc_rel x xs rs = (rec_exec x xs = rs)" + and ind2: + "\x rs. \x = map (\a. rec_exec a xs) gs; + primerec f (length gs)\ \ + rec_calc_rel f (map (\a. rec_exec a xs) gs) rs = + (rec_exec f (map (\a. rec_exec a xs) gs) = rs)" + and h: "primerec (Cn n f gs) (length xs)" + show "rec_calc_rel (Cn n f gs) xs rs = + (rec_exec (Cn n f gs) xs = rs)" + proof(auto simp: rec_exec.simps, erule_tac calc_cn_reverse, auto) + fix ys + assume g1:"\ka. rec_exec a xs) gs) rs = + (rec_exec f (map (\a. rec_exec a xs) gs) = rs)" + apply(rule_tac ind2, auto) + using h + apply(erule_tac prime_cn_reverse, simp) + done + moreover have "ys = (map (\a. rec_exec a xs) gs)" + proof(rule_tac nth_equalityI, auto simp: g2) + fix i + assume "i < length gs" thus "ys ! i = rec_exec (gs!i) xs" + using ind1[of "gs ! i" "ys ! i"] g1 h + apply(erule_tac prime_cn_reverse, simp) + done + qed + ultimately show "rec_exec f (map (\a. rec_exec a xs) gs) = rs" + using g3 + by(simp) + next + from h show + "rec_calc_rel (Cn n f gs) xs + (rec_exec f (map (\a. rec_exec a xs) gs))" + apply(rule_tac rs = "(map (\a. rec_exec a xs) gs)" in calc_cn, + auto) + apply(erule_tac [!] prime_cn_reverse, auto) + proof - + fix k + assume "k < length gs" "primerec f (length gs)" + "\iia. rec_exec a xs) gs) + (rec_exec f (map (\a. rec_exec a xs) gs))" + using ind2[of "(map (\a. rec_exec a xs) gs)" + "(rec_exec f (map (\a. rec_exec a xs) gs))"] + by simp + qed + qed +next + fix n f g xs rs + assume ind1: + "\rs. \last xs = 0; primerec f (length xs - Suc 0)\ + \ rec_calc_rel f (butlast xs) rs = + (rec_exec f (butlast xs) = rs)" + and ind2 : + "\rs. \0 < last xs; + primerec (Pr n f g) (Suc (length xs - Suc 0))\ \ + rec_calc_rel (Pr n f g) (butlast xs @ [last xs - Suc 0]) rs + = (rec_exec (Pr n f g) (butlast xs @ [last xs - Suc 0]) = rs)" + and ind3: + "\rs. \0 < last xs; primerec g (Suc (Suc (length xs - Suc 0)))\ + \ rec_calc_rel g (butlast xs @ + [last xs - Suc 0, rec_exec (Pr n f g) + (butlast xs @ [last xs - Suc 0])]) rs = + (rec_exec g (butlast xs @ [last xs - Suc 0, + rec_exec (Pr n f g) + (butlast xs @ [last xs - Suc 0])]) = rs)" + and h: "primerec (Pr n f g) (length (xs::nat list))" + show "rec_calc_rel (Pr n f g) xs rs = (rec_exec (Pr n f g) xs = rs)" + proof(auto) + assume "rec_calc_rel (Pr n f g) xs rs" + thus "rec_exec (Pr n f g) xs = rs" + proof(erule_tac calc_pr_reverse) + fix l + assume g: "xs = l @ [0]" + "rec_calc_rel f l rs" + "n = length l" + thus "rec_exec (Pr n f g) xs = rs" + using ind1[of rs] h + apply(simp add: rec_exec.simps, + erule_tac prime_pr_reverse, simp) + done + next + fix l y ry + assume d:"xs = l @ [Suc y]" + "rec_calc_rel (Pr (length l) f g) (l @ [y]) ry" + "n = length l" + "rec_calc_rel g (l @ [y, ry]) rs" + moreover hence "primerec g (Suc (Suc n))" using h + proof(erule_tac prime_pr_reverse) + assume "primerec g (Suc (Suc n))" "length xs = Suc n" + thus "?thesis" by simp + qed + ultimately show "rec_exec (Pr n f g) xs = rs" + apply(simp) + using ind3[of rs] + apply(simp add: rec_pr_Suc_simp_rewrite) + using ind2[of ry] h + apply(simp) + done + qed + next + show "rec_calc_rel (Pr n f g) xs (rec_exec (Pr n f g) xs)" + proof - + have "rec_calc_rel (Pr n f g) (butlast xs @ [last xs]) + (rec_exec (Pr n f g) (butlast xs @ [last xs]))" + using h + apply(erule_tac prime_pr_reverse, simp) + apply(case_tac "last xs", simp) + apply(rule_tac calc_pr_zero, simp) + using ind1[of "rec_exec (Pr n f g) (butlast xs @ [0])"] + apply(simp add: rec_exec.simps, simp, simp, simp) + thm calc_pr_ind + apply(rule_tac rk = "rec_exec (Pr n f g) + (butlast xs@[last xs - Suc 0])" in calc_pr_ind) + using ind2[of "rec_exec (Pr n f g) + (butlast xs @ [last xs - Suc 0])"] h + apply(simp, simp, simp) + proof - + fix nat + assume "length xs = Suc n" + "primerec g (Suc (Suc n))" + "last xs = Suc nat" + thus + "rec_calc_rel g (butlast xs @ [nat, rec_exec (Pr n f g) + (butlast xs @ [nat])]) (rec_exec (Pr n f g) (butlast xs @ [Suc nat]))" + using ind3[of "rec_exec (Pr n f g) + (butlast xs @ [Suc nat])"] + apply(simp add: rec_exec.simps) + done + qed + thus "rec_calc_rel (Pr n f g) xs (rec_exec (Pr n f g) xs)" + using h + apply(erule_tac prime_pr_reverse, simp) + apply(subgoal_tac "butlast xs @ [last xs] = xs", simp) + apply(case_tac xs, simp, simp) + done + qed + qed +next + fix n f xs rs + assume "primerec (Mn n f) (length (xs::nat list))" + thus "rec_calc_rel (Mn n f) xs rs = (rec_exec (Mn n f) xs = rs)" + by(erule_tac prime_mn_reverse) +qed + +declare numeral_2_eq_2[simp] numeral_3_eq_3[simp] + +lemma [intro]: "primerec rec_right (Suc 0)" +apply(simp add: rec_right_def rec_lo_def Let_def) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ +done + +lemma [simp]: +"rec_calc_rel rec_right [r] rs = (rec_exec rec_right [r] = rs)" +apply(rule_tac prime_rel_exec_eq, auto) +done + +lemma [intro]: "primerec rec_pi (Suc 0)" +apply(simp add: rec_pi_def rec_dummy_pi_def + rec_np_def rec_fac_def rec_prime_def + rec_Minr.simps Let_def get_fstn_args.simps + arity.simps + rec_all.simps rec_sigma.simps rec_accum.simps) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ +apply(simp add: rec_dummyfac_def) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ +done + +lemma [intro]: "primerec rec_trpl (Suc (Suc (Suc 0)))" +apply(simp add: rec_trpl_def) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ +done + +lemma [intro!]: "\0 < vl; n \ vl\ \ primerec (rec_listsum2 vl n) vl" +apply(induct n) +apply(simp_all add: rec_strt'.simps Let_def rec_listsum2.simps) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ +done + +lemma [elim]: "\0 < vl; n \ vl\ \ primerec (rec_strt' vl n) vl" +apply(induct n) +apply(simp_all add: rec_strt'.simps Let_def) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+) +done + +lemma [elim]: "vl > 0 \ primerec (rec_strt vl) vl" +apply(simp add: rec_strt.simps rec_strt'.simps) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ +done + +lemma [elim]: + "i < vl \ primerec ((map (\i. recf.id (Suc vl) (i)) + [Suc 0.. primerec (rec_inpt (Suc vl)) (Suc vl)" +apply(simp add: rec_inpt_def) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ +done + +lemma [intro]: "primerec rec_conf (Suc (Suc (Suc 0)))" +apply(simp add: rec_conf_def) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ +apply(auto simp: numeral_4_eq_4) +done + +lemma [simp]: + "rec_calc_rel rec_conf [m, r, t] rs = + (rec_exec rec_conf [m, r, t] = rs)" +apply(rule_tac prime_rel_exec_eq, auto) +done + +lemma [intro]: "primerec rec_lg (Suc (Suc 0))" +apply(simp add: rec_lg_def Let_def) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ +done + +lemma [intro]: "primerec rec_nonstop (Suc (Suc (Suc 0)))" +apply(simp add: rec_nonstop_def rec_NSTD_def rec_stat_def + rec_lo_def Let_def rec_left_def rec_right_def rec_newconf_def + rec_newstat_def) +apply(tactic {* resolve_tac [@{thm prime_cn}, + @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ +done + +lemma nonstop_eq[simp]: + "rec_calc_rel rec_nonstop [m, r, t] rs = + (rec_exec rec_nonstop [m, r, t] = rs)" +apply(rule prime_rel_exec_eq, auto) +done + +lemma halt_lemma': + "rec_calc_rel rec_halt [m, r] t = + (rec_calc_rel rec_nonstop [m, r, t] 0 \ + (\ t'< t. + (\ y. rec_calc_rel rec_nonstop [m, r, t'] y \ + y \ 0)))" +apply(auto simp: rec_halt_def) +apply(erule calc_mn_reverse, simp) +apply(erule_tac calc_mn_reverse) +apply(erule_tac x = t' in allE, simp) +apply(rule_tac calc_mn, simp_all) +done + +text {* + The following lemma gives the correctness of @{text "rec_halt"}. + It says: if @{text "rec_halt"} calculates that the TM coded by @{text "m"} + will reach a standard final configuration after @{text "t"} steps of execution, then it is indeed so. + *} +lemma halt_lemma: + "rec_calc_rel (rec_halt) [m, r] t = + (rec_exec rec_nonstop [m, r, t] = 0 \ + (\ t'< t. (\ y. rec_exec rec_nonstop [m, r, t'] = y + \ y \ 0)))" +using halt_lemma'[of m r t] +by simp + +text {*F: universal machine*} + +text {* + @{text "valu r"} extracts computing result out of the right number @{text "r"}. + *} +fun valu :: "nat \ nat" + where + "valu r = (lg (r + 1) 2) - 1" + +text {* + @{text "rec_valu"} is the recursive function implementing @{text "valu"}. +*} +definition rec_valu :: "recf" + where + "rec_valu = Cn 1 rec_minus [Cn 1 rec_lg [s, constn 2], constn 1]" + +text {* + The correctness of @{text "rec_valu"}. +*} +lemma value_lemma: "rec_exec rec_valu [r] = valu r" +apply(simp add: rec_exec.simps rec_valu_def lg_lemma) +done + +lemma [intro]: "primerec rec_valu (Suc 0)" +apply(simp add: rec_valu_def) +apply(rule_tac k = "Suc (Suc 0)" in prime_cn) +apply(auto simp: prime_s) +proof - + show "primerec rec_lg (Suc (Suc 0))" by auto +next + show "Suc (Suc 0) = Suc (Suc 0)" by simp +next + show "primerec (constn (Suc (Suc 0))) (Suc 0)" by auto +qed + +lemma [simp]: "rec_calc_rel rec_valu [r] rs = + (rec_exec rec_valu [r] = rs)" +apply(rule_tac prime_rel_exec_eq, auto) +done + +declare valu.simps[simp del] + +text {* + The definition of the universal function @{text "rec_F"}. + *} +definition rec_F :: "recf" + where + "rec_F = Cn (Suc (Suc 0)) rec_valu [Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) + rec_conf ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])]]" + +lemma get_fstn_args_nth: + "k < n \ (get_fstn_args m n ! k) = id m (k)" +apply(induct n, simp) +apply(case_tac "k = n", simp_all add: get_fstn_args.simps + nth_append) +done + +lemma [simp]: + "\ys \ []; k < length ys\ \ + (get_fstn_args (length ys) (length ys) ! k) = + id (length ys) (k)" +by(erule_tac get_fstn_args_nth) + +lemma calc_rel_get_pren: + "\ys \ []; k < length ys\ \ + rec_calc_rel (get_fstn_args (length ys) (length ys) ! k) ys + (ys ! k)" +apply(simp) +apply(rule_tac calc_id, auto) +done + +lemma [elim]: + "\xs \ []; k < Suc (length xs)\ \ + rec_calc_rel (get_fstn_args (Suc (length xs)) + (Suc (length xs)) ! k) (m # xs) ((m # xs) ! k)" +using calc_rel_get_pren[of "m#xs" k] +apply(simp) +done + +text {* + The correctness of @{text "rec_F"}, halt case. + *} +lemma F_lemma: + "rec_calc_rel rec_halt [m, r] t \ + rec_calc_rel rec_F [m, r] (valu (rght (conf m r t)))" +apply(simp add: rec_F_def) +apply(rule_tac rs = "[rght (conf m r t)]" in calc_cn, + auto simp: value_lemma) +apply(rule_tac rs = "[conf m r t]" in calc_cn, + auto simp: right_lemma) +apply(rule_tac rs = "[m, r, t]" in calc_cn, auto) +apply(subgoal_tac " k = 0 \ k = Suc 0 \ k = Suc (Suc 0)", + auto simp:nth_append) +apply(rule_tac [1-2] calc_id, simp_all add: conf_lemma) +done + + +text {* + The correctness of @{text "rec_F"}, nonhalt case. + *} +lemma F_lemma2: + "\ t. \ rec_calc_rel rec_halt [m, r] t \ + \ rs. \ rec_calc_rel rec_F [m, r] rs" +apply(auto simp: rec_F_def) +apply(erule_tac calc_cn_reverse, simp (no_asm_use))+ +proof - + fix rs rsa rsb rsc + assume h: + "\t. \ rec_calc_rel rec_halt [m, r] t" + "length rsa = Suc 0" + "rec_calc_rel rec_valu rsa rs" + "length rsb = Suc 0" + "rec_calc_rel rec_right rsb (rsa ! 0)" + "length rsc = (Suc (Suc (Suc 0)))" + "rec_calc_rel rec_conf rsc (rsb ! 0)" + and g: "\k nat \ nat" + where + "bl2nat [] n = 0" +| "bl2nat (Bk#bl) n = bl2nat bl (Suc n)" +| "bl2nat (Oc#bl) n = 2^n + bl2nat bl (Suc n)" + +fun bl2wc :: "block list \ nat" + where + "bl2wc xs = bl2nat xs 0" + +fun trpl_code :: "t_conf \ nat" + where + "trpl_code (st, l, r) = trpl (bl2wc l) st (bl2wc r)" + +declare bl2nat.simps[simp del] bl2wc.simps[simp del] + trpl_code.simps[simp del] + +fun action_map :: "taction \ nat" + where + "action_map W0 = 0" +| "action_map W1 = 1" +| "action_map L = 2" +| "action_map R = 3" +| "action_map Nop = 4" + +fun action_map_iff :: "nat \ taction" + where + "action_map_iff (0::nat) = W0" +| "action_map_iff (Suc 0) = W1" +| "action_map_iff (Suc (Suc 0)) = L" +| "action_map_iff (Suc (Suc (Suc 0))) = R" +| "action_map_iff n = Nop" + +fun block_map :: "block \ nat" + where + "block_map Bk = 0" +| "block_map Oc = 1" + +fun godel_code' :: "nat list \ nat \ nat" + where + "godel_code' [] n = 1" +| "godel_code' (x#xs) n = (Pi n)^x * godel_code' xs (Suc n) " + +fun godel_code :: "nat list \ nat" + where + "godel_code xs = (let lh = length xs in + 2^lh * (godel_code' xs (Suc 0)))" + +fun modify_tprog :: "tprog \ nat list" + where + "modify_tprog [] = []" +| "modify_tprog ((ac, ns)#nl) = action_map ac # ns # modify_tprog nl" + +text {* + @{text "code tp"} gives the Godel coding of TM program @{text "tp"}. + *} +fun code :: "tprog \ nat" + where + "code tp = (let nl = modify_tprog tp in + godel_code nl)" + +section {* Relating interperter functions to the execution of TMs *} + +lemma [simp]: "bl2wc [] = 0" by(simp add: bl2wc.simps bl2nat.simps) +term trpl + +lemma [simp]: "\fetch tp 0 b = (nact, ns)\ \ action_map nact = 4" +apply(simp add: fetch.simps) +done + +thm entry_lemma +lemma Pi_gr_1[simp]: "Pi n > Suc 0" +proof(induct n, auto simp: Pi.simps Np.simps) + fix n + let ?setx = "{y. y \ Suc (Pi n!) \ Pi n < y \ Prime y}" + have "finite ?setx" by auto + moreover have "?setx \ {}" + using prime_ex[of "Pi n"] + apply(auto) + done + ultimately show "Suc 0 < Min ?setx" + apply(simp add: Min_gr_iff) + apply(auto simp: Prime.simps) + done +qed + +lemma Pi_not_0[simp]: "Pi n > 0" +using Pi_gr_1[of n] +by arith + +declare godel_code.simps[simp del] + +lemma [simp]: "0 < godel_code' nl n" +apply(induct nl arbitrary: n) +apply(auto simp: godel_code'.simps) +done + +lemma godel_code_great: "godel_code nl > 0" +apply(simp add: godel_code.simps) +done + +lemma godel_code_eq_1: "(godel_code nl = 1) = (nl = [])" +apply(auto simp: godel_code.simps) +done + +lemma [elim]: + "\i < length nl; \ Suc 0 < godel_code nl\ \ nl ! i = 0" +using godel_code_great[of nl] godel_code_eq_1[of nl] +apply(simp) +done + +term set_of +lemma prime_coprime: "\Prime x; Prime y; x\y\ \ coprime x y" +proof(simp only: Prime.simps coprime_nat, auto simp: dvd_def, + rule_tac classical, simp) + fix d k ka + assume case_ka: "\uv d * ka" + and case_k: "\uv d * k" + and h: "(0::nat) < d" "d \ Suc 0" "Suc 0 < d * ka" + "ka \ k" "Suc 0 < d * k" + from h have "k > Suc 0 \ ka >Suc 0" + apply(auto) + apply(case_tac ka, simp, simp) + apply(case_tac k, simp, simp) + done + from this show "False" + proof(erule_tac disjE) + assume "(Suc 0::nat) < k" + hence "k < d*k \ d < d*k" + using h + by(auto) + thus "?thesis" + using case_k + apply(erule_tac x = d in allE) + apply(simp) + apply(erule_tac x = k in allE) + apply(simp) + done + next + assume "(Suc 0::nat) < ka" + hence "ka < d * ka \ d < d*ka" + using h by auto + thus "?thesis" + using case_ka + apply(erule_tac x = d in allE) + apply(simp) + apply(erule_tac x = ka in allE) + apply(simp) + done + qed +qed + +lemma Pi_inc: "Pi (Suc i) > Pi i" +proof(simp add: Pi.simps Np.simps) + let ?setx = "{y. y \ Suc (Pi i!) \ Pi i < y \ Prime y}" + have "finite ?setx" by simp + moreover have "?setx \ {}" + using prime_ex[of "Pi i"] + apply(auto) + done + ultimately show "Pi i < Min ?setx" + apply(simp add: Min_gr_iff) + done +qed + +lemma Pi_inc_gr: "i < j \ Pi i < Pi j" +proof(induct j, simp) + fix j + assume ind: "i < j \ Pi i < Pi j" + and h: "i < Suc j" + from h show "Pi i < Pi (Suc j)" + proof(cases "i < j") + case True thus "?thesis" + proof - + assume "i < j" + hence "Pi i < Pi j" by(erule_tac ind) + moreover have "Pi j < Pi (Suc j)" + apply(simp add: Pi_inc) + done + ultimately show "?thesis" + by simp + qed + next + assume "i < Suc j" "\ i < j" + hence "i = j" + by arith + thus "Pi i < Pi (Suc j)" + apply(simp add: Pi_inc) + done + qed +qed + +lemma Pi_notEq: "i \ j \ Pi i \ Pi j" +apply(case_tac "i < j") +using Pi_inc_gr[of i j] +apply(simp) +using Pi_inc_gr[of j i] +apply(simp) +done + +lemma [intro]: "Prime (Suc (Suc 0))" +apply(auto simp: Prime.simps) +apply(case_tac u, simp, case_tac nat, simp, simp) +done + +lemma Prime_Pi[intro]: "Prime (Pi n)" +proof(induct n, auto simp: Pi.simps Np.simps) + fix n + let ?setx = "{y. y \ Suc (Pi n!) \ Pi n < y \ Prime y}" + show "Prime (Min ?setx)" + proof - + have "finite ?setx" by simp + moreover have "?setx \ {}" + using prime_ex[of "Pi n"] + apply(simp) + done + ultimately show "?thesis" + apply(drule_tac Min_in, simp, simp) + done + qed +qed + +lemma Pi_coprime: "i \ j \ coprime (Pi i) (Pi j)" +using Prime_Pi[of i] +using Prime_Pi[of j] +apply(rule_tac prime_coprime, simp_all add: Pi_notEq) +done + +lemma Pi_power_coprime: "i \ j \ coprime ((Pi i)^m) ((Pi j)^n)" +by(rule_tac coprime_exp2_nat, erule_tac Pi_coprime) + +lemma coprime_dvd_mult_nat2: "\coprime (k::nat) n; k dvd n * m\ \ k dvd m" +apply(erule_tac coprime_dvd_mult_nat) +apply(simp add: dvd_def, auto) +apply(rule_tac x = ka in exI) +apply(subgoal_tac "n * m = m * n", simp) +apply(simp add: nat_mult_commute) +done + +declare godel_code'.simps[simp del] + +lemma godel_code'_butlast_last_id' : + "godel_code' (ys @ [y]) (Suc j) = godel_code' ys (Suc j) * + Pi (Suc (length ys + j)) ^ y" +proof(induct ys arbitrary: j, simp_all add: godel_code'.simps) +qed + +lemma godel_code'_butlast_last_id: +"xs \ [] \ godel_code' xs (Suc j) = + godel_code' (butlast xs) (Suc j) * Pi (length xs + j)^(last xs)" +apply(subgoal_tac "\ ys y. xs = ys @ [y]") +apply(erule_tac exE, erule_tac exE, simp add: + godel_code'_butlast_last_id') +apply(rule_tac x = "butlast xs" in exI) +apply(rule_tac x = "last xs" in exI, auto) +done + +lemma godel_code'_not0: "godel_code' xs n \ 0" +apply(induct xs, auto simp: godel_code'.simps) +done + +lemma godel_code_append_cons: + "length xs = i \ godel_code' (xs@y#ys) (Suc 0) + = godel_code' xs (Suc 0) * Pi (Suc i)^y * godel_code' ys (i + 2)" +proof(induct "length xs" arbitrary: i y ys xs, simp add: godel_code'.simps,simp) + fix x xs i y ys + assume ind: + "\xs i y ys. \x = i; length xs = i\ \ + godel_code' (xs @ y # ys) (Suc 0) + = godel_code' xs (Suc 0) * Pi (Suc i) ^ y * + godel_code' ys (Suc (Suc i))" + and h: "Suc x = i" + "length (xs::nat list) = i" + have + "godel_code' (butlast xs @ last xs # ((y::nat)#ys)) (Suc 0) = + godel_code' (butlast xs) (Suc 0) * Pi (Suc (i - 1))^(last xs) + * godel_code' (y#ys) (Suc (Suc (i - 1)))" + apply(rule_tac ind) + using h + by(auto) + moreover have + "godel_code' xs (Suc 0)= godel_code' (butlast xs) (Suc 0) * + Pi (i)^(last xs)" + using godel_code'_butlast_last_id[of xs] h + apply(case_tac "xs = []", simp, simp) + done + moreover have "butlast xs @ last xs # y # ys = xs @ y # ys" + using h + apply(case_tac xs, auto) + done + ultimately show + "godel_code' (xs @ y # ys) (Suc 0) = + godel_code' xs (Suc 0) * Pi (Suc i) ^ y * + godel_code' ys (Suc (Suc i))" + using h + apply(simp add: godel_code'_not0 Pi_not_0) + apply(simp add: godel_code'.simps) + done +qed + +lemma Pi_coprime_pre: + "length ps \ i \ coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" +proof(induct "length ps" arbitrary: ps, simp add: godel_code'.simps) + fix x ps + assume ind: + "\ps. \x = length ps; length ps \ i\ \ + coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" + and h: "Suc x = length ps" + "length (ps::nat list) \ i" + have g: "coprime (Pi (Suc i)) (godel_code' (butlast ps) (Suc 0))" + apply(rule_tac ind) + using h by auto + have k: "godel_code' ps (Suc 0) = + godel_code' (butlast ps) (Suc 0) * Pi (length ps)^(last ps)" + using godel_code'_butlast_last_id[of ps 0] h + by(case_tac ps, simp, simp) + from g have + "coprime (Pi (Suc i)) (godel_code' (butlast ps) (Suc 0) * + Pi (length ps)^(last ps)) " + proof(rule_tac coprime_mult_nat, simp) + show "coprime (Pi (Suc i)) (Pi (length ps) ^ last ps)" + apply(rule_tac coprime_exp_nat, rule prime_coprime, auto) + using Pi_notEq[of "Suc i" "length ps"] h by simp + qed + from this and k show "coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" + by simp +qed + +lemma Pi_coprime_suf: "i < j \ coprime (Pi i) (godel_code' ps j)" +proof(induct "length ps" arbitrary: ps, simp add: godel_code'.simps) + fix x ps + assume ind: + "\ps. \x = length ps; i < j\ \ + coprime (Pi i) (godel_code' ps j)" + and h: "Suc x = length (ps::nat list)" "i < j" + have g: "coprime (Pi i) (godel_code' (butlast ps) j)" + apply(rule ind) using h by auto + have k: "(godel_code' ps j) = godel_code' (butlast ps) j * + Pi (length ps + j - 1)^last ps" + using h godel_code'_butlast_last_id[of ps "j - 1"] + apply(case_tac "ps = []", simp, simp) + done + from g have + "coprime (Pi i) (godel_code' (butlast ps) j * + Pi (length ps + j - 1)^last ps)" + apply(rule_tac coprime_mult_nat, simp) + using Pi_power_coprime[of i "length ps + j - 1" 1 "last ps"] h + apply(auto) + done + from k and this show "coprime (Pi i) (godel_code' ps j)" + by auto +qed + +lemma godel_finite: + "finite {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" +proof(rule_tac n = "godel_code' nl (Suc 0)" in + bounded_nat_set_is_finite, auto, + case_tac "ia < godel_code' nl (Suc 0)", auto) + fix ia + assume g1: "Pi (Suc i) ^ ia dvd godel_code' nl (Suc 0)" + and g2: "\ ia < godel_code' nl (Suc 0)" + from g1 have "Pi (Suc i)^ia \ godel_code' nl (Suc 0)" + apply(erule_tac dvd_imp_le) + using godel_code'_not0[of nl "Suc 0"] by simp + moreover have "ia < Pi (Suc i)^ia" + apply(rule x_less_exp) + using Pi_gr_1 by auto + ultimately show "False" + using g2 + by(auto) +qed + + +lemma godel_code_in: + "i < length nl \ nl ! i \ {u. Pi (Suc i) ^ u dvd + godel_code' nl (Suc 0)}" +proof - + assume h: "i {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" + by(simp) +qed + +lemma godel_code'_get_nth: + "i < length nl \ Max {u. Pi (Suc i) ^ u dvd + godel_code' nl (Suc 0)} = nl ! i" +proof(rule_tac Max_eqI) + let ?gc = "godel_code' nl (Suc 0)" + assume h: "i < length nl" thus "finite {u. Pi (Suc i) ^ u dvd ?gc}" + by (simp add: godel_finite) +next + fix y + let ?suf ="godel_code' (drop (Suc i) nl) (i + 2)" + let ?pref = "godel_code' (take i nl) (Suc 0)" + assume h: "i < length nl" + "y \ {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" + moreover hence + "godel_code' (take i nl@(nl!i)#drop (Suc i) nl) (Suc 0) + = ?pref * Pi (Suc i)^(nl!i) * ?suf" + by(rule_tac godel_code_append_cons, simp) + moreover from h have "take i nl @ (nl!i) # drop (Suc i) nl = nl" + using upd_conv_take_nth_drop[of i nl "nl!i"] + by simp + ultimately show "y\nl!i" + proof(simp) + let ?suf' = "godel_code' (drop (Suc i) nl) (Suc (Suc i))" + assume mult_dvd: + "Pi (Suc i) ^ y dvd ?pref * Pi (Suc i) ^ nl ! i * ?suf'" + hence "Pi (Suc i) ^ y dvd ?pref * Pi (Suc i) ^ nl ! i" + proof(rule_tac coprime_dvd_mult_nat) + show "coprime (Pi (Suc i)^y) ?suf'" + proof - + have "coprime (Pi (Suc i) ^ y) (?suf'^(Suc 0))" + apply(rule_tac coprime_exp2_nat) + apply(rule_tac Pi_coprime_suf, simp) + done + thus "?thesis" by simp + qed + qed + hence "Pi (Suc i) ^ y dvd Pi (Suc i) ^ nl ! i" + proof(rule_tac coprime_dvd_mult_nat2) + show "coprime (Pi (Suc i) ^ y) ?pref" + proof - + have "coprime (Pi (Suc i)^y) (?pref^Suc 0)" + apply(rule_tac coprime_exp2_nat) + apply(rule_tac Pi_coprime_pre, simp) + done + thus "?thesis" by simp + qed + qed + hence "Pi (Suc i) ^ y \ Pi (Suc i) ^ nl ! i " + apply(rule_tac dvd_imp_le, auto) + done + thus "y \ nl ! i" + apply(rule_tac power_le_imp_le_exp, auto) + done + qed +next + assume h: "i {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" + by(rule_tac godel_code_in, simp) +qed + +lemma [simp]: + "{u. Pi (Suc i) ^ u dvd (Suc (Suc 0)) ^ length nl * + godel_code' nl (Suc 0)} = + {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" +apply(rule_tac Collect_cong, auto) +apply(rule_tac n = " (Suc (Suc 0)) ^ length nl" in + coprime_dvd_mult_nat2) +proof - + fix u + show "coprime (Pi (Suc i) ^ u) ((Suc (Suc 0)) ^ length nl)" + proof(rule_tac coprime_exp2_nat) + have "Pi 0 = (2::nat)" + apply(simp add: Pi.simps) + done + moreover have "coprime (Pi (Suc i)) (Pi 0)" + apply(rule_tac Pi_coprime, simp) + done + ultimately show "coprime (Pi (Suc i)) (Suc (Suc 0))" by simp + qed +qed + +lemma godel_code_get_nth: + "i < length nl \ + Max {u. Pi (Suc i) ^ u dvd godel_code nl} = nl ! i" +by(simp add: godel_code.simps godel_code'_get_nth) + +thm trpl.simps + +lemma "trpl l st r = godel_code' [l, st, r] 0" +apply(simp add: trpl.simps godel_code'.simps) +done + +lemma mod_dvd_simp: "(x mod y = (0::nat)) = (y dvd x)" +by(simp add: dvd_def, auto) + +lemma dvd_power_le: "\a > Suc 0; a ^ y dvd a ^ l\ \ y \ l" +apply(case_tac "y \ l", simp, simp) +apply(subgoal_tac "\ d. y = l + d", auto simp: power_add) +apply(rule_tac x = "y - l" in exI, simp) +done + + +lemma [elim]: "Pi n = 0 \ RR" + using Pi_not_0[of n] by simp + +lemma [elim]: "Pi n = Suc 0 \ RR" + using Pi_gr_1[of n] by simp + +lemma finite_power_dvd: + "\(a::nat) > Suc 0; y \ 0\ \ finite {u. a^u dvd y}" +apply(auto simp: dvd_def) +apply(rule_tac n = y in bounded_nat_set_is_finite, auto) +apply(case_tac k, simp,simp) +apply(rule_tac trans_less_add1) +apply(erule_tac x_less_exp) +done + +lemma conf_decode1: "\m \ n; m \ k; k \ n\ \ + Max {u. Pi m ^ u dvd Pi m ^ l * Pi n ^ st * Pi k ^ r} = l" +proof - + let ?setx = "{u. Pi m ^ u dvd Pi m ^ l * Pi n ^ st * Pi k ^ r}" + assume g: "m \ n" "m \ k" "k \ n" + show "Max ?setx = l" + proof(rule_tac Max_eqI) + show "finite ?setx" + apply(rule_tac finite_power_dvd, auto simp: Pi_gr_1) + done + next + fix y + assume h: "y \ ?setx" + have "Pi m ^ y dvd Pi m ^ l" + proof - + have "Pi m ^ y dvd Pi m ^ l * Pi n ^ st" + using h g + apply(rule_tac n = "Pi k^r" in coprime_dvd_mult_nat) + apply(rule Pi_power_coprime, simp, simp) + done + thus "Pi m^y dvd Pi m^l" + apply(rule_tac n = " Pi n ^ st" in coprime_dvd_mult_nat) + using g + apply(rule_tac Pi_power_coprime, simp, simp) + done + qed + thus "y \ (l::nat)" + apply(rule_tac a = "Pi m" in power_le_imp_le_exp) + apply(simp_all add: Pi_gr_1) + apply(rule_tac dvd_power_le, auto) + done + next + show "l \ ?setx" by simp + qed +qed + +lemma conf_decode2: + "\m \ n; m \ k; n \ k; + \ Suc 0 < Pi m ^ l * Pi n ^ st * Pi k ^ r\ \ l = 0" +apply(case_tac "Pi m ^ l * Pi n ^ st * Pi k ^ r", auto) +done + +lemma [simp]: "left (trpl l st r) = l" +apply(simp add: left.simps trpl.simps lo.simps + loR.simps mod_dvd_simp, auto simp: conf_decode1) +apply(case_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r", + auto) +apply(erule_tac x = l in allE, auto) +done + +lemma [simp]: "stat (trpl l st r) = st" +apply(simp add: stat.simps trpl.simps lo.simps + loR.simps mod_dvd_simp, auto) +apply(subgoal_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r + = Pi (Suc 0)^st * Pi 0 ^ l * Pi (Suc (Suc 0)) ^ r") +apply(simp (no_asm_simp) add: conf_decode1, simp) +apply(case_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * + Pi (Suc (Suc 0)) ^ r", auto) +apply(erule_tac x = st in allE, auto) +done + +lemma [simp]: "rght (trpl l st r) = r" +apply(simp add: rght.simps trpl.simps lo.simps + loR.simps mod_dvd_simp, auto) +apply(subgoal_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r + = Pi (Suc (Suc 0))^r * Pi 0 ^ l * Pi (Suc 0) ^ st") +apply(simp (no_asm_simp) add: conf_decode1, simp) +apply(case_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r", + auto) +apply(erule_tac x = r in allE, auto) +done + +lemma max_lor: + "i < length nl \ Max {u. loR [godel_code nl, Pi (Suc i), u]} + = nl ! i" +apply(simp add: loR.simps godel_code_get_nth mod_dvd_simp) +done + +lemma godel_decode: + "i < length nl \ Entry (godel_code nl) i = nl ! i" +apply(auto simp: Entry.simps lo.simps max_lor) +apply(erule_tac x = "nl!i" in allE) +using max_lor[of i nl] godel_finite[of i nl] +apply(simp) +apply(drule_tac Max_in, auto simp: loR.simps + godel_code.simps mod_dvd_simp) +using godel_code_in[of i nl] +apply(simp) +done + +lemma Four_Suc: "4 = Suc (Suc (Suc (Suc 0)))" +by auto + +declare numeral_2_eq_2[simp del] + +lemma modify_tprog_fetch_even: + "\st \ length tp div 2; st > 0\ \ + modify_tprog tp ! (4 * (st - Suc 0) ) = + action_map (fst (tp ! (2 * (st - Suc 0))))" +proof(induct st arbitrary: tp, simp) + fix tp st + assume ind: + "\tp. \st \ length tp div 2; 0 < st\ \ + modify_tprog tp ! (4 * (st - Suc 0)) = + action_map (fst ((tp::tprog) ! (2 * (st - Suc 0))))" + and h: "Suc st \ length (tp::tprog) div 2" "0 < Suc st" + thus "modify_tprog tp ! (4 * (Suc st - Suc 0)) = + action_map (fst (tp ! (2 * (Suc st - Suc 0))))" + proof(cases "st = 0") + case True thus "?thesis" + using h + apply(auto) + apply(cases tp, simp, case_tac a, simp add: modify_tprog.simps) + done + next + case False + assume g: "st \ 0" + hence "\ aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" + using h + apply(case_tac tp, simp, case_tac list, simp, simp) + done + from this obtain aa ab ba bb tp' where g1: + "tp = (aa, ab) # (ba, bb) # tp'" by blast + hence g2: + "modify_tprog tp' ! (4 * (st - Suc 0)) = + action_map (fst ((tp'::tprog) ! (2 * (st - Suc 0))))" + apply(rule_tac ind) + using h g by auto + thus "?thesis" + using g1 g + apply(case_tac st, simp, simp add: Four_Suc) + done + qed +qed + +lemma modify_tprog_fetch_odd: + "\st \ length tp div 2; st > 0\ \ + modify_tprog tp ! (Suc (Suc (4 * (st - Suc 0)))) = + action_map (fst (tp ! (Suc (2 * (st - Suc 0)))))" +proof(induct st arbitrary: tp, simp) + fix tp st + assume ind: + "\tp. \st \ length tp div 2; 0 < st\ \ + modify_tprog tp ! Suc (Suc (4 * (st - Suc 0))) = + action_map (fst (tp ! Suc (2 * (st - Suc 0))))" + and h: "Suc st \ length (tp::tprog) div 2" "0 < Suc st" + thus "modify_tprog tp ! Suc (Suc (4 * (Suc st - Suc 0))) + = action_map (fst (tp ! Suc (2 * (Suc st - Suc 0))))" + proof(cases "st = 0") + case True thus "?thesis" + using h + apply(auto) + apply(cases tp, simp, case_tac a, simp add: modify_tprog.simps) + apply(case_tac list, simp, case_tac ab, + simp add: modify_tprog.simps) + done + next + case False + assume g: "st \ 0" + hence "\ aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" + using h + apply(case_tac tp, simp, case_tac list, simp, simp) + done + from this obtain aa ab ba bb tp' where g1: + "tp = (aa, ab) # (ba, bb) # tp'" by blast + hence g2: "modify_tprog tp' ! Suc (Suc (4 * (st - Suc 0))) = + action_map (fst (tp' ! Suc (2 * (st - Suc 0))))" + apply(rule_tac ind) + using h g by auto + thus "?thesis" + using g1 g + apply(case_tac st, simp, simp add: Four_Suc) + done + qed +qed + +lemma modify_tprog_fetch_action: + "\st \ length tp div 2; st > 0; b = 1 \ b = 0\ \ + modify_tprog tp ! (4 * (st - Suc 0) + 2* b) = + action_map (fst (tp ! ((2 * (st - Suc 0)) + b)))" +apply(erule_tac disjE, auto elim: modify_tprog_fetch_odd + modify_tprog_fetch_even) +done + +lemma length_modify: "length (modify_tprog tp) = 2 * length tp" +apply(induct tp, auto) +done + +declare fetch.simps[simp del] + +lemma fetch_action_eq: + "\block_map b = scan r; fetch tp st b = (nact, ns); + st \ length tp div 2\ \ actn (code tp) st r = action_map nact" +proof(simp add: actn.simps, auto) + let ?i = "4 * (st - Suc 0) + 2 * (r mod 2)" + assume h: "block_map b = r mod 2" "fetch tp st b = (nact, ns)" + "st \ length tp div 2" "0 < st" + have "?i < length (modify_tprog tp)" + proof - + have "length (modify_tprog tp) = 2 * length tp" + by(simp add: length_modify) + thus "?thesis" + using h + by(auto) + qed + hence + "Entry (godel_code (modify_tprog tp))?i = + (modify_tprog tp) ! ?i" + by(erule_tac godel_decode) + thm modify_tprog.simps + moreover have + "modify_tprog tp ! ?i = + action_map (fst (tp ! (2 * (st - Suc 0) + r mod 2)))" + apply(rule_tac modify_tprog_fetch_action) + using h + by(auto) + moreover have "(fst (tp ! (2 * (st - Suc 0) + r mod 2))) = nact" + using h + apply(simp add: fetch.simps nth_of.simps) + apply(case_tac b, auto simp: block_map.simps nth_of.simps split: if_splits) + done + ultimately show + "Entry (godel_code (modify_tprog tp)) + (4 * (st - Suc 0) + 2 * (r mod 2)) + = action_map nact" + by simp +qed + +lemma [simp]: "fetch tp 0 b = (nact, ns) \ ns = 0" +by(simp add: fetch.simps) + +lemma Five_Suc: "5 = Suc 4" by simp + +lemma modify_tprog_fetch_state: + "\st \ length tp div 2; st > 0; b = 1 \ b = 0\ \ + modify_tprog tp ! Suc (4 * (st - Suc 0) + 2 * b) = + (snd (tp ! (2 * (st - Suc 0) + b)))" +proof(induct st arbitrary: tp, simp) + fix st tp + assume ind: + "\tp. \st \ length tp div 2; 0 < st; b = 1 \ b = 0\ \ + modify_tprog tp ! Suc (4 * (st - Suc 0) + 2 * b) = + snd (tp ! (2 * (st - Suc 0) + b))" + and h: + "Suc st \ length (tp::tprog) div 2" + "0 < Suc st" + "b = 1 \ b = 0" + show "modify_tprog tp ! Suc (4 * (Suc st - Suc 0) + 2 * b) = + snd (tp ! (2 * (Suc st - Suc 0) + b))" + proof(cases "st = 0") + case True + thus "?thesis" + using h + apply(cases tp, simp, case_tac a, simp add: modify_tprog.simps) + apply(case_tac list, simp, case_tac ab, + simp add: modify_tprog.simps, auto) + done + next + case False + assume g: "st \ 0" + hence "\ aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" + using h + apply(case_tac tp, simp, case_tac list, simp, simp) + done + from this obtain aa ab ba bb tp' where g1: + "tp = (aa, ab) # (ba, bb) # tp'" by blast + hence g2: + "modify_tprog tp' ! Suc (4 * (st - Suc 0) + 2 * b) = + snd (tp' ! (2 * (st - Suc 0) + b))" + apply(rule_tac ind) + using h g by auto + thus "?thesis" + using g1 g + apply(case_tac st, simp, simp) + done + qed +qed + +lemma fetch_state_eq: + "\block_map b = scan r; + fetch tp st b = (nact, ns); + st \ length tp div 2\ \ newstat (code tp) st r = ns" +proof(simp add: newstat.simps, auto) + let ?i = "Suc (4 * (st - Suc 0) + 2 * (r mod 2))" + assume h: "block_map b = r mod 2" "fetch tp st b = + (nact, ns)" "st \ length tp div 2" "0 < st" + have "?i < length (modify_tprog tp)" + proof - + have "length (modify_tprog tp) = 2 * length tp" + apply(simp add: length_modify) + done + thus "?thesis" + using h + by(auto) + qed + hence "Entry (godel_code (modify_tprog tp)) (?i) = + (modify_tprog tp) ! ?i" + by(erule_tac godel_decode) + thm modify_tprog.simps + moreover have + "modify_tprog tp ! ?i = + (snd (tp ! (2 * (st - Suc 0) + r mod 2)))" + apply(rule_tac modify_tprog_fetch_state) + using h + by(auto) + moreover have "(snd (tp ! (2 * (st - Suc 0) + r mod 2))) = ns" + using h + apply(simp add: fetch.simps nth_of.simps) + apply(case_tac b, auto simp: block_map.simps nth_of.simps + split: if_splits) + done + ultimately show "Entry (godel_code (modify_tprog tp)) (?i) + = ns" + by simp +qed + + +lemma [intro!]: + "\a = a'; b = b'; c = c'\ \ trpl a b c = trpl a' b' c'" +by simp + +lemma [simp]: "bl2wc [Bk] = 0" +by(simp add: bl2wc.simps bl2nat.simps) + +lemma bl2nat_double: "bl2nat xs (Suc n) = 2 * bl2nat xs n" +proof(induct xs arbitrary: n) + case Nil thus "?case" + by(simp add: bl2nat.simps) +next + case (Cons x xs) thus "?case" + proof - + assume ind: "\n. bl2nat xs (Suc n) = 2 * bl2nat xs n " + show "bl2nat (x # xs) (Suc n) = 2 * bl2nat (x # xs) n" + proof(cases x) + case Bk thus "?thesis" + apply(simp add: bl2nat.simps) + using ind[of "Suc n"] by simp + next + case Oc thus "?thesis" + apply(simp add: bl2nat.simps) + using ind[of "Suc n"] by simp + qed + qed +qed + + +lemma [simp]: "c \ [] \ 2 * bl2wc (tl c) = bl2wc c - bl2wc c mod 2 " +apply(case_tac c, simp, case_tac a) +apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) +done + +lemma [simp]: + "c \ [] \ bl2wc (Oc # tl c) = Suc (bl2wc c) - bl2wc c mod 2 " +apply(case_tac c, simp, case_tac a) +apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) +done + +lemma [simp]: "bl2wc (Bk # c) = 2*bl2wc (c)" +apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double) +done + +lemma [simp]: "bl2wc [Oc] = Suc 0" + by(simp add: bl2wc.simps bl2nat.simps) + +lemma [simp]: "b \ [] \ bl2wc (tl b) = bl2wc b div 2" +apply(case_tac b, simp, case_tac a) +apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) +done + +lemma [simp]: "b \ [] \ bl2wc ([hd b]) = bl2wc b mod 2" +apply(case_tac b, simp, case_tac a) +apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) +done + +lemma [simp]: "\b \ []; c \ []\ \ bl2wc (hd b # c) = 2 * bl2wc c + bl2wc b mod 2" +apply(case_tac b, simp, case_tac a) +apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) +done + +lemma [simp]: " 2 * (bl2wc c div 2) = bl2wc c - bl2wc c mod 2" + by(simp add: mult_div_cancel) + +lemma [simp]: "bl2wc (Oc # list) mod 2 = Suc 0" + by(simp add: bl2wc.simps bl2nat.simps bl2nat_double) + + +declare code.simps[simp del] +declare nth_of.simps[simp del] +declare new_tape.simps[simp del] + +text {* + The lemma relates the one step execution of TMs with the interpreter function @{text "rec_newconf"}. + *} +lemma rec_t_eq_step: + "(\ (s, l, r). s \ length tp div 2) c \ + trpl_code (tstep c tp) = + rec_exec rec_newconf [code tp, trpl_code c]" +apply(cases c, auto simp: tstep.simps) +proof(case_tac "fetch tp a (case c of [] \ Bk | x # xs \ x)", + simp add: newconf.simps trpl_code.simps) + fix a b c aa ba + assume h: "(a::nat) \ length tp div 2" + "fetch tp a (case c of [] \ Bk | x # xs \ x) = (aa, ba)" + moreover hence "actn (code tp) a (bl2wc c) = action_map aa" + apply(rule_tac b = "(case c of [] \ Bk | x # xs \ x)" + in fetch_action_eq, auto) + apply(auto split: list.splits) + apply(case_tac ab, auto) + done + moreover from h have "(newstat (code tp) a (bl2wc c)) = ba" + apply(rule_tac b = "(case c of [] \ Bk | x # xs \ x)" + in fetch_state_eq, auto split: list.splits) + apply(case_tac ab, auto) + done + ultimately show + "trpl_code (ba, new_tape aa (b, c)) = + trpl (newleft (bl2wc b) (bl2wc c) (actn (code tp) a (bl2wc c))) + (newstat (code tp) a (bl2wc c)) (newrght (bl2wc b) (bl2wc c) + (actn (code tp) a (bl2wc c)))" + by(auto simp: new_tape.simps trpl_code.simps + newleft.simps newrght.simps split: taction.splits) +qed + +lemma [simp]: "a\<^bsup>0\<^esup> = []" +apply(simp add: exp_zero) +done +lemma [simp]: "bl2nat (Oc # Oc\<^bsup>x\<^esup>) 0 = (2 * 2 ^ x - Suc 0)" +apply(induct x) +apply(simp add: bl2nat.simps) +apply(simp add: bl2nat.simps bl2nat_double exp_ind_def) +done + +lemma [simp]: "bl2nat (Oc\<^bsup>y\<^esup>) 0 = 2^y - Suc 0" +apply(induct y, auto simp: bl2nat.simps exp_ind_def bl2nat_double) +apply(case_tac "(2::nat)^y", auto) +done + +lemma [simp]: "bl2nat (Bk\<^bsup>l\<^esup>) n = 0" +apply(induct l, auto simp: bl2nat.simps bl2nat_double exp_ind_def) +done + +lemma bl2nat_cons_bk: "bl2nat (ks @ [Bk]) 0 = bl2nat ks 0" +apply(induct ks, auto simp: bl2nat.simps split: block.splits) +apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) +done + +lemma bl2nat_cons_oc: + "bl2nat (ks @ [Oc]) 0 = bl2nat ks 0 + 2 ^ length ks" +apply(induct ks, auto simp: bl2nat.simps split: block.splits) +apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) +done + +lemma bl2nat_append: + "bl2nat (xs @ ys) 0 = bl2nat xs 0 + bl2nat ys (length xs) " +proof(induct "length xs" arbitrary: xs ys, simp add: bl2nat.simps) + fix x xs ys + assume ind: + "\xs ys. x = length xs \ + bl2nat (xs @ ys) 0 = bl2nat xs 0 + bl2nat ys (length xs)" + and h: "Suc x = length (xs::block list)" + have "\ ks k. xs = ks @ [k]" + apply(rule_tac x = "butlast xs" in exI, + rule_tac x = "last xs" in exI) + using h + apply(case_tac xs, auto) + done + from this obtain ks k where "xs = ks @ [k]" by blast + moreover hence + "bl2nat (ks @ (k # ys)) 0 = bl2nat ks 0 + + bl2nat (k # ys) (length ks)" + apply(rule_tac ind) using h by simp + ultimately show "bl2nat (xs @ ys) 0 = + bl2nat xs 0 + bl2nat ys (length xs)" + apply(case_tac k, simp_all add: bl2nat.simps) + apply(simp_all only: bl2nat_cons_bk bl2nat_cons_oc) + done +qed + +lemma bl2nat_exp: "n \ 0 \ bl2nat bl n = 2^n * bl2nat bl 0" +apply(induct bl) +apply(auto simp: bl2nat.simps) +apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) +done + +lemma nat_minus_eq: "\a = b; c = d\ \ a - c = b - d" +by auto + +lemma tape_of_nat_list_butlast_last: + "ys \ [] \ = @ Bk # Oc\<^bsup>Suc y\<^esup>" +apply(induct ys, simp, simp) +apply(case_tac "ys = []", simp add: tape_of_nl_abv + tape_of_nat_list.simps) +apply(simp) +done + +lemma listsum2_append: + "\n \ length xs\ \ listsum2 (xs @ ys) n = listsum2 xs n" +apply(induct n) +apply(auto simp: listsum2.simps nth_append) +done + +lemma strt'_append: + "\n \ length xs\ \ strt' xs n = strt' (xs @ ys) n" +proof(induct n arbitrary: xs ys) + fix xs ys + show "strt' xs 0 = strt' (xs @ ys) 0" by(simp add: strt'.simps) +next + fix n xs ys + assume ind: + "\ xs ys. n \ length xs \ strt' xs n = strt' (xs @ ys) n" + and h: "Suc n \ length (xs::nat list)" + show "strt' xs (Suc n) = strt' (xs @ ys) (Suc n)" + using ind[of xs ys] h + apply(simp add: strt'.simps nth_append listsum2_append) + done +qed + +lemma length_listsum2_eq: + "\length (ys::nat list) = k\ + \ length () = listsum2 (map Suc ys) k + k - 1" +apply(induct k arbitrary: ys, simp_all add: listsum2.simps) +apply(subgoal_tac "\ xs x. ys = xs @ [x]", auto) +proof - + fix xs x + assume ind: "\ys. length ys = length xs \ length () + = listsum2 (map Suc ys) (length xs) + + length (xs::nat list) - Suc 0" + have "length () + = listsum2 (map Suc xs) (length xs) + length xs - Suc 0" + apply(rule_tac ind, simp) + done + thus "length () = + Suc (listsum2 (map Suc xs @ [Suc x]) (length xs) + x + length xs)" + apply(case_tac "xs = []") + apply(simp add: tape_of_nl_abv listsum2.simps + tape_of_nat_list.simps) + apply(simp add: tape_of_nat_list_butlast_last) + using listsum2_append[of "length xs" "map Suc xs" "[Suc x]"] + apply(simp) + done +next + fix k ys + assume "length ys = Suc k" + thus "\xs x. ys = xs @ [x]" + apply(rule_tac x = "butlast ys" in exI, + rule_tac x = "last ys" in exI) + apply(case_tac ys, auto) + done +qed + +lemma tape_of_nat_list_length: + "length (<(ys::nat list)>) = + listsum2 (map Suc ys) (length ys) + length ys - 1" + using length_listsum2_eq[of ys "length ys"] + apply(simp) + done + + + +lemma [simp]: + "trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp 0) = + rec_exec rec_conf [code tp, bl2wc (), 0]" +apply(simp add: steps.simps rec_exec.simps conf_lemma conf.simps + inpt.simps trpl_code.simps bl2wc.simps) +done + +text {* + The following lemma relates the multi-step interpreter function @{text "rec_conf"} + with the multi-step execution of TMs. + *} +lemma rec_t_eq_steps: + "turing_basic.t_correct tp \ + trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp) = + rec_exec rec_conf [code tp, bl2wc (), stp]" +proof(induct stp) + case 0 thus "?case" by(simp) +next + case (Suc n) thus "?case" + proof - + assume ind: + "t_correct tp \ trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp n) + = rec_exec rec_conf [code tp, bl2wc (), n]" + and h: "t_correct tp" + show + "trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp (Suc n)) = + rec_exec rec_conf [code tp, bl2wc (), Suc n]" + proof(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp n", + simp only: tstep_red conf_lemma conf.simps) + fix a b c + assume g: "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp n = (a, b, c) " + hence "conf (code tp) (bl2wc ()) n= trpl_code (a, b, c)" + using ind h + apply(simp add: conf_lemma) + done + moreover hence + "trpl_code (tstep (a, b, c) tp) = + rec_exec rec_newconf [code tp, trpl_code (a, b, c)]" + thm rec_t_eq_step + apply(rule_tac rec_t_eq_step) + using h g + apply(simp add: s_keep) + done + ultimately show + "trpl_code (tstep (a, b, c) tp) = + newconf (code tp) (conf (code tp) (bl2wc ()) n)" + by(simp add: newconf_lemma) + qed + qed +qed + +lemma [simp]: "bl2wc (Bk\<^bsup>m\<^esup>) = 0" +apply(induct m) +apply(simp, simp) +done + +lemma [simp]: "bl2wc (Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>) = bl2wc (Oc\<^bsup>rs\<^esup>)" +apply(induct rs, simp, + simp add: bl2wc.simps bl2nat.simps bl2nat_double) +done + +lemma lg_power: "x > Suc 0 \ lg (x ^ rs) x = rs" +proof(simp add: lg.simps, auto) + fix xa + assume h: "Suc 0 < x" + show "Max {ya. ya \ x ^ rs \ lgR [x ^ rs, x, ya]} = rs" + apply(rule_tac Max_eqI, simp_all add: lgR.simps) + apply(simp add: h) + using x_less_exp[of x rs] h + apply(simp) + done +next + assume "\ Suc 0 < x ^ rs" "Suc 0 < x" + thus "rs = 0" + apply(case_tac "x ^ rs", simp, simp) + done +next + assume "Suc 0 < x" "\xa. \ lgR [x ^ rs, x, xa]" + thus "rs = 0" + apply(simp only:lgR.simps) + apply(erule_tac x = rs in allE, simp) + done +qed + +text {* + The following lemma relates execution of TMs with + the multi-step interpreter function @{text "rec_nonstop"}. Note, + @{text "rec_nonstop"} is constructed using @{text "rec_conf"}. + *} +lemma nonstop_t_eq: + "\steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>); + turing_basic.t_correct tp; + rs > 0\ + \ rec_exec rec_nonstop [code tp, bl2wc (), stp] = 0" +proof(simp add: nonstop_lemma nonstop.simps nstd.simps) + assume h: "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" + and tc_t: "turing_basic.t_correct tp" "rs > 0" + have g: "rec_exec rec_conf [code tp, bl2wc (), stp] = + trpl_code (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)" + using rec_t_eq_steps[of tp l lm stp] tc_t h + by(simp) + thus "\ NSTD (conf (code tp) (bl2wc ()) stp)" + proof(auto simp: NSTD.simps) + show "stat (conf (code tp) (bl2wc ()) stp) = 0" + using g + by(auto simp: conf_lemma trpl_code.simps) + next + show "left (conf (code tp) (bl2wc ()) stp) = 0" + using g + by(simp add: conf_lemma trpl_code.simps) + next + show "rght (conf (code tp) (bl2wc ()) stp) = + 2 ^ lg (Suc (rght (conf (code tp) (bl2wc ()) stp))) 2 - Suc 0" + using g h + proof(simp add: conf_lemma trpl_code.simps) + have "2 ^ lg (Suc (bl2wc (Oc\<^bsup>rs\<^esup>))) 2 = Suc (bl2wc (Oc\<^bsup>rs\<^esup>))" + apply(simp add: bl2wc.simps lg_power) + done + thus "bl2wc (Oc\<^bsup>rs\<^esup>) = 2 ^ lg (Suc (bl2wc (Oc\<^bsup>rs\<^esup>))) 2 - Suc 0" + apply(simp) + done + qed + next + show "0 < rght (conf (code tp) (bl2wc ()) stp)" + using g h tc_t + apply(simp add: conf_lemma trpl_code.simps bl2wc.simps + bl2nat.simps) + apply(case_tac rs, simp, simp add: bl2nat.simps) + done + qed +qed + +lemma [simp]: "actn m 0 r = 4" +by(simp add: actn.simps) + +lemma [simp]: "newstat m 0 r = 0" +by(simp add: newstat.simps) + +declare exp_def[simp del] + +lemma halt_least_step: + "\steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs \<^esup> @ Bk\<^bsup>n\<^esup>); + turing_basic.t_correct tp; + 0 \ + \ stp. (nonstop (code tp) (bl2wc ()) stp = 0 \ + (\ stp'. nonstop (code tp) (bl2wc ()) stp' = 0 \ stp \ stp'))" +proof(induct stp, simp add: steps.simps, simp) + fix stp + assume ind: + "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) \ + \stp. nonstop (code tp) (bl2wc ()) stp = 0 \ + (\stp'. nonstop (code tp) (bl2wc ()) stp' = 0 \ stp \ stp')" + and h: + "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp (Suc stp) = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" + "turing_basic.t_correct tp" + "0 < rs" + from h show + "\stp. nonstop (code tp) (bl2wc ()) stp = 0 + \ (\stp'. nonstop (code tp) (bl2wc ()) stp' = 0 \ stp \ stp')" + proof(simp add: tstep_red, + case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp", simp, + case_tac a, simp add: tstep_0) + assume "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" + thus "\stp. nonstop (code tp) (bl2wc ()) stp = 0 \ + (\stp'. nonstop (code tp) (bl2wc ()) stp' = 0 \ stp \ stp')" + apply(erule_tac ind) + done + next + fix a b c nat + assume "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (a, b, c)" + "a = Suc nat" + thus "\stp. nonstop (code tp) (bl2wc ()) stp = 0 \ + (\stp'. nonstop (code tp) (bl2wc ()) stp' = 0 \ stp \ stp')" + using h + apply(rule_tac x = "Suc stp" in exI, auto) + apply(drule_tac nonstop_t_eq, simp_all add: nonstop_lemma) + proof - + fix stp' + assume g:"steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (Suc nat, b, c)" + "nonstop (code tp) (bl2wc ()) stp' = 0" + thus "Suc stp \ stp'" + proof(case_tac "Suc stp \ stp'", simp, simp) + assume "\ Suc stp \ stp'" + hence "stp' \ stp" by simp + hence "\ isS0 (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp')" + using g + apply(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp'",auto, + simp add: isS0_def) + apply(subgoal_tac "\ n. stp = stp' + n", + auto simp: steps_add steps_0) + apply(rule_tac x = "stp - stp'" in exI, simp) + done + hence "nonstop (code tp) (bl2wc ()) stp' = 1" + proof(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp'", + simp add: isS0_def nonstop.simps) + fix a b c + assume k: + "0 < a" "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp' = (a, b, c)" + thus " NSTD (conf (code tp) (bl2wc ()) stp')" + using rec_t_eq_steps[of tp l lm stp'] h + proof(simp add: conf_lemma) + assume "trpl_code (a, b, c) = conf (code tp) (bl2wc ()) stp'" + moreover have "NSTD (trpl_code (a, b, c))" + using k + apply(auto simp: trpl_code.simps NSTD.simps) + done + ultimately show "NSTD (conf (code tp) (bl2wc ()) stp')" by simp + qed + qed + thus "False" using g by simp + qed + qed + qed +qed + +(* +lemma halt_steps_ex: + "\steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>); + lm \ []; turing_basic.t_correct tp; 0 \ + \ t. rec_calc_rel (rec_halt (length lm)) (code tp # lm) t" +apply(drule_tac halt_least_step, auto) +apply(rule_tac x = stp in exI) +apply(simp add: halt_lemma nonstop_lemma) +apply(auto) +done*) +thm loR.simps + +lemma conf_trpl_ex: "\ p q r. conf m (bl2wc ()) stp = trpl p q r" +apply(induct stp, auto simp: conf.simps inpt.simps trpl.simps + newconf.simps) +apply(rule_tac x = 0 in exI, rule_tac x = 1 in exI, + rule_tac x = "bl2wc ()" in exI) +apply(simp) +done + +lemma nonstop_rgt_ex: + "nonstop m (bl2wc ()) stpa = 0 \ \ r. conf m (bl2wc ()) stpa = trpl 0 0 r" +apply(auto simp: nonstop.simps NSTD.simps split: if_splits) +using conf_trpl_ex[of m lm stpa] +apply(auto) +done + +lemma [elim]: "x > Suc 0 \ Max {u. x ^ u dvd x ^ r} = r" +proof(rule_tac Max_eqI) + assume "x > Suc 0" + thus "finite {u. x ^ u dvd x ^ r}" + apply(rule_tac finite_power_dvd, auto) + done +next + fix y + assume "Suc 0 < x" "y \ {u. x ^ u dvd x ^ r}" + thus "y \ r" + apply(case_tac "y\ r", simp) + apply(subgoal_tac "\ d. y = r + d") + apply(auto simp: power_add) + apply(rule_tac x = "y - r" in exI, simp) + done +next + show "r \ {u. x ^ u dvd x ^ r}" by simp +qed + +lemma lo_power: "x > Suc 0 \ lo (x ^ r) x = r" +apply(auto simp: lo.simps loR.simps mod_dvd_simp) +apply(case_tac "x^r", simp_all) +done + +lemma lo_rgt: "lo (trpl 0 0 r) (Pi 2) = r" +apply(simp add: trpl.simps lo_power) +done + +lemma conf_keep: + "conf m lm stp = trpl 0 0 r \ + conf m lm (stp + n) = trpl 0 0 r" +apply(induct n) +apply(auto simp: conf.simps newconf.simps newleft.simps + newrght.simps rght.simps lo_rgt) +done + +lemma halt_state_keep_steps_add: + "\nonstop m (bl2wc ()) stpa = 0\ \ + conf m (bl2wc ()) stpa = conf m (bl2wc ()) (stpa + n)" +apply(drule_tac nonstop_rgt_ex, auto simp: conf_keep) +done + +lemma halt_state_keep: + "\nonstop m (bl2wc ()) stpa = 0; nonstop m (bl2wc ()) stpb = 0\ \ + conf m (bl2wc ()) stpa = conf m (bl2wc ()) stpb" +apply(case_tac "stpa > stpb") +using halt_state_keep_steps_add[of m lm stpb "stpa - stpb"] +apply simp +using halt_state_keep_steps_add[of m lm stpa "stpb - stpa"] +apply(simp) +done + +thm halt_lemma + +text {* + The correntess of @{text "rec_F"} which relates the interpreter function @{text "rec_F"} with the + execution of of TMs. + *} +lemma F_t_halt_eq: + "\steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>); + turing_basic.t_correct tp; + 0 + \ rec_calc_rel rec_F [code tp, (bl2wc ())] (rs - Suc 0)" +apply(frule_tac halt_least_step, auto) +apply(frule_tac nonstop_t_eq, auto simp: nonstop_lemma) +using rec_t_eq_steps[of tp l lm stp] +apply(simp add: conf_lemma) +proof - + fix stpa + assume h: + "nonstop (code tp) (bl2wc ()) stpa = 0" + "\stp'. nonstop (code tp) (bl2wc ()) stp' = 0 \ stpa \ stp'" + "nonstop (code tp) (bl2wc ()) stp = 0" + "trpl_code (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) = conf (code tp) (bl2wc ()) stp" + "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" + hence g1: "conf (code tp) (bl2wc ()) stpa = trpl_code (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" + using halt_state_keep[of "code tp" lm stpa stp] + by(simp) + moreover have g2: + "rec_calc_rel rec_halt [code tp, (bl2wc ())] stpa" + using h + apply(simp add: halt_lemma nonstop_lemma, auto) + done + show + "rec_calc_rel rec_F [code tp, (bl2wc ())] (rs - Suc 0)" + proof - + have + "rec_calc_rel rec_F [code tp, (bl2wc ())] + (valu (rght (conf (code tp) (bl2wc ()) stpa)))" + apply(rule F_lemma) using g2 h by auto + moreover have + "valu (rght (conf (code tp) (bl2wc ()) stpa)) = rs - Suc 0" + using g1 + apply(simp add: valu.simps trpl_code.simps + bl2wc.simps bl2nat_append lg_power) + done + ultimately show "?thesis" by simp + qed +qed + + +end \ No newline at end of file diff -r cbb4ac6c8081 -r 1ce04eb1c8ad utm/UTM.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/utm/UTM.thy Sat Sep 29 12:38:12 2012 +0000 @@ -0,0 +1,4704 @@ +theory UTM +imports Main uncomputable recursive abacus UF GCD +begin + +section {* Wang coding of input arguments *} + +text {* + The direct compilation of the universal function @{text "rec_F"} can not give us UTM, because @{text "rec_F"} is of arity 2, + where the first argument represents the Godel coding of the TM being simulated and the second argument represents the right number (in Wang's coding) of the TM tape. + (Notice, left number is always @{text "0"} at the very beginning). However, UTM needs to simulate the execution of any TM which may + very well take many input arguments. Therefore, a initialization TM needs to run before the TM compiled from @{text "rec_F"}, and the sequential + composition of these two TMs will give rise to the UTM we are seeking. The purpose of this initialization TM is to transform the multiple + input arguments of the TM being simulated into Wang's coding, so that it can be consumed by the TM compiled from @{text "rec_F"} as the second + argument. + + However, this initialization TM (named @{text "t_wcode"}) can not be constructed by compiling from any resurve function, because every recursive + function takes a fixed number of input arguments, while @{text "t_wcode"} needs to take varying number of arguments and tranform them into + Wang's coding. Therefore, this section give a direct construction of @{text "t_wcode"} with just some parts being obtained from recursive functions. +*} + +definition rec_twice :: "recf" + where + "rec_twice = Cn 1 rec_mult [id 1 0, constn 2]" + +definition rec_fourtimes :: "recf" + where + "rec_fourtimes = Cn 1 rec_mult [id 1 0, constn 4]" + +definition abc_twice :: "abc_prog" + where + "abc_twice = (let (aprog, ary, fp) = rec_ci rec_twice in + aprog [+] dummy_abc ((Suc 0)))" + +definition abc_fourtimes :: "abc_prog" + where + "abc_fourtimes = (let (aprog, ary, fp) = rec_ci rec_fourtimes in + aprog [+] dummy_abc ((Suc 0)))" + +definition twice_ly :: "nat list" + where + "twice_ly = layout_of abc_twice" + +definition fourtimes_ly :: "nat list" + where + "fourtimes_ly = layout_of abc_fourtimes" + +definition t_twice :: "tprog" + where + "t_twice = change_termi_state (tm_of (abc_twice) @ (tMp 1 (start_of twice_ly (length abc_twice) - Suc 0)))" + +definition t_fourtimes :: "tprog" + where + "t_fourtimes = change_termi_state (tm_of (abc_fourtimes) @ + (tMp 1 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)))" + + +definition t_twice_len :: "nat" + where + "t_twice_len = length t_twice div 2" + +definition t_wcode_main_first_part:: "tprog" + where + "t_wcode_main_first_part \ + [(L, 1), (L, 2), (L, 7), (R, 3), + (R, 4), (W0, 3), (R, 4), (R, 5), + (W1, 6), (R, 5), (R, 13), (L, 6), + (R, 0), (R, 8), (R, 9), (Nop, 8), + (R, 10), (W0, 9), (R, 10), (R, 11), + (W1, 12), (R, 11), (R, t_twice_len + 14), (L, 12)]" + +definition t_wcode_main :: "tprog" + where + "t_wcode_main = (t_wcode_main_first_part @ tshift t_twice 12 @ [(L, 1), (L, 1)] + @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])" + +fun bl_bin :: "block list \ nat" + where + "bl_bin [] = 0" +| "bl_bin (Bk # xs) = 2 * bl_bin xs" +| "bl_bin (Oc # xs) = Suc (2 * bl_bin xs)" + +declare bl_bin.simps[simp del] + +type_synonym bin_inv_t = "block list \ nat \ tape \ bool" + +fun wcode_before_double :: "bin_inv_t" + where + "wcode_before_double ires rs (l, r) = + (\ ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \ + r = Oc\<^bsup>(Suc (Suc rs))\<^esup> @ Bk\<^bsup>rn \<^esup>)" + +declare wcode_before_double.simps[simp del] + +fun wcode_after_double :: "bin_inv_t" + where + "wcode_after_double ires rs (l, r) = + (\ ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \ + r = Oc\<^bsup>Suc (Suc (Suc 2*rs))\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +declare wcode_after_double.simps[simp del] + +fun wcode_on_left_moving_1_B :: "bin_inv_t" + where + "wcode_on_left_moving_1_B ires rs (l, r) = + (\ ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Oc # ires \ + r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr > Suc 0 \ mr > 0)" + +declare wcode_on_left_moving_1_B.simps[simp del] + +fun wcode_on_left_moving_1_O :: "bin_inv_t" + where + "wcode_on_left_moving_1_O ires rs (l, r) = + (\ ln rn. + l = Oc # ires \ + r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +declare wcode_on_left_moving_1_O.simps[simp del] + +fun wcode_on_left_moving_1 :: "bin_inv_t" + where + "wcode_on_left_moving_1 ires rs (l, r) = + (wcode_on_left_moving_1_B ires rs (l, r) \ wcode_on_left_moving_1_O ires rs (l, r))" + +declare wcode_on_left_moving_1.simps[simp del] + +fun wcode_on_checking_1 :: "bin_inv_t" + where + "wcode_on_checking_1 ires rs (l, r) = + (\ ln rn. l = ires \ + r = Oc # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wcode_erase1 :: "bin_inv_t" + where +"wcode_erase1 ires rs (l, r) = + (\ ln rn. l = Oc # ires \ + tl r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +declare wcode_erase1.simps [simp del] + +fun wcode_on_right_moving_1 :: "bin_inv_t" + where + "wcode_on_right_moving_1 ires rs (l, r) = + (\ ml mr rn. + l = Bk\<^bsup>ml\<^esup> @ Oc # ires \ + r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr > Suc 0)" + +declare wcode_on_right_moving_1.simps [simp del] + +declare wcode_on_right_moving_1.simps[simp del] + +fun wcode_goon_right_moving_1 :: "bin_inv_t" + where + "wcode_goon_right_moving_1 ires rs (l, r) = + (\ ml mr ln rn. + l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \ + r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr = Suc rs)" + +declare wcode_goon_right_moving_1.simps[simp del] + +fun wcode_backto_standard_pos_B :: "bin_inv_t" + where + "wcode_backto_standard_pos_B ires rs (l, r) = + (\ ln rn. l = Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \ + r = Bk # Oc\<^bsup>(Suc (Suc rs))\<^esup> @ Bk\<^bsup>rn \<^esup>)" + +declare wcode_backto_standard_pos_B.simps[simp del] + +fun wcode_backto_standard_pos_O :: "bin_inv_t" + where + "wcode_backto_standard_pos_O ires rs (l, r) = + (\ ml mr ln rn. + l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \ + r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr = Suc (Suc rs) \ mr > 0)" + +declare wcode_backto_standard_pos_O.simps[simp del] + +fun wcode_backto_standard_pos :: "bin_inv_t" + where + "wcode_backto_standard_pos ires rs (l, r) = (wcode_backto_standard_pos_B ires rs (l, r) \ + wcode_backto_standard_pos_O ires rs (l, r))" + +declare wcode_backto_standard_pos.simps[simp del] + +lemma [simp]: "<0::nat> = [Oc]" +apply(simp add: tape_of_nat_abv exponent_def tape_of_nat_list.simps) +done + +lemma tape_of_Suc_nat: " = replicate a Oc @ [Oc, Oc]" +apply(simp add: tape_of_nat_abv exp_ind tape_of_nat_list.simps) +apply(simp only: exp_ind_def[THEN sym]) +apply(simp only: exp_ind, simp, simp add: exponent_def) +done + +lemma [simp]: "length () = Suc a" +apply(simp add: tape_of_nat_abv tape_of_nat_list.simps) +done + +lemma [simp]: "<[a::nat]> = " +apply(simp add: tape_of_nat_abv tape_of_nl_abv exponent_def + tape_of_nat_list.simps) +done + +lemma bin_wc_eq: "bl_bin xs = bl2wc xs" +proof(induct xs) + show " bl_bin [] = bl2wc []" + apply(simp add: bl_bin.simps) + done +next + fix a xs + assume "bl_bin xs = bl2wc xs" + thus " bl_bin (a # xs) = bl2wc (a # xs)" + apply(case_tac a, simp_all add: bl_bin.simps bl2wc.simps) + apply(simp_all add: bl2nat.simps bl2nat_double) + done +qed + +declare exp_def[simp del] + +lemma bl_bin_nat_Suc: + "bl_bin () = bl_bin () + 2^(Suc a)" +apply(simp add: tape_of_nat_abv bin_wc_eq) +apply(simp add: bl2wc.simps) +done +lemma [simp]: " rev (a\<^bsup>aa\<^esup>) = a\<^bsup>aa\<^esup>" +apply(simp add: exponent_def) +done + +declare tape_of_nl_abv_cons[simp del] + +lemma tape_of_nl_rev: "rev () = ()" +apply(induct lm rule: list_tl_induct, simp) +apply(case_tac "list = []", simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(simp add: tape_of_nat_list_butlast_last tape_of_nl_abv_cons) +done +lemma [simp]: "a\<^bsup>Suc 0\<^esup> = [a]" +by(simp add: exp_def) +lemma tape_of_nl_cons_app1: "() = (Oc\<^bsup>Suc a\<^esup> @ Bk # ())" +apply(case_tac xs, simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma bl_bin_bk_oc[simp]: + "bl_bin (xs @ [Bk, Oc]) = + bl_bin xs + 2*2^(length xs)" +apply(simp add: bin_wc_eq) +using bl2nat_cons_oc[of "xs @ [Bk]"] +apply(simp add: bl2nat_cons_bk bl2wc.simps) +done + +lemma tape_of_nat[simp]: "() = Oc\<^bsup>Suc a\<^esup>" +apply(simp add: tape_of_nat_abv) +done +lemma tape_of_nl_cons_app2: "() = ( @ Bk # Oc\<^bsup>Suc b\<^esup>)" +proof(induct "length xs" arbitrary: xs c, + simp add: tape_of_nl_abv tape_of_nat_list.simps) + fix x xs c + assume ind: "\xs c. x = length xs \ = + @ Bk # Oc\<^bsup>Suc b\<^esup>" + and h: "Suc x = length (xs::nat list)" + show " = @ Bk # Oc\<^bsup>Suc b\<^esup>" + proof(case_tac xs, simp add: tape_of_nl_abv tape_of_nat_list.simps) + fix a list + assume g: "xs = a # list" + hence k: " = @ Bk # Oc\<^bsup>Suc b\<^esup>" + apply(rule_tac ind) + using h + apply(simp) + done + from g and k show " = @ Bk # Oc\<^bsup>Suc b\<^esup>" + apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) + done + qed +qed + +lemma [simp]: "length () = Suc (Suc aa) + length ()" +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma [simp]: "bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista) @ [Bk, Oc]) = + bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)) + + 2* 2^(length (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)))" +using bl_bin_bk_oc[of "Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)"] +apply(simp) +done + +lemma [simp]: + "bl_bin () + (4 * rs + 4) * 2 ^ (length () - Suc 0) + = bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # ) + rs * (2 * 2 ^ (aa + length ()))" +apply(case_tac "list", simp add: add_mult_distrib, simp) +apply(simp add: tape_of_nl_cons_app2 add_mult_distrib) +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma tape_of_nl_app_Suc: "(()) = () @ [Oc]" +apply(induct list) +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind) +apply(case_tac list) +apply(simp_all add:tape_of_nl_abv tape_of_nat_list.simps exp_ind) +done + +lemma [simp]: "bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # @ [Oc]) + = bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # ) + + 2^(length (Oc # Oc\<^bsup>aa\<^esup> @ Bk # ))" +apply(simp add: bin_wc_eq) +apply(simp add: bl2nat_cons_oc bl2wc.simps) +using bl2nat_cons_oc[of "Oc # Oc\<^bsup>aa\<^esup> @ Bk # "] +apply(simp) +done +lemma [simp]: "bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # ) + (4 * 2 ^ (aa + length ()) + + 4 * (rs * 2 ^ (aa + length ()))) = + bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # ) + + rs * (2 * 2 ^ (aa + length ()))" +apply(simp add: tape_of_nl_app_Suc) +done + +declare tape_of_nat[simp del] + +text{* double case*} +fun wcode_double_case_inv :: "nat \ bin_inv_t" + where + "wcode_double_case_inv st ires rs (l, r) = + (if st = Suc 0 then wcode_on_left_moving_1 ires rs (l, r) + else if st = Suc (Suc 0) then wcode_on_checking_1 ires rs (l, r) + else if st = 3 then wcode_erase1 ires rs (l, r) + else if st = 4 then wcode_on_right_moving_1 ires rs (l, r) + else if st = 5 then wcode_goon_right_moving_1 ires rs (l, r) + else if st = 6 then wcode_backto_standard_pos ires rs (l, r) + else if st = 13 then wcode_before_double ires rs (l, r) + else False)" + +declare wcode_double_case_inv.simps[simp del] + +fun wcode_double_case_state :: "t_conf \ nat" + where + "wcode_double_case_state (st, l, r) = + 13 - st" + +fun wcode_double_case_step :: "t_conf \ nat" + where + "wcode_double_case_step (st, l, r) = + (if st = Suc 0 then (length l) + else if st = Suc (Suc 0) then (length r) + else if st = 3 then + if hd r = Oc then 1 else 0 + else if st = 4 then (length r) + else if st = 5 then (length r) + else if st = 6 then (length l) + else 0)" + +fun wcode_double_case_measure :: "t_conf \ nat \ nat" + where + "wcode_double_case_measure (st, l, r) = + (wcode_double_case_state (st, l, r), + wcode_double_case_step (st, l, r))" + +definition wcode_double_case_le :: "(t_conf \ t_conf) set" + where "wcode_double_case_le \ (inv_image lex_pair wcode_double_case_measure)" + +lemma [intro]: "wf lex_pair" +by(auto intro:wf_lex_prod simp:lex_pair_def) + +lemma wf_wcode_double_case_le[intro]: "wf wcode_double_case_le" +by(auto intro:wf_inv_image simp: wcode_double_case_le_def ) +term fetch + +lemma [simp]: "fetch t_wcode_main (Suc 0) Bk = (L, Suc 0)" +apply(simp add: t_wcode_main_def t_wcode_main_first_part_def + fetch.simps nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main (Suc 0) Oc = (L, Suc (Suc 0))" +apply(simp add: t_wcode_main_def t_wcode_main_first_part_def + fetch.simps nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main (Suc (Suc 0)) Oc = (R, 3)" +apply(simp add: t_wcode_main_def t_wcode_main_first_part_def + fetch.simps nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main (Suc (Suc (Suc 0))) Bk = (R, 4)" +apply(simp add: t_wcode_main_def t_wcode_main_first_part_def + fetch.simps nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main (Suc (Suc (Suc 0))) Oc = (W0, 3)" +apply(simp add: t_wcode_main_def t_wcode_main_first_part_def + fetch.simps nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 4 Bk = (R, 4)" +apply(simp add: t_wcode_main_def t_wcode_main_first_part_def + fetch.simps nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 4 Oc = (R, 5)" +apply(simp add: t_wcode_main_def t_wcode_main_first_part_def + fetch.simps nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 5 Oc = (R, 5)" +apply(simp add: t_wcode_main_def t_wcode_main_first_part_def + fetch.simps nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 5 Bk = (W1, 6)" +apply(simp add: t_wcode_main_def t_wcode_main_first_part_def + fetch.simps nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 6 Bk = (R, 13)" +apply(simp add: t_wcode_main_def t_wcode_main_first_part_def + fetch.simps nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 6 Oc = (L, 6)" +apply(simp add: t_wcode_main_def t_wcode_main_first_part_def + fetch.simps nth_of.simps) +done +lemma [elim]: "Bk\<^bsup>mr\<^esup> = [] \ mr = 0" +apply(case_tac mr, auto simp: exponent_def) +done + +lemma [simp]: "wcode_on_left_moving_1 ires rs (b, []) = False" +apply(simp add: wcode_on_left_moving_1.simps wcode_on_left_moving_1_B.simps + wcode_on_left_moving_1_O.simps, auto) +done + + +declare wcode_on_checking_1.simps[simp del] + +lemmas wcode_double_case_inv_simps = + wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps + wcode_on_left_moving_1_B.simps wcode_on_checking_1.simps + wcode_erase1.simps wcode_on_right_moving_1.simps + wcode_goon_right_moving_1.simps wcode_backto_standard_pos.simps + + +lemma [simp]: "wcode_on_left_moving_1 ires rs (b, r) \ b \ []" +apply(simp add: wcode_double_case_inv_simps, auto) +done + + +lemma [elim]: "\wcode_on_left_moving_1 ires rs (b, Bk # list); + tl b = aa \ hd b # Bk # list = ba\ \ + wcode_on_left_moving_1 ires rs (aa, ba)" +apply(simp only: wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps + wcode_on_left_moving_1_B.simps) +apply(erule_tac disjE) +apply(erule_tac exE)+ +apply(case_tac ml, simp) +apply(rule_tac x = "mr - Suc (Suc 0)" in exI, rule_tac x = rn in exI) +apply(case_tac mr, simp, case_tac nat, simp, simp add: exp_ind) +apply(rule_tac disjI1) +apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, + simp add: exp_ind_def) +apply(erule_tac exE)+ +apply(simp) +done + + +lemma [elim]: + "\wcode_on_left_moving_1 ires rs (b, Oc # list); tl b = aa \ hd b # Oc # list = ba\ + \ wcode_on_checking_1 ires rs (aa, ba)" +apply(simp only: wcode_double_case_inv_simps) +apply(erule_tac disjE) +apply(erule_tac [!] exE)+ +apply(case_tac mr, simp, simp add: exp_ind_def) +apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp) +done + + +lemma [simp]: "wcode_on_checking_1 ires rs (b, []) = False" +apply(auto simp: wcode_double_case_inv_simps) +done + +lemma [simp]: "wcode_on_checking_1 ires rs (b, Bk # list) = False" +apply(auto simp: wcode_double_case_inv_simps) +done + +lemma [elim]: "\wcode_on_checking_1 ires rs (b, Oc # ba);Oc # b = aa \ list = ba\ + \ wcode_erase1 ires rs (aa, ba)" +apply(simp only: wcode_double_case_inv_simps) +apply(erule_tac exE)+ +apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp) +done + + +lemma [simp]: "wcode_on_checking_1 ires rs (b, []) = False" +apply(simp add: wcode_double_case_inv_simps) +done + +lemma [simp]: "wcode_on_checking_1 ires rs ([], Bk # list) = False" +apply(simp add: wcode_double_case_inv_simps) +done + +lemma [simp]: "wcode_erase1 ires rs (b, []) = False" +apply(simp add: wcode_double_case_inv_simps) +done + +lemma [simp]: "wcode_on_right_moving_1 ires rs (b, []) = False" +apply(simp add: wcode_double_case_inv_simps exp_ind_def) +done + +lemma [simp]: "wcode_on_right_moving_1 ires rs (b, []) = False" +apply(simp add: wcode_double_case_inv_simps exp_ind_def) +done + +lemma [elim]: "\wcode_on_right_moving_1 ires rs (b, Bk # ba); Bk # b = aa \ list = b\ \ + wcode_on_right_moving_1 ires rs (aa, ba)" +apply(simp only: wcode_double_case_inv_simps) +apply(erule_tac exE)+ +apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI, + rule_tac x = rn in exI) +apply(simp add: exp_ind_def) +apply(case_tac mr, simp, simp add: exp_ind_def) +done + +lemma [elim]: + "\wcode_on_right_moving_1 ires rs (b, Oc # ba); Oc # b = aa \ list = ba\ + \ wcode_goon_right_moving_1 ires rs (aa, ba)" +apply(simp only: wcode_double_case_inv_simps) +apply(erule_tac exE)+ +apply(rule_tac x = "Suc 0" in exI, rule_tac x = "rs" in exI, + rule_tac x = "ml - Suc (Suc 0)" in exI, rule_tac x = rn in exI) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(case_tac ml, simp, case_tac nat, simp, simp) +apply(simp add: exp_ind_def) +done + +lemma [simp]: + "wcode_on_right_moving_1 ires rs (b, []) \ False" +apply(simp add: wcode_double_case_inv_simps exponent_def) +done + +lemma [elim]: "\wcode_erase1 ires rs (b, Bk # ba); Bk # b = aa \ list = ba; c = Bk # ba\ + \ wcode_on_right_moving_1 ires rs (aa, ba)" +apply(simp only: wcode_double_case_inv_simps) +apply(erule_tac exE)+ +apply(rule_tac x = "Suc 0" in exI, rule_tac x = "Suc (Suc ln)" in exI, + rule_tac x = rn in exI, simp add: exp_ind) +done + +lemma [elim]: "\wcode_erase1 ires rs (aa, Oc # list); b = aa \ Bk # list = ba\ \ + wcode_erase1 ires rs (aa, ba)" +apply(simp only: wcode_double_case_inv_simps) +apply(erule_tac exE)+ +apply(rule_tac x = ln in exI, rule_tac x = rn in exI, auto) +done + +lemma [elim]: "\wcode_goon_right_moving_1 ires rs (aa, []); b = aa \ [Oc] = ba\ + \ wcode_backto_standard_pos ires rs (aa, ba)" +apply(simp only: wcode_double_case_inv_simps) +apply(erule_tac exE)+ +apply(rule_tac disjI2) +apply(simp only:wcode_backto_standard_pos_O.simps) +apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI, + rule_tac x = rn in exI, simp) +apply(case_tac mr, simp_all add: exponent_def) +done + +lemma [elim]: + "\wcode_goon_right_moving_1 ires rs (aa, Bk # list); b = aa \ Oc # list = ba\ + \ wcode_backto_standard_pos ires rs (aa, ba)" +apply(simp only: wcode_double_case_inv_simps) +apply(erule_tac exE)+ +apply(rule_tac disjI2) +apply(simp only:wcode_backto_standard_pos_O.simps) +apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI, + rule_tac x = "rn - Suc 0" in exI, simp) +apply(case_tac mr, simp, case_tac rn, simp, simp_all add: exp_ind_def) +done + +lemma [elim]: "\wcode_goon_right_moving_1 ires rs (b, Oc # ba); Oc # b = aa \ list = ba\ + \ wcode_goon_right_moving_1 ires rs (aa, ba)" +apply(simp only: wcode_double_case_inv_simps) +apply(erule_tac exE)+ +apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI, + rule_tac x = ln in exI, rule_tac x = rn in exI) +apply(simp add: exp_ind_def) +apply(case_tac mr, simp, case_tac rn, simp_all add: exp_ind_def) +done + +lemma [elim]: "\wcode_backto_standard_pos ires rs (b, []); Bk # b = aa\ \ False" +apply(auto simp: wcode_double_case_inv_simps wcode_backto_standard_pos_O.simps + wcode_backto_standard_pos_B.simps) +apply(case_tac mr, simp_all add: exp_ind_def) +done + +lemma [elim]: "\wcode_backto_standard_pos ires rs (b, Bk # ba); Bk # b = aa \ list = ba\ + \ wcode_before_double ires rs (aa, ba)" +apply(simp only: wcode_double_case_inv_simps wcode_backto_standard_pos_B.simps + wcode_backto_standard_pos_O.simps wcode_before_double.simps) +apply(erule_tac disjE) +apply(erule_tac exE)+ +apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp) +apply(auto) +apply(case_tac [!] mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "wcode_backto_standard_pos ires rs ([], Oc # list) = False" +apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps + wcode_backto_standard_pos_O.simps) +done + +lemma [simp]: "wcode_backto_standard_pos ires rs (b, []) = False" +apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps + wcode_backto_standard_pos_O.simps) +apply(case_tac mr, simp, simp add: exp_ind_def) +done + +lemma [elim]: "\wcode_backto_standard_pos ires rs (b, Oc # list); tl b = aa; hd b # Oc # list = ba\ + \ wcode_backto_standard_pos ires rs (aa, ba)" +apply(simp only: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps + wcode_backto_standard_pos_O.simps) +apply(erule_tac disjE) +apply(simp) +apply(erule_tac exE)+ +apply(case_tac ml, simp) +apply(rule_tac disjI1, rule_tac conjI) +apply(rule_tac x = ln in exI, simp, rule_tac x = rn in exI, simp) +apply(rule_tac disjI2) +apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = ln in exI, + rule_tac x = rn in exI, simp) +apply(simp add: exp_ind_def) +done + +declare new_tape.simps[simp del] nth_of.simps[simp del] fetch.simps[simp del] +lemma wcode_double_case_first_correctness: + "let P = (\ (st, l, r). st = 13) in + let Q = (\ (st, l, r). wcode_double_case_inv st ires rs (l, r)) in + let f = (\ stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp) in + \ n .P (f n) \ Q (f (n::nat))" +proof - + let ?P = "(\ (st, l, r). st = 13)" + let ?Q = "(\ (st, l, r). wcode_double_case_inv st ires rs (l, r))" + let ?f = "(\ stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp)" + have "\ n. ?P (?f n) \ ?Q (?f (n::nat))" + proof(rule_tac halt_lemma2) + show "wf wcode_double_case_le" + by auto + next + show "\ na. \ ?P (?f na) \ ?Q (?f na) \ + ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_double_case_le" + proof(rule_tac allI, case_tac "?f na", simp add: tstep_red) + fix na a b c + show "a \ 13 \ wcode_double_case_inv a ires rs (b, c) \ + (case tstep (a, b, c) t_wcode_main of (st, x) \ + wcode_double_case_inv st ires rs x) \ + (tstep (a, b, c) t_wcode_main, a, b, c) \ wcode_double_case_le" + apply(rule_tac impI, simp add: wcode_double_case_inv.simps) + apply(auto split: if_splits simp: tstep.simps, + case_tac [!] c, simp_all, case_tac [!] "(c::block list)!0") + apply(simp_all add: new_tape.simps wcode_double_case_inv.simps wcode_double_case_le_def + lex_pair_def) + apply(auto split: if_splits) + done + qed + next + show "?Q (?f 0)" + apply(simp add: steps.simps wcode_double_case_inv.simps + wcode_on_left_moving_1.simps + wcode_on_left_moving_1_B.simps) + apply(rule_tac disjI1) + apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def) + apply(rule_tac x = "Suc 0" in exI, simp add: exp_ind_def) + apply(auto) + done + next + show "\ ?P (?f 0)" + apply(simp add: steps.simps) + done + qed + thus "let P = \(st, l, r). st = 13; + Q = \(st, l, r). wcode_double_case_inv st ires rs (l, r); + f = steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main + in \n. P (f n) \ Q (f n)" + apply(simp add: Let_def) + done +qed + +lemma [elim]: "t_ncorrect tp + \ t_ncorrect (abacus.tshift tp a)" +apply(simp add: t_ncorrect.simps shift_length) +done + +lemma tshift_fetch: "\ fetch tp a b = (aa, st'); 0 < st'\ + \ fetch (abacus.tshift tp (length tp1 div 2)) a b + = (aa, st' + length tp1 div 2)" +apply(subgoal_tac "a > 0") +apply(auto simp: fetch.simps nth_of.simps shift_length nth_map + tshift.simps split: block.splits if_splits) +done + +lemma t_steps_steps_eq: "\steps (st, l, r) tp stp = (st', l', r'); + 0 < st'; + 0 < st \ st \ length tp div 2; + t_ncorrect tp1; + t_ncorrect tp\ + \ t_steps (st + length tp1 div 2, l, r) (tshift tp (length tp1 div 2), + length tp1 div 2) stp + = (st' + length tp1 div 2, l', r')" +apply(induct stp arbitrary: st' l' r', simp add: steps.simps t_steps.simps, + simp add: tstep_red stepn) +apply(case_tac "(steps (st, l, r) tp stp)", simp) +proof - + fix stp st' l' r' a b c + assume ind: "\st' l' r'. + \a = st' \ b = l' \ c = r'; 0 < st'\ + \ t_steps (st + length tp1 div 2, l, r) + (abacus.tshift tp (length tp1 div 2), length tp1 div 2) stp = + (st' + length tp1 div 2, l', r')" + and h: "tstep (a, b, c) tp = (st', l', r')" "0 < st'" "t_ncorrect tp1" "t_ncorrect tp" + have k: "t_steps (st + length tp1 div 2, l, r) (abacus.tshift tp (length tp1 div 2), + length tp1 div 2) stp = (a + length tp1 div 2, b, c)" + apply(rule_tac ind, simp) + using h + apply(case_tac a, simp_all add: tstep.simps fetch.simps) + done + from h and this show "t_step (t_steps (st + length tp1 div 2, l, r) (abacus.tshift tp (length tp1 div 2), length tp1 div 2) stp) + (abacus.tshift tp (length tp1 div 2), length tp1 div 2) = + (st' + length tp1 div 2, l', r')" + apply(simp add: k) + apply(simp add: tstep.simps t_step.simps) + apply(case_tac "fetch tp a (case c of [] \ Bk | x # xs \ x)", simp) + apply(subgoal_tac "fetch (abacus.tshift tp (length tp1 div 2)) a + (case c of [] \ Bk | x # xs \ x) = (aa, st' + length tp1 div 2)", simp) + apply(simp add: tshift_fetch) + done +qed + +lemma t_tshift_lemma: "\ steps (st, l, r) tp stp = (st', l', r'); + st' \ 0; + stp > 0; + 0 < st \ st \ length tp div 2; + t_ncorrect tp1; + t_ncorrect tp; + t_ncorrect tp2 + \ + \ \ stp>0. steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2) stp + = (st' + length tp1 div 2, l', r')" +proof - + assume h: "steps (st, l, r) tp stp = (st', l', r')" + "st' \ 0" "stp > 0" + "0 < st \ st \ length tp div 2" + "t_ncorrect tp1" + "t_ncorrect tp" + "t_ncorrect tp2" + from h have + "\stp>0. t_steps (st + length tp1 div 2, l, r) (tp1 @ abacus.tshift tp (length tp1 div 2) @ tp2, 0) stp = + (st' + length tp1 div 2, l', r')" + apply(rule_tac stp = stp in turing_shift, simp_all add: shift_length) + apply(simp add: t_steps_steps_eq) + apply(simp add: t_ncorrect.simps shift_length) + done + thus "\ stp>0. steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2) stp + = (st' + length tp1 div 2, l', r')" + apply(erule_tac exE) + apply(rule_tac x = stp in exI, simp) + apply(subgoal_tac "length (tp1 @ abacus.tshift tp (length tp1 div 2) @ tp2) mod 2 = 0") + apply(simp only: steps_eq) + using h + apply(auto simp: t_ncorrect.simps shift_length) + apply arith + done +qed + + +lemma t_twice_len_ge: "Suc 0 \ length t_twice div 2" +apply(simp add: t_twice_def tMp.simps shift_length) +done + +lemma [intro]: "rec_calc_rel (recf.id (Suc 0) 0) [rs] rs" + apply(rule_tac calc_id, simp_all) + done + +lemma [intro]: "rec_calc_rel (constn 2) [rs] 2" +using prime_rel_exec_eq[of "constn 2" "[rs]" 2] +apply(subgoal_tac "primerec (constn 2) 1", auto) +done + +lemma [intro]: "rec_calc_rel rec_mult [rs, 2] (2 * rs)" +using prime_rel_exec_eq[of "rec_mult" "[rs, 2]" "2*rs"] +apply(subgoal_tac "primerec rec_mult (Suc (Suc 0))", auto) +done +lemma t_twice_correct: "\stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) stp = + (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +proof(case_tac "rec_ci rec_twice") + fix a b c + assume h: "rec_ci rec_twice = (a, b, c)" + have "\stp m l. steps (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<^bsup>n\<^esup>) (tm_of abc_twice @ tMp (Suc 0) + (start_of twice_ly (length abc_twice) - 1)) stp = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2*rs)\<^esup> @ Bk\<^bsup>l\<^esup>)" + proof(rule_tac t_compiled_by_rec) + show "rec_ci rec_twice = (a, b, c)" by (simp add: h) + next + show "rec_calc_rel rec_twice [rs] (2 * rs)" + apply(simp add: rec_twice_def) + apply(rule_tac rs = "[rs, 2]" in calc_cn, simp_all) + apply(rule_tac allI, case_tac k, auto) + done + next + show "length [rs] = Suc 0" by simp + next + show "layout_of (a [+] dummy_abc (Suc 0)) = layout_of (a [+] dummy_abc (Suc 0))" + by simp + next + show "start_of twice_ly (length abc_twice) = + start_of (layout_of (a [+] dummy_abc (Suc 0))) (length (a [+] dummy_abc (Suc 0)))" + using h + apply(simp add: twice_ly_def abc_twice_def) + done + next + show "tm_of abc_twice = tm_of (a [+] dummy_abc (Suc 0))" + using h + apply(simp add: abc_twice_def) + done + qed + thus "\stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) stp = + (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) + done +qed + +lemma change_termi_state_fetch: "\fetch ap a b = (aa, st); st > 0\ + \ fetch (change_termi_state ap) a b = (aa, st)" +apply(case_tac b, auto simp: fetch.simps nth_of.simps change_termi_state.simps nth_map + split: if_splits block.splits) +done + +lemma change_termi_state_exec_in_range: + "\steps (st, l, r) ap stp = (st', l', r'); st' \ 0\ + \ steps (st, l, r) (change_termi_state ap) stp = (st', l', r')" +proof(induct stp arbitrary: st l r st' l' r', simp add: steps.simps) + fix stp st l r st' l' r' + assume ind: "\st l r st' l' r'. + \steps (st, l, r) ap stp = (st', l', r'); st' \ 0\ \ + steps (st, l, r) (change_termi_state ap) stp = (st', l', r')" + and h: "steps (st, l, r) ap (Suc stp) = (st', l', r')" "st' \ 0" + from h show "steps (st, l, r) (change_termi_state ap) (Suc stp) = (st', l', r')" + proof(simp add: tstep_red, case_tac "steps (st, l, r) ap stp", simp) + fix a b c + assume g: "steps (st, l, r) ap stp = (a, b, c)" + "tstep (a, b, c) ap = (st', l', r')" "0 < st'" + hence "steps (st, l, r) (change_termi_state ap) stp = (a, b, c)" + apply(rule_tac ind, simp) + apply(case_tac a, simp_all add: tstep_0) + done + from g and this show "tstep (steps (st, l, r) (change_termi_state ap) stp) + (change_termi_state ap) = (st', l', r')" + apply(simp add: tstep.simps) + apply(case_tac "fetch ap a (case c of [] \ Bk | x # xs \ x)", simp) + apply(subgoal_tac "fetch (change_termi_state ap) a (case c of [] \ Bk | x # xs \ x) + = (aa, st')", simp) + apply(simp add: change_termi_state_fetch) + done + qed +qed + +lemma change_termi_state_fetch0: + "\0 < a; a \ length ap div 2; t_correct ap; fetch ap a b = (aa, 0)\ + \ fetch (change_termi_state ap) a b = (aa, Suc (length ap div 2))" +apply(case_tac b, auto simp: fetch.simps nth_of.simps change_termi_state.simps nth_map + split: if_splits block.splits) +done + +lemma turing_change_termi_state: + "\steps (Suc 0, l, r) ap stp = (0, l', r'); t_correct ap\ + \ \ stp. steps (Suc 0, l, r) (change_termi_state ap) stp = + (Suc (length ap div 2), l', r')" +apply(drule first_halt_point) +apply(erule_tac exE) +apply(rule_tac x = "Suc stp" in exI, simp add: tstep_red) +apply(case_tac "steps (Suc 0, l, r) ap stp") +apply(simp add: isS0_def change_termi_state_exec_in_range) +apply(subgoal_tac "steps (Suc 0, l, r) (change_termi_state ap) stp = (a, b, c)", simp) +apply(simp add: tstep.simps) +apply(case_tac "fetch ap a (case c of [] \ Bk | x # xs \ x)", simp) +apply(subgoal_tac "fetch (change_termi_state ap) a + (case c of [] \ Bk | x # xs \ x) = (aa, Suc (length ap div 2))", simp) +apply(rule_tac ap = ap in change_termi_state_fetch0, simp_all) +apply(rule_tac tp = "(l, r)" and l = b and r = c and stp = stp and A = ap in s_keep, simp_all) +apply(simp add: change_termi_state_exec_in_range) +done + +lemma t_twice_change_term_state: + "\ stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp + = (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +using t_twice_correct[of ires rs n] +apply(erule_tac exE) +apply(erule_tac exE) +apply(erule_tac exE) +proof(drule_tac turing_change_termi_state) + fix stp ln rn + show "t_correct (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0))" + apply(rule_tac t_compiled_correct, simp_all) + apply(simp add: twice_ly_def) + done +next + fix stp ln rn + show "\stp. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + (change_termi_state (tm_of abc_twice @ tMp (Suc 0) + (start_of twice_ly (length abc_twice) - Suc 0))) stp = + (Suc (length (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) div 2), + Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) \ + \stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp = + (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(erule_tac exE) + apply(simp add: t_twice_len_def t_twice_def) + apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp) + done +qed + +lemma t_twice_append_pre: + "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp + = (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) + \ \ stp>0. steps (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + (t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @ + ([(L, 1), (L, 1)] @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp + = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +proof(rule_tac t_tshift_lemma, simp_all add: t_twice_len_ge) + assume "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp = + (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + thus "0 < stp" + apply(case_tac stp, simp add: steps.simps t_twice_len_ge t_twice_len_def) + using t_twice_len_ge + apply(simp, simp) + done +next + show "t_ncorrect t_wcode_main_first_part" + apply(simp add: t_ncorrect.simps t_wcode_main_first_part_def) + done +next + show "t_ncorrect t_twice" + using length_tm_even[of abc_twice] + apply(auto simp: t_ncorrect.simps t_twice_def) + apply(arith) + done +next + show "t_ncorrect ((L, Suc 0) # (L, Suc 0) # + abacus.tshift t_fourtimes (t_twice_len + 13) @ [(L, Suc 0), (L, Suc 0)])" + using length_tm_even[of abc_fourtimes] + apply(simp add: t_ncorrect.simps shift_length t_fourtimes_def) + apply arith + done +qed + +lemma t_twice_append: + "\ stp ln rn. steps (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + (t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @ + ([(L, 1), (L, 1)] @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp + = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + using t_twice_change_term_state[of ires rs n] + apply(erule_tac exE) + apply(erule_tac exE) + apply(erule_tac exE) + apply(drule_tac t_twice_append_pre) + apply(erule_tac exE) + apply(rule_tac x = stpa in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) + apply(simp) + done + +lemma [simp]: "fetch t_wcode_main (Suc (t_twice_len + length t_wcode_main_first_part div 2)) Oc + = (L, Suc 0)" +apply(subgoal_tac "length (t_twice) mod 2 = 0") +apply(simp add: t_wcode_main_def nth_append fetch.simps t_wcode_main_first_part_def + nth_of.simps shift_length t_twice_len_def, auto) +apply(simp add: t_twice_def) +apply(subgoal_tac "length (tm_of abc_twice) mod 2 = 0") +apply arith +apply(rule_tac tm_even) +done + +lemma wcode_jump1: + "\ stp ln rn. steps (Suc (t_twice_len) + length t_wcode_main_first_part div 2, + Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>n\<^esup>) + t_wcode_main stp + = (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +apply(rule_tac x = "Suc 0" in exI, rule_tac x = "m" in exI, rule_tac x = n in exI) +apply(simp add: steps.simps tstep.simps exp_ind_def new_tape.simps) +apply(case_tac m, simp, simp add: exp_ind_def) +apply(simp add: exp_ind_def[THEN sym] exp_ind[THEN sym]) +done + +lemma wcode_main_first_part_len: + "length t_wcode_main_first_part = 24" + apply(simp add: t_wcode_main_first_part_def) + done + +lemma wcode_double_case: + shows "\stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +proof - + have "\stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (13, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + using wcode_double_case_first_correctness[of ires rs m n] + apply(simp) + apply(erule_tac exE) + apply(case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, + Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na", + auto simp: wcode_double_case_inv.simps + wcode_before_double.simps) + apply(rule_tac x = na in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) + apply(simp) + done + from this obtain stpa lna rna where stp1: + "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa = + (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast + have "\ stp ln rn. steps (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp = + (13 + t_twice_len, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rn\<^esup>)" + using t_twice_append[of "Bk\<^bsup>lna\<^esup> @ Oc # ires" "Suc rs" rna] + apply(erule_tac exE) + apply(erule_tac exE) + apply(erule_tac exE) + apply(simp add: wcode_main_first_part_len) + apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI, + rule_tac x = rn in exI) + apply(simp add: t_wcode_main_def) + apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym]) + done + from this obtain stpb lnb rnb where stp2: + "steps (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb = + (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>)" by blast + have "\stp ln rn. steps (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, + Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stp = + (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rn\<^esup>)" + using wcode_jump1[of lnb "Oc # ires" "Suc rs" rnb] + apply(erule_tac exE) + apply(erule_tac exE) + apply(erule_tac exE) + apply(rule_tac x = stp in exI, + rule_tac x = ln in exI, + rule_tac x = rn in exI, simp add:wcode_main_first_part_len t_wcode_main_def) + apply(subgoal_tac "Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc # ires = Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires", simp) + apply(simp add: exp_ind_def[THEN sym] exp_ind[THEN sym]) + apply(simp) + apply(case_tac lnb, simp, simp add: exp_ind_def[THEN sym] exp_ind) + done + from this obtain stpc lnc rnc where stp3: + "steps (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, + Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stpc = + (Suc 0, Bk # Bk\<^bsup>lnc\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnc\<^esup>)" + by blast + from stp1 stp2 stp3 show "?thesis" + apply(rule_tac x = "stpa + stpb + stpc" in exI, rule_tac x = lnc in exI, + rule_tac x = rnc in exI) + apply(simp add: steps_add) + done +qed + + +(* Begin: fourtime_case*) +fun wcode_on_left_moving_2_B :: "bin_inv_t" + where + "wcode_on_left_moving_2_B ires rs (l, r) = + (\ ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Bk # Oc # ires \ + r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr > Suc 0 \ mr > 0)" + +fun wcode_on_left_moving_2_O :: "bin_inv_t" + where + "wcode_on_left_moving_2_O ires rs (l, r) = + (\ ln rn. l = Bk # Oc # ires \ + r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wcode_on_left_moving_2 :: "bin_inv_t" + where + "wcode_on_left_moving_2 ires rs (l, r) = + (wcode_on_left_moving_2_B ires rs (l, r) \ + wcode_on_left_moving_2_O ires rs (l, r))" + +fun wcode_on_checking_2 :: "bin_inv_t" + where + "wcode_on_checking_2 ires rs (l, r) = + (\ ln rn. l = Oc#ires \ + r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wcode_goon_checking :: "bin_inv_t" + where + "wcode_goon_checking ires rs (l, r) = + (\ ln rn. l = ires \ + r = Oc # Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wcode_right_move :: "bin_inv_t" + where + "wcode_right_move ires rs (l, r) = + (\ ln rn. l = Oc # ires \ + r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wcode_erase2 :: "bin_inv_t" + where + "wcode_erase2 ires rs (l, r) = + (\ ln rn. l = Bk # Oc # ires \ + tl r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wcode_on_right_moving_2 :: "bin_inv_t" + where + "wcode_on_right_moving_2 ires rs (l, r) = + (\ ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # ires \ + r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \ ml + mr > Suc 0)" + +fun wcode_goon_right_moving_2 :: "bin_inv_t" + where + "wcode_goon_right_moving_2 ires rs (l, r) = + (\ ml mr ln rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \ + r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ ml + mr = Suc rs)" + +fun wcode_backto_standard_pos_2_B :: "bin_inv_t" + where + "wcode_backto_standard_pos_2_B ires rs (l, r) = + (\ ln rn. l = Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \ + r = Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wcode_backto_standard_pos_2_O :: "bin_inv_t" + where + "wcode_backto_standard_pos_2_O ires rs (l, r) = + (\ ml mr ln rn. l = Oc\<^bsup>ml \<^esup>@ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \ + r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr = (Suc (Suc rs)) \ mr > 0)" + +fun wcode_backto_standard_pos_2 :: "bin_inv_t" + where + "wcode_backto_standard_pos_2 ires rs (l, r) = + (wcode_backto_standard_pos_2_O ires rs (l, r) \ + wcode_backto_standard_pos_2_B ires rs (l, r))" + +fun wcode_before_fourtimes :: "bin_inv_t" + where + "wcode_before_fourtimes ires rs (l, r) = + (\ ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \ + r = Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +declare wcode_on_left_moving_2_B.simps[simp del] wcode_on_left_moving_2.simps[simp del] + wcode_on_left_moving_2_O.simps[simp del] wcode_on_checking_2.simps[simp del] + wcode_goon_checking.simps[simp del] wcode_right_move.simps[simp del] + wcode_erase2.simps[simp del] + wcode_on_right_moving_2.simps[simp del] wcode_goon_right_moving_2.simps[simp del] + wcode_backto_standard_pos_2_B.simps[simp del] wcode_backto_standard_pos_2_O.simps[simp del] + wcode_backto_standard_pos_2.simps[simp del] + +lemmas wcode_fourtimes_invs = + wcode_on_left_moving_2_B.simps wcode_on_left_moving_2.simps + wcode_on_left_moving_2_O.simps wcode_on_checking_2.simps + wcode_goon_checking.simps wcode_right_move.simps + wcode_erase2.simps + wcode_on_right_moving_2.simps wcode_goon_right_moving_2.simps + wcode_backto_standard_pos_2_B.simps wcode_backto_standard_pos_2_O.simps + wcode_backto_standard_pos_2.simps + +fun wcode_fourtimes_case_inv :: "nat \ bin_inv_t" + where + "wcode_fourtimes_case_inv st ires rs (l, r) = + (if st = Suc 0 then wcode_on_left_moving_2 ires rs (l, r) + else if st = Suc (Suc 0) then wcode_on_checking_2 ires rs (l, r) + else if st = 7 then wcode_goon_checking ires rs (l, r) + else if st = 8 then wcode_right_move ires rs (l, r) + else if st = 9 then wcode_erase2 ires rs (l, r) + else if st = 10 then wcode_on_right_moving_2 ires rs (l, r) + else if st = 11 then wcode_goon_right_moving_2 ires rs (l, r) + else if st = 12 then wcode_backto_standard_pos_2 ires rs (l, r) + else if st = t_twice_len + 14 then wcode_before_fourtimes ires rs (l, r) + else False)" + +declare wcode_fourtimes_case_inv.simps[simp del] + +fun wcode_fourtimes_case_state :: "t_conf \ nat" + where + "wcode_fourtimes_case_state (st, l, r) = 13 - st" + +fun wcode_fourtimes_case_step :: "t_conf \ nat" + where + "wcode_fourtimes_case_step (st, l, r) = + (if st = Suc 0 then length l + else if st = 9 then + (if hd r = Oc then 1 + else 0) + else if st = 10 then length r + else if st = 11 then length r + else if st = 12 then length l + else 0)" + +fun wcode_fourtimes_case_measure :: "t_conf \ nat \ nat" + where + "wcode_fourtimes_case_measure (st, l, r) = + (wcode_fourtimes_case_state (st, l, r), + wcode_fourtimes_case_step (st, l, r))" + +definition wcode_fourtimes_case_le :: "(t_conf \ t_conf) set" + where "wcode_fourtimes_case_le \ (inv_image lex_pair wcode_fourtimes_case_measure)" + +lemma wf_wcode_fourtimes_case_le[intro]: "wf wcode_fourtimes_case_le" +by(auto intro:wf_inv_image simp: wcode_fourtimes_case_le_def) + +lemma [simp]: "fetch t_wcode_main (Suc (Suc 0)) Bk = (L, 7)" +apply(simp add: t_wcode_main_def fetch.simps + t_wcode_main_first_part_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 7 Oc = (R, 8)" +apply(simp add: t_wcode_main_def fetch.simps + t_wcode_main_first_part_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 8 Bk = (R, 9)" +apply(simp add: t_wcode_main_def fetch.simps + t_wcode_main_first_part_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 9 Bk = (R, 10)" +apply(simp add: t_wcode_main_def fetch.simps + t_wcode_main_first_part_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 9 Oc = (W0, 9)" +apply(simp add: t_wcode_main_def fetch.simps + t_wcode_main_first_part_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 10 Bk = (R, 10)" +apply(simp add: t_wcode_main_def fetch.simps + t_wcode_main_first_part_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 10 Oc = (R, 11)" +apply(simp add: t_wcode_main_def fetch.simps + t_wcode_main_first_part_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 11 Bk = (W1, 12)" +apply(simp add: t_wcode_main_def fetch.simps + t_wcode_main_first_part_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 11 Oc = (R, 11)" +apply(simp add: t_wcode_main_def fetch.simps + t_wcode_main_first_part_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 12 Oc = (L, 12)" +apply(simp add: t_wcode_main_def fetch.simps + t_wcode_main_first_part_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_main 12 Bk = (R, t_twice_len + 14)" +apply(simp add: t_wcode_main_def fetch.simps + t_wcode_main_first_part_def nth_of.simps) +done + + +lemma [simp]: "wcode_on_left_moving_2 ires rs (b, []) = False" +apply(auto simp: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_on_checking_2 ires rs (b, []) = False" +apply(auto simp: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_goon_checking ires rs (b, []) = False" +apply(auto simp: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_right_move ires rs (b, []) = False" +apply(auto simp: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_erase2 ires rs (b, []) = False" +apply(auto simp: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_on_right_moving_2 ires rs (b, []) = False" +apply(auto simp: wcode_fourtimes_invs exponent_def) +done + +lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, []) = False" +apply(auto simp: wcode_fourtimes_invs exponent_def) +done + +lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \ b \ []" +apply(simp add: wcode_fourtimes_invs, auto) +done + +lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \ wcode_on_left_moving_2 ires rs (tl b, hd b # Bk # list)" +apply(simp only: wcode_fourtimes_invs) +apply(erule_tac disjE) +apply(erule_tac exE)+ +apply(case_tac ml, simp) +apply(rule_tac x = "mr - (Suc (Suc 0))" in exI, rule_tac x = rn in exI, simp) +apply(case_tac mr, simp, case_tac nat, simp, simp add: exp_ind) +apply(rule_tac disjI1) +apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, + simp add: exp_ind_def) +apply(simp) +done + +lemma [simp]: "wcode_on_checking_2 ires rs (b, Bk # list) \ b \ []" +apply(auto simp: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_on_checking_2 ires rs (b, Bk # list) + \ wcode_goon_checking ires rs (tl b, hd b # Bk # list)" +apply(simp only: wcode_fourtimes_invs) +apply(auto) +done + +lemma [simp]: "wcode_goon_checking ires rs (b, Bk # list) = False" +apply(simp add: wcode_fourtimes_invs) +done + +lemma [simp]: " wcode_right_move ires rs (b, Bk # list) \ b\ []" +apply(simp add: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_right_move ires rs (b, Bk # list) \ wcode_erase2 ires rs (Bk # b, list)" +apply(auto simp:wcode_fourtimes_invs ) +apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp) +done + +lemma [simp]: "wcode_erase2 ires rs (b, Bk # list) \ b \ []" +apply(auto simp: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_erase2 ires rs (b, Bk # list) \ wcode_on_right_moving_2 ires rs (Bk # b, list)" +apply(auto simp:wcode_fourtimes_invs ) +apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: exp_ind) +apply(rule_tac x = "Suc (Suc ln)" in exI, simp add: exp_ind, auto) +done + +lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list) \ b \ []" +apply(auto simp:wcode_fourtimes_invs ) +done + +lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list) + \ wcode_on_right_moving_2 ires rs (Bk # b, list)" +apply(auto simp: wcode_fourtimes_invs) +apply(rule_tac x = "Suc ml" in exI, simp add: exp_ind_def) +apply(rule_tac x = "mr - 1" in exI, case_tac mr, auto simp: exp_ind_def) +done + +lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \ b \ []" +apply(auto simp: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \ + wcode_backto_standard_pos_2 ires rs (b, Oc # list)" +apply(simp add: wcode_fourtimes_invs, auto) +apply(rule_tac x = ml in exI, auto) +apply(rule_tac x = "Suc 0" in exI, simp) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(rule_tac x = "rn - 1" in exI, simp) +apply(case_tac rn, simp, simp add: exp_ind_def) +done + +lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list) \ b \ []" +apply(simp add: wcode_fourtimes_invs, auto) +done + +lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Oc # list) \ b \ []" +apply(simp add: wcode_fourtimes_invs, auto) +done + +lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Oc # list) \ + wcode_on_checking_2 ires rs (tl b, hd b # Oc # list)" +apply(auto simp: wcode_fourtimes_invs) +apply(case_tac [!] mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, []) \ b \ []" +apply(auto simp: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, []) \ + wcode_backto_standard_pos_2 ires rs (b, [Oc])" +apply(simp only: wcode_fourtimes_invs) +apply(erule_tac exE)+ +apply(rule_tac disjI1) +apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, + rule_tac x = ln in exI, rule_tac x = rn in exI, simp) +apply(case_tac mr, simp, simp add: exp_ind_def) +done + +lemma "wcode_backto_standard_pos_2 ires rs (b, Bk # list) + \ (\ln. b = Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires) \ (\rn. list = Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +apply(auto simp: wcode_fourtimes_invs) +apply(case_tac [!] mr, auto simp: exp_ind_def) +done + + +lemma [simp]: "wcode_on_checking_2 ires rs (b, Oc # list) \ False" +apply(simp add: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_goon_checking ires rs (b, Oc # list) \ + (b = [] \ wcode_right_move ires rs ([Oc], list)) \ + (b \ [] \ wcode_right_move ires rs (Oc # b, list))" +apply(simp only: wcode_fourtimes_invs) +apply(erule_tac exE)+ +apply(auto) +done + +lemma [simp]: "wcode_right_move ires rs (b, Oc # list) = False" +apply(auto simp: wcode_fourtimes_invs) +done + +lemma [simp]: " wcode_erase2 ires rs (b, Oc # list) \ b \ []" +apply(simp add: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_erase2 ires rs (b, Oc # list) + \ wcode_erase2 ires rs (b, Bk # list)" +apply(auto simp: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Oc # list) \ b \ []" +apply(simp only: wcode_fourtimes_invs) +apply(auto) +done + +lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Oc # list) + \ wcode_goon_right_moving_2 ires rs (Oc # b, list)" +apply(auto simp: wcode_fourtimes_invs) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(rule_tac x = "Suc 0" in exI, auto) +apply(rule_tac x = "ml - 2" in exI) +apply(case_tac ml, simp, case_tac nat, simp_all add: exp_ind_def) +done + +lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \ b \ []" +apply(simp only:wcode_fourtimes_invs, auto) +done + +lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list) + \ (\ln. b = Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires) \ (\rn. list = Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +apply(simp add: wcode_fourtimes_invs, auto) +apply(case_tac [!] mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "wcode_on_checking_2 ires rs (b, Oc # list) = False" +apply(simp add: wcode_fourtimes_invs) +done + +lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \ + wcode_goon_right_moving_2 ires rs (Oc # b, list)" +apply(simp only:wcode_fourtimes_invs, auto) +apply(rule_tac x = "Suc ml" in exI, auto simp: exp_ind_def) +apply(rule_tac x = "mr - 1" in exI) +apply(case_tac mr, case_tac rn, auto simp: exp_ind_def) +done + +lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list) \ b \ []" +apply(simp only: wcode_fourtimes_invs, auto) +done + +lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list) + \ wcode_backto_standard_pos_2 ires rs (tl b, hd b # Oc # list)" +apply(simp only: wcode_fourtimes_invs) +apply(erule_tac disjE) +apply(erule_tac exE)+ +apply(case_tac ml, simp) +apply(rule_tac disjI2) +apply(rule_tac conjI, rule_tac x = ln in exI, simp) +apply(rule_tac x = rn in exI, simp) +apply(rule_tac disjI1) +apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, + rule_tac x = ln in exI, rule_tac x = rn in exI, simp add: exp_ind_def) +apply(simp) +done + +lemma wcode_fourtimes_case_first_correctness: + shows "let P = (\ (st, l, r). st = t_twice_len + 14) in + let Q = (\ (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r)) in + let f = (\ stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp) in + \ n .P (f n) \ Q (f (n::nat))" +proof - + let ?P = "(\ (st, l, r). st = t_twice_len + 14)" + let ?Q = "(\ (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r))" + let ?f = "(\ stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp)" + have "\ n . ?P (?f n) \ ?Q (?f (n::nat))" + proof(rule_tac halt_lemma2) + show "wf wcode_fourtimes_case_le" + by auto + next + show "\ na. \ ?P (?f na) \ ?Q (?f na) \ + ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_fourtimes_case_le" + apply(rule_tac allI, + case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na", simp, + rule_tac impI) + apply(simp add: tstep_red tstep.simps, case_tac c, simp, case_tac [2] aa, simp_all) + + apply(simp_all add: wcode_fourtimes_case_inv.simps new_tape.simps + wcode_fourtimes_case_le_def lex_pair_def split: if_splits) + done + next + show "?Q (?f 0)" + apply(simp add: steps.simps wcode_fourtimes_case_inv.simps) + apply(simp add: wcode_on_left_moving_2.simps wcode_on_left_moving_2_B.simps + wcode_on_left_moving_2_O.simps) + apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def) + apply(rule_tac x ="Suc 0" in exI, auto) + done + next + show "\ ?P (?f 0)" + apply(simp add: steps.simps) + done + qed + thus "?thesis" + apply(erule_tac exE, simp) + done +qed + +definition t_fourtimes_len :: "nat" + where + "t_fourtimes_len = (length t_fourtimes div 2)" + +lemma t_fourtimes_len_gr: "t_fourtimes_len > 0" +apply(simp add: t_fourtimes_len_def t_fourtimes_def) +done + +lemma t_fourtimes_correct: + "\stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) stp = + (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +proof(case_tac "rec_ci rec_fourtimes") + fix a b c + assume h: "rec_ci rec_fourtimes = (a, b, c)" + have "\stp m l. steps (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<^bsup>n\<^esup>) (tm_of abc_fourtimes @ tMp (Suc 0) + (start_of fourtimes_ly (length abc_fourtimes) - 1)) stp = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4*rs)\<^esup> @ Bk\<^bsup>l\<^esup>)" + proof(rule_tac t_compiled_by_rec) + show "rec_ci rec_fourtimes = (a, b, c)" by (simp add: h) + next + show "rec_calc_rel rec_fourtimes [rs] (4 * rs)" + using prime_rel_exec_eq [of rec_fourtimes "[rs]" "4 * rs"] + apply(subgoal_tac "primerec rec_fourtimes (length [rs])") + apply(simp_all add: rec_fourtimes_def rec_exec.simps) + apply(auto) + apply(simp only: Nat.One_nat_def[THEN sym], auto) + done + next + show "length [rs] = Suc 0" by simp + next + show "layout_of (a [+] dummy_abc (Suc 0)) = layout_of (a [+] dummy_abc (Suc 0))" + by simp + next + show "start_of fourtimes_ly (length abc_fourtimes) = + start_of (layout_of (a [+] dummy_abc (Suc 0))) (length (a [+] dummy_abc (Suc 0)))" + using h + apply(simp add: fourtimes_ly_def abc_fourtimes_def) + done + next + show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc (Suc 0))" + using h + apply(simp add: abc_fourtimes_def) + done + qed + thus "\stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) stp = + (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) + done +qed + +lemma t_fourtimes_change_term_state: + "\ stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp + = (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +using t_fourtimes_correct[of ires rs n] +apply(erule_tac exE) +apply(erule_tac exE) +apply(erule_tac exE) +proof(drule_tac turing_change_termi_state) + fix stp ln rn + show "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) + (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))" + apply(rule_tac t_compiled_correct, auto simp: fourtimes_ly_def) + done +next + fix stp ln rn + show "\stp. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) + (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) stp = + (Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly + (length abc_fourtimes) - Suc 0)) div 2), Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) \ + \stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp = + (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(erule_tac exE) + apply(simp add: t_fourtimes_len_def t_fourtimes_def) + apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp) + done +qed + +lemma t_fourtimes_append_pre: + "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp + = (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) + \ \ stp>0. steps (Suc 0 + length (t_wcode_main_first_part @ + tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, + Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + ((t_wcode_main_first_part @ + tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) @ + tshift t_fourtimes (length (t_wcode_main_first_part @ + tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2) @ ([(L, 1), (L, 1)])) stp + = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ + tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, + Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +proof(rule_tac t_tshift_lemma, auto) + assume "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp = + (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + thus "0 < stp" + using t_fourtimes_len_gr + apply(case_tac stp, simp_all add: steps.simps) + done +next + show "Suc 0 \ length t_fourtimes div 2" + apply(simp add: t_fourtimes_def shift_length tMp.simps) + done +next + show "t_ncorrect (t_wcode_main_first_part @ + abacus.tshift t_twice (length t_wcode_main_first_part div 2) @ + [(L, Suc 0), (L, Suc 0)])" + apply(simp add: t_ncorrect.simps t_wcode_main_first_part_def shift_length + t_twice_def) + using tm_even[of abc_twice] + by arith +next + show "t_ncorrect t_fourtimes" + apply(simp add: t_fourtimes_def steps.simps t_ncorrect.simps) + using tm_even[of abc_fourtimes] + by arith +next + show "t_ncorrect [(L, Suc 0), (L, Suc 0)]" + apply(simp add: t_ncorrect.simps) + done +qed + +lemma [simp]: "length t_wcode_main_first_part = 24" +apply(simp add: t_wcode_main_first_part_def) +done + +lemma [simp]: "(26 + length t_twice) div 2 = (length t_twice) div 2 + 13" +using tm_even[of abc_twice] +apply(simp add: t_twice_def) +done + +lemma [simp]: "((26 + length (abacus.tshift t_twice 12)) div 2) + = (length (abacus.tshift t_twice 12) div 2 + 13)" +using tm_even[of abc_twice] +apply(simp add: t_twice_def) +done + +lemma [simp]: "t_twice_len + 14 = 14 + length (abacus.tshift t_twice 12) div 2" +using tm_even[of abc_twice] +apply(simp add: t_twice_def t_twice_len_def shift_length) +done + +lemma t_fourtimes_append: + "\ stp ln rn. + steps (Suc 0 + length (t_wcode_main_first_part @ tshift t_twice + (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, + Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + ((t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @ + [(L, 1), (L, 1)]) @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)]) stp + = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ tshift t_twice + (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, + Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + using t_fourtimes_change_term_state[of ires rs n] + apply(erule_tac exE) + apply(erule_tac exE) + apply(erule_tac exE) + apply(drule_tac t_fourtimes_append_pre) + apply(erule_tac exE) + apply(rule_tac x = stpa in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) + apply(simp add: t_twice_len_def shift_length) + done + +lemma t_wcode_main_len: "length t_wcode_main = length t_twice + length t_fourtimes + 28" +apply(simp add: t_wcode_main_def shift_length) +done + +lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) b + = (L, Suc 0)" +using tm_even[of "abc_twice"] tm_even[of "abc_fourtimes"] +apply(case_tac b) +apply(simp_all only: fetch.simps) +apply(auto simp: nth_of.simps t_wcode_main_len t_twice_len_def + t_fourtimes_def t_twice_def t_fourtimes_def t_fourtimes_len_def) +apply(auto simp: t_wcode_main_def t_wcode_main_first_part_def shift_length t_twice_def nth_append + t_fourtimes_def) +done + +lemma wcode_jump2: + "\ stp ln rn. steps (t_twice_len + 14 + t_fourtimes_len + , Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4 * rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stp = + (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4 * rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +apply(rule_tac x = "Suc 0" in exI) +apply(simp add: steps.simps shift_length) +apply(rule_tac x = lnb in exI, rule_tac x = rnb in exI) +apply(simp add: tstep.simps new_tape.simps) +done + +lemma wcode_fourtimes_case: + shows "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)" +proof - + have "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (t_twice_len + 14, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + using wcode_fourtimes_case_first_correctness[of ires rs m n] + apply(simp add: wcode_fourtimes_case_inv.simps, auto) + apply(rule_tac x = na in exI, rule_tac x = ln in exI, + rule_tac x = rn in exI) + apply(simp) + done + from this obtain stpa lna rna where stp1: + "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa = + (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast + have "\stp ln rn. steps (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>) + t_wcode_main stp = + (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + using t_fourtimes_append[of " Bk\<^bsup>lna\<^esup> @ Oc # ires" "rs + 1" rna] + apply(erule_tac exE) + apply(erule_tac exE) + apply(erule_tac exE) + apply(simp add: t_wcode_main_def) + apply(rule_tac x = stp in exI, + rule_tac x = "ln + lna" in exI, + rule_tac x = rn in exI, simp) + apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym]) + done + from this obtain stpb lnb rnb where stp2: + "steps (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>) + t_wcode_main stpb = + (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>)" + by blast + have "\stp ln rn. steps (t_twice_len + 14 + t_fourtimes_len, + Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>) + t_wcode_main stp = + (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(rule wcode_jump2) + done + from this obtain stpc lnc rnc where stp3: + "steps (t_twice_len + 14 + t_fourtimes_len, + Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>) + t_wcode_main stpc = + (Suc 0, Bk # Bk\<^bsup>lnc\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnc\<^esup>)" + by blast + from stp1 stp2 stp3 show "?thesis" + apply(rule_tac x = "stpa + stpb + stpc" in exI, + rule_tac x = lnc in exI, rule_tac x = rnc in exI) + apply(simp add: steps_add) + done +qed + +(**********************************************************) + +fun wcode_on_left_moving_3_B :: "bin_inv_t" + where + "wcode_on_left_moving_3_B ires rs (l, r) = + (\ ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Bk # Bk # ires \ + r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr > Suc 0 \ mr > 0 )" + +fun wcode_on_left_moving_3_O :: "bin_inv_t" + where + "wcode_on_left_moving_3_O ires rs (l, r) = + (\ ln rn. l = Bk # Bk # ires \ + r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wcode_on_left_moving_3 :: "bin_inv_t" + where + "wcode_on_left_moving_3 ires rs (l, r) = + (wcode_on_left_moving_3_B ires rs (l, r) \ + wcode_on_left_moving_3_O ires rs (l, r))" + +fun wcode_on_checking_3 :: "bin_inv_t" + where + "wcode_on_checking_3 ires rs (l, r) = + (\ ln rn. l = Bk # ires \ + r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wcode_goon_checking_3 :: "bin_inv_t" + where + "wcode_goon_checking_3 ires rs (l, r) = + (\ ln rn. l = ires \ + r = Bk # Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wcode_stop :: "bin_inv_t" + where + "wcode_stop ires rs (l, r) = + (\ ln rn. l = Bk # ires \ + r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wcode_halt_case_inv :: "nat \ bin_inv_t" + where + "wcode_halt_case_inv st ires rs (l, r) = + (if st = 0 then wcode_stop ires rs (l, r) + else if st = Suc 0 then wcode_on_left_moving_3 ires rs (l, r) + else if st = Suc (Suc 0) then wcode_on_checking_3 ires rs (l, r) + else if st = 7 then wcode_goon_checking_3 ires rs (l, r) + else False)" + +fun wcode_halt_case_state :: "t_conf \ nat" + where + "wcode_halt_case_state (st, l, r) = + (if st = 1 then 5 + else if st = Suc (Suc 0) then 4 + else if st = 7 then 3 + else 0)" + +fun wcode_halt_case_step :: "t_conf \ nat" + where + "wcode_halt_case_step (st, l, r) = + (if st = 1 then length l + else 0)" + +fun wcode_halt_case_measure :: "t_conf \ nat \ nat" + where + "wcode_halt_case_measure (st, l, r) = + (wcode_halt_case_state (st, l, r), + wcode_halt_case_step (st, l, r))" + +definition wcode_halt_case_le :: "(t_conf \ t_conf) set" + where "wcode_halt_case_le \ (inv_image lex_pair wcode_halt_case_measure)" + +lemma wf_wcode_halt_case_le[intro]: "wf wcode_halt_case_le" +by(auto intro:wf_inv_image simp: wcode_halt_case_le_def) + +declare wcode_on_left_moving_3_B.simps[simp del] wcode_on_left_moving_3_O.simps[simp del] + wcode_on_checking_3.simps[simp del] wcode_goon_checking_3.simps[simp del] + wcode_on_left_moving_3.simps[simp del] wcode_stop.simps[simp del] + +lemmas wcode_halt_invs = + wcode_on_left_moving_3_B.simps wcode_on_left_moving_3_O.simps + wcode_on_checking_3.simps wcode_goon_checking_3.simps + wcode_on_left_moving_3.simps wcode_stop.simps + +lemma [simp]: "fetch t_wcode_main 7 Bk = (R, 0)" +apply(simp add: fetch.simps t_wcode_main_def nth_append nth_of.simps + t_wcode_main_first_part_def) +done + +lemma [simp]: "wcode_on_left_moving_3 ires rs (b, []) = False" +apply(simp only: wcode_halt_invs) +apply(simp add: exp_ind_def) +done + +lemma [simp]: "wcode_on_checking_3 ires rs (b, []) = False" +apply(simp add: wcode_halt_invs) +done + +lemma [simp]: "wcode_goon_checking_3 ires rs (b, []) = False" +apply(simp add: wcode_halt_invs) +done + +lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Bk # list) + \ wcode_on_left_moving_3 ires rs (tl b, hd b # Bk # list)" +apply(simp only: wcode_halt_invs) +apply(erule_tac disjE) +apply(erule_tac exE)+ +apply(case_tac ml, simp) +apply(rule_tac x = "mr - 2" in exI, rule_tac x = rn in exI) +apply(case_tac mr, simp, simp add: exp_ind, simp add: exp_ind[THEN sym]) +apply(rule_tac disjI1) +apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, + rule_tac x = rn in exI, simp add: exp_ind_def) +apply(simp) +done + +lemma [simp]: "wcode_goon_checking_3 ires rs (b, Bk # list) \ + (b = [] \ wcode_stop ires rs ([Bk], list)) \ + (b \ [] \ wcode_stop ires rs (Bk # b, list))" +apply(auto simp: wcode_halt_invs) +done + +lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Oc # list) \ b \ []" +apply(auto simp: wcode_halt_invs) +done + +lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Oc # list) \ + wcode_on_checking_3 ires rs (tl b, hd b # Oc # list)" +apply(simp add:wcode_halt_invs, auto) +apply(case_tac [!] mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "wcode_on_checking_3 ires rs (b, Oc # list) = False" +apply(auto simp: wcode_halt_invs) +done + +lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Bk # list) \ b \ []" +apply(simp add: wcode_halt_invs, auto) +done + + +lemma [simp]: "wcode_on_checking_3 ires rs (b, Bk # list) \ b \ []" +apply(auto simp: wcode_halt_invs) +done + +lemma [simp]: "wcode_on_checking_3 ires rs (b, Bk # list) \ + wcode_goon_checking_3 ires rs (tl b, hd b # Bk # list)" +apply(auto simp: wcode_halt_invs) +done + +lemma [simp]: "wcode_goon_checking_3 ires rs (b, Oc # list) = False" +apply(simp add: wcode_goon_checking_3.simps) +done + +lemma t_halt_case_correctness: +shows "let P = (\ (st, l, r). st = 0) in + let Q = (\ (st, l, r). wcode_halt_case_inv st ires rs (l, r)) in + let f = (\ stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp) in + \ n .P (f n) \ Q (f (n::nat))" +proof - + let ?P = "(\ (st, l, r). st = 0)" + let ?Q = "(\ (st, l, r). wcode_halt_case_inv st ires rs (l, r))" + let ?f = "(\ stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp)" + have "\ n. ?P (?f n) \ ?Q (?f (n::nat))" + proof(rule_tac halt_lemma2) + show "wf wcode_halt_case_le" by auto + next + show "\ na. \ ?P (?f na) \ ?Q (?f na) \ + ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_halt_case_le" + apply(rule_tac allI, rule_tac impI, case_tac "?f na") + apply(simp add: tstep_red tstep.simps) + apply(case_tac c, simp, case_tac [2] aa) + apply(simp_all split: if_splits add: new_tape.simps wcode_halt_case_le_def lex_pair_def) + done + next + show "?Q (?f 0)" + apply(simp add: steps.simps wcode_halt_invs) + apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def) + apply(rule_tac x = "Suc 0" in exI, auto) + done + next + show "\ ?P (?f 0)" + apply(simp add: steps.simps) + done + qed + thus "?thesis" + apply(auto) + done +qed + +declare wcode_halt_case_inv.simps[simp del] +lemma [intro]: "\ xs. ( :: block list) = Oc # xs" +apply(case_tac "rev list", simp) +apply(simp add: tape_of_nat_abv tape_of_nat_list.simps exp_ind_def) +apply(case_tac list, simp, simp) +done + +lemma wcode_halt_case: + "\stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + using t_halt_case_correctness[of ires rs m n] +apply(simp) +apply(erule_tac exE) +apply(case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, + Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na") +apply(auto simp: wcode_halt_case_inv.simps wcode_stop.simps) +apply(rule_tac x = na in exI, rule_tac x = ln in exI, + rule_tac x = rn in exI, simp) +done + +lemma bl_bin_one: "bl_bin [Oc] = Suc 0" +apply(simp add: bl_bin.simps) +done + +lemma t_wcode_main_lemma_pre: + "\args \ []; lm = \ \ + \ stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main + stp + = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2^(length lm - 1) \<^esup> @ Bk\<^bsup>rn\<^esup>)" +proof(induct "length args" arbitrary: args lm rs m n, simp) + fix x args lm rs m n + assume ind: + "\args lm rs m n. + \x = length args; (args::nat list) \ []; lm = \ + \ \stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + + and h: "Suc x = length args" "(args::nat list) \ []" "lm = " + from h have "\ (a::nat) xs. args = xs @ [a]" + apply(rule_tac x = "last args" in exI) + apply(rule_tac x = "butlast args" in exI, auto) + done + from this obtain a xs where "args = xs @ [a]" by blast + from h and this show + "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + proof(case_tac "xs::nat list", simp) + show "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev () @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin () + rs * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)" + proof(induct "a" arbitrary: m n rs ires, simp) + fix m n rs ires + show "\stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) + t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin [Oc] + rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(simp add: bl_bin_one) + apply(rule_tac wcode_halt_case) + done + next + fix a m n rs ires + assume ind2: + "\m n rs ires. + \stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev () @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin () + rs * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)" + show "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev () @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin () + rs * 2 ^ Suc a\<^esup> @ Bk\<^bsup>rn\<^esup>)" + proof - + have "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev () @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ rev () @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(simp add: tape_of_nat) + using wcode_double_case[of m "Oc\<^bsup>a\<^esup> @ Bk # Bk # ires" rs n] + apply(simp add: exp_ind_def) + done + from this obtain stpa lna rna where stp1: + "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev () @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa = + (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev () @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast + moreover have + "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev () @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp = + (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin () + (2*rs + 2) * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)" + using ind2[of lna ires "2*rs + 2" rna] by simp + from this obtain stpb lnb rnb where stp2: + "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev () @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb = + (0, Bk # ires, Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin () + (2*rs + 2) * 2 ^ a\<^esup> @ Bk\<^bsup>rnb\<^esup>)" + by blast + from stp1 and stp2 show "?thesis" + apply(rule_tac x = "stpa + stpb" in exI, + rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp) + apply(simp add: steps_add bl_bin_nat_Suc exponent_def) + done + qed + qed + next + fix aa list + assume g: "Suc x = length args" "args \ []" "lm = " "args = xs @ [a::nat]" "xs = (aa::nat) # list" + thus "\stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + proof(induct a arbitrary: m n rs args lm, simp_all add: tape_of_nl_rev, + simp only: tape_of_nl_cons_app1, simp) + fix m n rs args lm + have "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # rev (<(aa::nat) # list>) @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ rev () @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + proof(simp add: tape_of_nl_rev) + have "\ xs. () = Oc # xs" by auto + from this obtain xs where "() = Oc # xs" .. + thus "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ @ Bk # Bk # ires, Bk # Oc\<^bsup>5 + 4 * rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(simp) + using wcode_fourtimes_case[of m "xs @ Bk # Bk # ires" rs n] + apply(simp) + done + qed + from this obtain stpa lna rna where stp1: + "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # rev () @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa = + (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev () @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast + from g have + "\ stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp = (0, Bk # ires, + Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin ()+ (4*rs + 4) * 2^(length () - 1) \<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(rule_tac args = "(aa::nat)#list" in ind, simp_all) + done + from this obtain stpb lnb rnb where stp2: + "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb = (0, Bk # ires, + Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin ()+ (4*rs + 4) * 2^(length () - 1) \<^esup> @ Bk\<^bsup>rnb\<^esup>)" + by blast + from stp1 and stp2 and h + show "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # + Bk # Oc\<^bsup>bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # ) + rs * (2 * 2 ^ (aa + length ()))\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, + rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_rev) + done + next + fix ab m n rs args lm + assume ind2: + "\ m n rs args lm. + \lm = ; args = aa # list @ [ab]\ + \ \stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # + Bk # Oc\<^bsup>bl_bin () + rs * 2 ^ (length () - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + and k: "args = aa # list @ [Suc ab]" "lm = " + show "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (0, Bk # ires,Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # + Bk # Oc\<^bsup>bl_bin () + rs * 2 ^ (length () - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + proof(simp add: tape_of_nl_cons_app1) + have "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # @ Bk # Bk # ires, + Bk # Oc # Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp + = (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc\<^bsup>Suc ab\<^esup> @ Bk # @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + using wcode_double_case[of m "Oc\<^bsup>ab\<^esup> @ Bk # @ Bk # Bk # ires" + rs n] + apply(simp add: exp_ind_def) + done + from this obtain stpa lna rna where stp1: + "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # @ Bk # Bk # ires, + Bk # Oc # Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa + = (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ Oc\<^bsup>Suc ab\<^esup> @ Bk # @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast + from k have + "\ stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp + = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # + Bk # Oc\<^bsup>bl_bin ( ) + (2*rs + 2)* 2^(length () - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(rule_tac ind2, simp_all) + done + from this obtain stpb lnb rnb where stp2: + "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb + = (0, Bk # ires, Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk # + Bk # Oc\<^bsup>bl_bin ( ) + (2*rs + 2)* 2^(length () - Suc 0)\<^esup> @ Bk\<^bsup>rnb\<^esup>)" + by blast + from stp1 and stp2 show + "\stp ln rn. + steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # @ Bk # Bk # ires, + Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = + (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # + Oc\<^bsup>bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # ) + rs * (2 * 2 ^ (aa + length ()))\<^esup> + @ Bk\<^bsup>rn\<^esup>)" + apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, + rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_cons_app1 exp_ind_def) + done + qed + qed + qed + qed + + + +(* turing_shift can be used*) +term t_wcode_main +definition t_wcode_prepare :: "tprog" + where + "t_wcode_prepare \ + [(W1, 2), (L, 1), (L, 3), (R, 2), (R, 4), (W0, 3), + (R, 4), (R, 5), (R, 6), (R, 5), (R, 7), (R, 5), + (W1, 7), (L, 0)]" + +fun wprepare_add_one :: "nat \ nat list \ tape \ bool" + where + "wprepare_add_one m lm (l, r) = + (\ rn. l = [] \ + (r = @ Bk\<^bsup>rn\<^esup> \ + r = Bk # @ Bk\<^bsup>rn\<^esup>))" + +fun wprepare_goto_first_end :: "nat \ nat list \ tape \ bool" + where + "wprepare_goto_first_end m lm (l, r) = + (\ ml mr rn. l = Oc\<^bsup>ml\<^esup> \ + r = Oc\<^bsup>mr\<^esup> @ Bk # @ Bk\<^bsup>rn\<^esup> \ + ml + mr = Suc (Suc m))" + +fun wprepare_erase :: "nat \ nat list \ tape \ bool" + where + "wprepare_erase m lm (l, r) = + (\ rn. l = Oc\<^bsup>Suc m\<^esup> \ + tl r = Bk # @ Bk\<^bsup>rn\<^esup>)" + +fun wprepare_goto_start_pos_B :: "nat \ nat list \ tape \ bool" + where + "wprepare_goto_start_pos_B m lm (l, r) = + (\ rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Bk # @ Bk\<^bsup>rn\<^esup>)" + +fun wprepare_goto_start_pos_O :: "nat \ nat list \ tape \ bool" + where + "wprepare_goto_start_pos_O m lm (l, r) = + (\ rn. l = Bk # Bk # Oc\<^bsup>Suc m\<^esup> \ + r = @ Bk\<^bsup>rn\<^esup>)" + +fun wprepare_goto_start_pos :: "nat \ nat list \ tape \ bool" + where + "wprepare_goto_start_pos m lm (l, r) = + (wprepare_goto_start_pos_B m lm (l, r) \ + wprepare_goto_start_pos_O m lm (l, r))" + +fun wprepare_loop_start_on_rightmost :: "nat \ nat list \ tape \ bool" + where + "wprepare_loop_start_on_rightmost m lm (l, r) = + (\ rn mr. rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # @ Bk\<^bsup>rn\<^esup> \ l \ [] \ + r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wprepare_loop_start_in_middle :: "nat \ nat list \ tape \ bool" + where + "wprepare_loop_start_in_middle m lm (l, r) = + (\ rn (mr:: nat) (lm1::nat list). + rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # @ Bk\<^bsup>rn\<^esup> \ l \ [] \ + r = Oc\<^bsup>mr\<^esup> @ Bk # @ Bk\<^bsup>rn\<^esup> \ lm1 \ [])" + +fun wprepare_loop_start :: "nat \ nat list \ tape \ bool" + where + "wprepare_loop_start m lm (l, r) = (wprepare_loop_start_on_rightmost m lm (l, r) \ + wprepare_loop_start_in_middle m lm (l, r))" + +fun wprepare_loop_goon_on_rightmost :: "nat \ nat list \ tape \ bool" + where + "wprepare_loop_goon_on_rightmost m lm (l, r) = + (\ rn. l = Bk # @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Bk\<^bsup>rn\<^esup>)" + +fun wprepare_loop_goon_in_middle :: "nat \ nat list \ tape \ bool" + where + "wprepare_loop_goon_in_middle m lm (l, r) = + (\ rn (mr:: nat) (lm1::nat list). + rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # @ Bk\<^bsup>rn\<^esup> \ l \ [] \ + (if lm1 = [] then r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> + else r = Oc\<^bsup>mr\<^esup> @ Bk # @ Bk\<^bsup>rn\<^esup>) \ mr > 0)" + +fun wprepare_loop_goon :: "nat \ nat list \ tape \ bool" + where + "wprepare_loop_goon m lm (l, r) = + (wprepare_loop_goon_in_middle m lm (l, r) \ + wprepare_loop_goon_on_rightmost m lm (l, r))" + +fun wprepare_add_one2 :: "nat \ nat list \ tape \ bool" + where + "wprepare_add_one2 m lm (l, r) = + (\ rn. l = Bk # Bk # @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \ + (r = [] \ tl r = Bk\<^bsup>rn\<^esup>))" + +fun wprepare_stop :: "nat \ nat list \ tape \ bool" + where + "wprepare_stop m lm (l, r) = + (\ rn. l = Bk # @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Bk # Oc # Bk\<^bsup>rn\<^esup>)" + +fun wprepare_inv :: "nat \ nat \ nat list \ tape \ bool" + where + "wprepare_inv st m lm (l, r) = + (if st = 0 then wprepare_stop m lm (l, r) + else if st = Suc 0 then wprepare_add_one m lm (l, r) + else if st = Suc (Suc 0) then wprepare_goto_first_end m lm (l, r) + else if st = Suc (Suc (Suc 0)) then wprepare_erase m lm (l, r) + else if st = 4 then wprepare_goto_start_pos m lm (l, r) + else if st = 5 then wprepare_loop_start m lm (l, r) + else if st = 6 then wprepare_loop_goon m lm (l, r) + else if st = 7 then wprepare_add_one2 m lm (l, r) + else False)" + +fun wprepare_stage :: "t_conf \ nat" + where + "wprepare_stage (st, l, r) = + (if st \ 1 \ st \ 4 then 3 + else if st = 5 \ st = 6 then 2 + else 1)" + +fun wprepare_state :: "t_conf \ nat" + where + "wprepare_state (st, l, r) = + (if st = 1 then 4 + else if st = Suc (Suc 0) then 3 + else if st = Suc (Suc (Suc 0)) then 2 + else if st = 4 then 1 + else if st = 7 then 2 + else 0)" + +fun wprepare_step :: "t_conf \ nat" + where + "wprepare_step (st, l, r) = + (if st = 1 then (if hd r = Oc then Suc (length l) + else 0) + else if st = Suc (Suc 0) then length r + else if st = Suc (Suc (Suc 0)) then (if hd r = Oc then 1 + else 0) + else if st = 4 then length r + else if st = 5 then Suc (length r) + else if st = 6 then (if r = [] then 0 else Suc (length r)) + else if st = 7 then (if (r \ [] \ hd r = Oc) then 0 + else 1) + else 0)" + +fun wcode_prepare_measure :: "t_conf \ nat \ nat \ nat" + where + "wcode_prepare_measure (st, l, r) = + (wprepare_stage (st, l, r), + wprepare_state (st, l, r), + wprepare_step (st, l, r))" + +definition wcode_prepare_le :: "(t_conf \ t_conf) set" + where "wcode_prepare_le \ (inv_image lex_triple wcode_prepare_measure)" + +lemma [intro]: "wf lex_pair" +by(auto intro:wf_lex_prod simp:lex_pair_def) + +lemma wf_wcode_prepare_le[intro]: "wf wcode_prepare_le" +by(auto intro:wf_inv_image simp: wcode_prepare_le_def + recursive.lex_triple_def) + +declare wprepare_add_one.simps[simp del] wprepare_goto_first_end.simps[simp del] + wprepare_erase.simps[simp del] wprepare_goto_start_pos.simps[simp del] + wprepare_loop_start.simps[simp del] wprepare_loop_goon.simps[simp del] + wprepare_add_one2.simps[simp del] + +lemmas wprepare_invs = wprepare_add_one.simps wprepare_goto_first_end.simps + wprepare_erase.simps wprepare_goto_start_pos.simps + wprepare_loop_start.simps wprepare_loop_goon.simps + wprepare_add_one2.simps + +declare wprepare_inv.simps[simp del] +lemma [simp]: "fetch t_wcode_prepare (Suc 0) Bk = (W1, 2)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare (Suc 0) Oc = (L, 1)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare (Suc (Suc 0)) Bk = (L, 3)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare (Suc (Suc 0)) Oc = (R, 2)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Bk = (R, 4)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Oc = (W0, 3)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare 4 Bk = (R, 4)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare 4 Oc = (R, 5)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare 5 Oc = (R, 5)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare 5 Bk = (R, 6)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare 6 Oc = (R, 5)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare 6 Bk = (R, 7)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare 7 Oc = (L, 0)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_prepare 7 Bk = (W1, 7)" +apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) +done + +lemma tape_of_nl_not_null: "lm \ [] \ \ []" +apply(case_tac lm, auto) +apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def) +done + +lemma [simp]: "lm \ [] \ wprepare_add_one m lm (b, []) = False" +apply(simp add: wprepare_invs) +apply(simp add: tape_of_nl_not_null) +done + +lemma [simp]: "lm \ [] \ wprepare_goto_first_end m lm (b, []) = False" +apply(simp add: wprepare_invs) +done + +lemma [simp]: "lm \ [] \ wprepare_erase m lm (b, []) = False" +apply(simp add: wprepare_invs) +done + + + +lemma [simp]: "lm \ [] \ wprepare_goto_start_pos m lm (b, []) = False" +apply(simp add: wprepare_invs tape_of_nl_not_null) +done + +lemma [simp]: "\lm \ []; wprepare_loop_start m lm (b, [])\ \ b \ []" +apply(simp add: wprepare_invs tape_of_nl_not_null, auto) +done + +lemma [simp]: "\lm \ []; wprepare_loop_start m lm (b, [])\ \ + wprepare_loop_goon m lm (Bk # b, [])" +apply(simp only: wprepare_invs tape_of_nl_not_null) +apply(erule_tac disjE) +apply(rule_tac disjI2) +apply(simp add: wprepare_loop_start_on_rightmost.simps + wprepare_loop_goon_on_rightmost.simps, auto) +apply(rule_tac rev_eq, simp add: tape_of_nl_rev) +done + +lemma [simp]: "\lm \ []; wprepare_loop_goon m lm (b, [])\ \ b \ []" +apply(simp only: wprepare_invs tape_of_nl_not_null, auto) +done + +lemma [simp]:"\lm \ []; wprepare_loop_goon m lm (b, [])\ \ + wprepare_add_one2 m lm (Bk # b, [])" +apply(simp only: wprepare_invs tape_of_nl_not_null, auto split: if_splits) +apply(case_tac mr, simp, simp add: exp_ind_def) +done + +lemma [simp]: "wprepare_add_one2 m lm (b, []) \ b \ []" +apply(simp only: wprepare_invs tape_of_nl_not_null, auto) +done + +lemma [simp]: "wprepare_add_one2 m lm (b, []) \ wprepare_add_one2 m lm (b, [Oc])" +apply(simp only: wprepare_invs tape_of_nl_not_null, auto) +done + +lemma [simp]: "Bk # list = <(m::nat) # lm> @ ys = False" +apply(case_tac lm, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def) +done + +lemma [simp]: "\lm \ []; wprepare_add_one m lm (b, Bk # list)\ + \ (b = [] \ wprepare_goto_first_end m lm ([], Oc # list)) \ + (b \ [] \ wprepare_goto_first_end m lm (b, Oc # list))" +apply(simp only: wprepare_invs, auto) +apply(rule_tac x = 0 in exI, simp add: exp_ind_def) +apply(case_tac lm, simp, simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def) +apply(rule_tac x = rn in exI, simp) +done + +lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \ b \ []" +apply(simp only: wprepare_invs tape_of_nl_not_null, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \ + wprepare_erase m lm (tl b, hd b # Bk # list)" +apply(simp only: wprepare_invs tape_of_nl_not_null, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(case_tac mr, auto simp: exp_ind_def) +done + +lemma [simp]: "wprepare_erase m lm (b, Bk # list) \ b \ []" +apply(simp only: wprepare_invs exp_ind_def, auto) +done + +lemma [simp]: "wprepare_erase m lm (b, Bk # list) \ + wprepare_goto_start_pos m lm (Bk # b, list)" +apply(simp only: wprepare_invs, auto) +done + +lemma [simp]: "\wprepare_add_one m lm (b, Bk # list)\ \ list \ []" +apply(simp only: wprepare_invs) +apply(case_tac lm, simp_all add: tape_of_nl_abv + tape_of_nat_list.simps exp_ind_def, auto) +done + +lemma [simp]: "\lm \ []; wprepare_goto_first_end m lm (b, Bk # list)\ \ list \ []" +apply(simp only: wprepare_invs, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(simp add: tape_of_nl_not_null) +done + +lemma [simp]: "\lm \ []; wprepare_goto_first_end m lm (b, Bk # list)\ \ b \ []" +apply(simp only: wprepare_invs, auto) +apply(case_tac mr, simp_all add: exp_ind_def tape_of_nl_not_null) +done + +lemma [simp]: "\lm \ []; wprepare_erase m lm (b, Bk # list)\ \ list \ []" +apply(simp only: wprepare_invs, auto) +done + +lemma [simp]: "\lm \ []; wprepare_erase m lm (b, Bk # list)\ \ b \ []" +apply(simp only: wprepare_invs, auto simp: exp_ind_def) +done + +lemma [simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Bk # list)\ \ list \ []" +apply(simp only: wprepare_invs, auto) +apply(simp add: tape_of_nl_not_null) +apply(case_tac lm, simp, case_tac list) +apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def) +done + +lemma [simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Bk # list)\ \ b \ []" +apply(simp only: wprepare_invs) +apply(auto) +done + +lemma [simp]: "\lm \ []; wprepare_loop_goon m lm (b, Bk # list)\ \ b \ []" +apply(simp only: wprepare_invs, auto) +done + +lemma [simp]: "\lm \ []; wprepare_loop_goon m lm (b, Bk # list)\ \ + (list = [] \ wprepare_add_one2 m lm (Bk # b, [])) \ + (list \ [] \ wprepare_add_one2 m lm (Bk # b, list))" +apply(simp only: wprepare_invs, simp) +apply(case_tac list, simp_all split: if_splits, auto) +apply(case_tac [1-3] mr, simp_all add: exp_ind_def) +apply(case_tac mr, simp_all add: exp_ind_def tape_of_nl_not_null) +apply(case_tac [1-2] mr, simp_all add: exp_ind_def) +apply(case_tac rn, simp, case_tac nat, auto simp: exp_ind_def) +done + +lemma [simp]: "wprepare_add_one2 m lm (b, Bk # list) \ b \ []" +apply(simp only: wprepare_invs, simp) +done + +lemma [simp]: "wprepare_add_one2 m lm (b, Bk # list) \ + (list = [] \ wprepare_add_one2 m lm (b, [Oc])) \ + (list \ [] \ wprepare_add_one2 m lm (b, Oc # list))" +apply(simp only: wprepare_invs, auto) +done + +lemma [simp]: "wprepare_goto_first_end m lm (b, Oc # list) + \ (b = [] \ wprepare_goto_first_end m lm ([Oc], list)) \ + (b \ [] \ wprepare_goto_first_end m lm (Oc # b, list))" +apply(simp only: wprepare_invs, auto) +apply(rule_tac x = 1 in exI, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(case_tac ml, simp_all add: exp_ind_def) +apply(rule_tac x = rn in exI, simp) +apply(rule_tac x = "Suc ml" in exI, simp_all add: exp_ind_def) +apply(rule_tac x = "mr - 1" in exI, simp) +apply(case_tac mr, simp_all add: exp_ind_def, auto) +done + +lemma [simp]: "wprepare_erase m lm (b, Oc # list) \ b \ []" +apply(simp only: wprepare_invs, auto simp: exp_ind_def) +done + +lemma [simp]: "wprepare_erase m lm (b, Oc # list) + \ wprepare_erase m lm (b, Bk # list)" +apply(simp only:wprepare_invs, auto simp: exp_ind_def) +done + +lemma [simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Bk # list)\ + \ wprepare_goto_start_pos m lm (Bk # b, list)" +apply(simp only:wprepare_invs, auto) +apply(case_tac [!] lm, simp, simp_all) +done + +lemma [simp]: "wprepare_loop_start m lm (b, aa) \ b \ []" +apply(simp only:wprepare_invs, auto) +done +lemma [elim]: "Bk # list = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ \rn. list = Bk\<^bsup>rn\<^esup>" +apply(case_tac mr, simp_all) +apply(case_tac rn, simp_all add: exp_ind_def, auto) +done + +lemma rev_equal_iff: "x = y \ rev x = rev y" +by simp + +lemma tape_of_nl_false1: + "lm \ [] \ rev b @ [Bk] \ Bk\<^bsup>ln\<^esup> @ Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # " +apply(auto) +apply(drule_tac rev_equal_iff, simp add: tape_of_nl_rev) +apply(case_tac "rev lm") +apply(case_tac [2] list, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def) +done + +lemma [simp]: "wprepare_loop_start_in_middle m lm (b, [Bk]) = False" +apply(simp add: wprepare_loop_start_in_middle.simps, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(case_tac lm1, simp, simp add: tape_of_nl_not_null) +done + +declare wprepare_loop_start_in_middle.simps[simp del] + +declare wprepare_loop_start_on_rightmost.simps[simp del] + wprepare_loop_goon_in_middle.simps[simp del] + wprepare_loop_goon_on_rightmost.simps[simp del] + +lemma [simp]: "wprepare_loop_goon_in_middle m lm (Bk # b, []) = False" +apply(simp add: wprepare_loop_goon_in_middle.simps, auto) +done + +lemma [simp]: "\lm \ []; wprepare_loop_start m lm (b, [Bk])\ \ + wprepare_loop_goon m lm (Bk # b, [])" +apply(simp only: wprepare_invs, simp) +apply(simp add: wprepare_loop_goon_on_rightmost.simps + wprepare_loop_start_on_rightmost.simps, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(rule_tac rev_eq) +apply(simp add: tape_of_nl_rev) +apply(simp add: exp_ind_def[THEN sym] exp_ind) +done + +lemma [simp]: "wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista) + \ wprepare_loop_goon_in_middle m lm (Bk # b, a # lista) = False" +apply(auto simp: wprepare_loop_start_on_rightmost.simps + wprepare_loop_goon_in_middle.simps) +apply(case_tac [!] mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "\lm \ []; wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)\ + \ wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista)" +apply(simp only: wprepare_loop_start_on_rightmost.simps + wprepare_loop_goon_on_rightmost.simps, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(simp add: tape_of_nl_rev) +apply(simp add: exp_ind_def[THEN sym] exp_ind) +done + +lemma [simp]: "\lm \ []; wprepare_loop_start_in_middle m lm (b, Bk # a # lista)\ + \ wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista) = False" +apply(simp add: wprepare_loop_start_in_middle.simps + wprepare_loop_goon_on_rightmost.simps, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(case_tac "lm1::nat list", simp_all, case_tac list, simp) +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps tape_of_nat_abv exp_ind_def) +apply(case_tac [!] rna, simp_all add: exp_ind_def) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(case_tac lm1, simp, case_tac list, simp) +apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def tape_of_nat_abv) +done + +lemma [simp]: + "\lm \ []; wprepare_loop_start_in_middle m lm (b, Bk # a # lista)\ + \ wprepare_loop_goon_in_middle m lm (Bk # b, a # lista)" +apply(simp add: wprepare_loop_start_in_middle.simps + wprepare_loop_goon_in_middle.simps, auto) +apply(rule_tac x = rn in exI, simp) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(case_tac lm1, simp) +apply(rule_tac x = "Suc aa" in exI, simp) +apply(rule_tac x = list in exI) +apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma [simp]: "\lm \ []; wprepare_loop_start m lm (b, Bk # a # lista)\ \ + wprepare_loop_goon m lm (Bk # b, a # lista)" +apply(simp add: wprepare_loop_start.simps + wprepare_loop_goon.simps) +apply(erule_tac disjE, simp, auto) +done + +lemma start_2_goon: + "\lm \ []; wprepare_loop_start m lm (b, Bk # list)\ \ + (list = [] \ wprepare_loop_goon m lm (Bk # b, [])) \ + (list \ [] \ wprepare_loop_goon m lm (Bk # b, list))" +apply(case_tac list, auto) +done + +lemma add_one_2_add_one: "wprepare_add_one m lm (b, Oc # list) + \ (hd b = Oc \ (b = [] \ wprepare_add_one m lm ([], Bk # Oc # list)) \ + (b \ [] \ wprepare_add_one m lm (tl b, Oc # Oc # list))) \ + (hd b \ Oc \ (b = [] \ wprepare_add_one m lm ([], Bk # Oc # list)) \ + (b \ [] \ wprepare_add_one m lm (tl b, hd b # Oc # list)))" +apply(simp only: wprepare_add_one.simps, auto) +done + +lemma [simp]: "wprepare_loop_start m lm (b, Oc # list) \ b \ []" +apply(simp) +done + +lemma [simp]: "wprepare_loop_start_on_rightmost m lm (b, Oc # list) \ + wprepare_loop_start_on_rightmost m lm (Oc # b, list)" +apply(simp add: wprepare_loop_start_on_rightmost.simps, auto) +apply(rule_tac x = rn in exI, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(case_tac rn, auto simp: exp_ind_def) +done + +lemma [simp]: "wprepare_loop_start_in_middle m lm (b, Oc # list) \ + wprepare_loop_start_in_middle m lm (Oc # b, list)" +apply(simp add: wprepare_loop_start_in_middle.simps, auto) +apply(rule_tac x = rn in exI, auto) +apply(case_tac mr, simp, simp add: exp_ind_def) +apply(rule_tac x = nat in exI, simp) +apply(rule_tac x = lm1 in exI, simp) +done + +lemma start_2_start: "wprepare_loop_start m lm (b, Oc # list) \ + wprepare_loop_start m lm (Oc # b, list)" +apply(simp add: wprepare_loop_start.simps) +apply(erule_tac disjE, simp_all ) +done + +lemma [simp]: "wprepare_loop_goon m lm (b, Oc # list) \ b \ []" +apply(simp add: wprepare_loop_goon.simps + wprepare_loop_goon_in_middle.simps + wprepare_loop_goon_on_rightmost.simps) +apply(auto) +done + +lemma [simp]: "wprepare_goto_start_pos m lm (b, Oc # list) \ b \ []" +apply(simp add: wprepare_goto_start_pos.simps) +done + +lemma [simp]: "wprepare_loop_goon_on_rightmost m lm (b, Oc # list) = False" +apply(simp add: wprepare_loop_goon_on_rightmost.simps) +done +lemma wprepare_loop1: "\rev b @ Oc\<^bsup>mr\<^esup> = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ; + b \ []; 0 < mr; Oc # list = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>\ + \ wprepare_loop_start_on_rightmost m lm (Oc # b, list)" +apply(simp add: wprepare_loop_start_on_rightmost.simps) +apply(rule_tac x = rn in exI, simp) +apply(case_tac mr, simp, simp add: exp_ind_def, auto) +done + +lemma wprepare_loop2: "\rev b @ Oc\<^bsup>mr\<^esup> @ Bk # = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ; + b \ []; Oc # list = Oc\<^bsup>mr\<^esup> @ Bk # <(a::nat) # lista> @ Bk\<^bsup>rn\<^esup>\ + \ wprepare_loop_start_in_middle m lm (Oc # b, list)" +apply(simp add: wprepare_loop_start_in_middle.simps) +apply(rule_tac x = rn in exI, simp) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(rule_tac x = nat in exI, simp) +apply(rule_tac x = "a#lista" in exI, simp) +done + +lemma [simp]: "wprepare_loop_goon_in_middle m lm (b, Oc # list) \ + wprepare_loop_start_on_rightmost m lm (Oc # b, list) \ + wprepare_loop_start_in_middle m lm (Oc # b, list)" +apply(simp add: wprepare_loop_goon_in_middle.simps split: if_splits) +apply(case_tac lm1, simp_all add: wprepare_loop1 wprepare_loop2) +done + +lemma [simp]: "wprepare_loop_goon m lm (b, Oc # list) + \ wprepare_loop_start m lm (Oc # b, list)" +apply(simp add: wprepare_loop_goon.simps + wprepare_loop_start.simps) +done + +lemma [simp]: "wprepare_add_one m lm (b, Oc # list) + \ b = [] \ wprepare_add_one m lm ([], Bk # Oc # list)" +apply(auto) +apply(simp add: wprepare_add_one.simps) +done + +lemma [simp]: "wprepare_goto_start_pos m [a] (b, Oc # list) + \ wprepare_loop_start_on_rightmost m [a] (Oc # b, list) " +apply(auto simp: wprepare_goto_start_pos.simps + wprepare_loop_start_on_rightmost.simps) +apply(rule_tac x = rn in exI, simp) +apply(simp add: tape_of_nat_abv tape_of_nat_list.simps exp_ind_def, auto) +done + +lemma [simp]: "wprepare_goto_start_pos m (a # aa # listaa) (b, Oc # list) + \wprepare_loop_start_in_middle m (a # aa # listaa) (Oc # b, list)" +apply(auto simp: wprepare_goto_start_pos.simps + wprepare_loop_start_in_middle.simps) +apply(rule_tac x = rn in exI, simp) +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def) +apply(rule_tac x = a in exI, rule_tac x = "aa#listaa" in exI, simp) +done + +lemma [simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Oc # list)\ + \ wprepare_loop_start m lm (Oc # b, list)" +apply(case_tac lm, simp_all) +apply(case_tac lista, simp_all add: wprepare_loop_start.simps) +done + +lemma [simp]: "wprepare_add_one2 m lm (b, Oc # list) \ b \ []" +apply(auto simp: wprepare_add_one2.simps) +done + +lemma add_one_2_stop: + "wprepare_add_one2 m lm (b, Oc # list) + \ wprepare_stop m lm (tl b, hd b # Oc # list)" +apply(simp add: wprepare_stop.simps wprepare_add_one2.simps) +done + +declare wprepare_stop.simps[simp del] + +lemma wprepare_correctness: + assumes h: "lm \ []" + shows "let P = (\ (st, l, r). st = 0) in + let Q = (\ (st, l, r). wprepare_inv st m lm (l, r)) in + let f = (\ stp. steps (Suc 0, [], ()) t_wcode_prepare stp) in + \ n .P (f n) \ Q (f n)" +proof - + let ?P = "(\ (st, l, r). st = 0)" + let ?Q = "(\ (st, l, r). wprepare_inv st m lm (l, r))" + let ?f = "(\ stp. steps (Suc 0, [], ()) t_wcode_prepare stp)" + have "\ n. ?P (?f n) \ ?Q (?f n)" + proof(rule_tac halt_lemma2) + show "wf wcode_prepare_le" by auto + next + show "\ n. \ ?P (?f n) \ ?Q (?f n) \ + ?Q (?f (Suc n)) \ (?f (Suc n), ?f n) \ wcode_prepare_le" + using h + apply(rule_tac allI, rule_tac impI, case_tac "?f n", + simp add: tstep_red tstep.simps) + apply(case_tac c, simp, case_tac [2] aa) + apply(simp_all add: wprepare_inv.simps wcode_prepare_le_def new_tape.simps + lex_triple_def lex_pair_def + + split: if_splits) + apply(simp_all add: start_2_goon start_2_start + add_one_2_add_one add_one_2_stop) + apply(auto simp: wprepare_add_one2.simps) + done + next + show "?Q (?f 0)" + apply(simp add: steps.simps wprepare_inv.simps wprepare_invs) + done + next + show "\ ?P (?f 0)" + apply(simp add: steps.simps) + done + qed + thus "?thesis" + apply(auto) + done +qed + +lemma [intro]: "t_correct t_wcode_prepare" +apply(simp add: t_correct.simps t_wcode_prepare_def iseven_def) +apply(rule_tac x = 7 in exI, simp) +done + +lemma twice_len_even: "length (tm_of abc_twice) mod 2 = 0" +apply(simp add: tm_even) +done + +lemma fourtimes_len_even: "length (tm_of abc_fourtimes) mod 2 = 0" +apply(simp add: tm_even) +done + +lemma t_correct_termi: "t_correct tp \ + list_all (\(acn, st). (st \ Suc (length tp div 2))) (change_termi_state tp)" +apply(auto simp: t_correct.simps List.list_all_length) +apply(erule_tac x = n in allE, simp) +apply(case_tac "tp!n", auto simp: change_termi_state.simps split: if_splits) +done + + +lemma t_correct_shift: + "list_all (\(acn, st). (st \ y)) tp \ + list_all (\(acn, st). (st \ y + off)) (tshift tp off) " +apply(auto simp: t_correct.simps List.list_all_length) +apply(erule_tac x = n in allE, simp add: shift_length) +apply(case_tac "tp!n", auto simp: tshift.simps) +done + +lemma [intro]: + "t_correct (tm_of abc_twice @ tMp (Suc 0) + (start_of twice_ly (length abc_twice) - Suc 0))" +apply(rule_tac t_compiled_correct, simp_all) +apply(simp add: twice_ly_def) +done + +lemma [intro]: "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) + (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))" +apply(rule_tac t_compiled_correct, simp_all) +apply(simp add: fourtimes_ly_def) +done + + +lemma [intro]: "t_correct t_wcode_main" +apply(auto simp: t_wcode_main_def t_correct.simps shift_length + t_twice_def t_fourtimes_def) +proof - + show "iseven (60 + (length (tm_of abc_twice) + + length (tm_of abc_fourtimes)))" + using twice_len_even fourtimes_len_even + apply(auto simp: iseven_def) + apply(rule_tac x = "30 + q + qa" in exI, simp) + done +next + show " list_all (\(acn, s). s \ (60 + (length (tm_of abc_twice) + + length (tm_of abc_fourtimes))) div 2) t_wcode_main_first_part" + apply(auto simp: t_wcode_main_first_part_def shift_length t_twice_def) + done +next + have "list_all (\(acn, s). s \ Suc (length (tm_of abc_twice @ tMp (Suc 0) + (start_of twice_ly (length abc_twice) - Suc 0)) div 2)) + (change_termi_state (tm_of abc_twice @ tMp (Suc 0) + (start_of twice_ly (length abc_twice) - Suc 0)))" + apply(rule_tac t_correct_termi, auto) + done + hence "list_all (\(acn, s). s \ Suc (length (tm_of abc_twice @ tMp (Suc 0) + (start_of twice_ly (length abc_twice) - Suc 0)) div 2) + 12) + (abacus.tshift (change_termi_state (tm_of abc_twice @ tMp (Suc 0) + (start_of twice_ly (length abc_twice) - Suc 0))) 12)" + apply(rule_tac t_correct_shift, simp) + done + thus "list_all (\(acn, s). s \ + (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2) + (abacus.tshift (change_termi_state (tm_of abc_twice @ tMp (Suc 0) + (start_of twice_ly (length abc_twice) - Suc 0))) 12)" + apply(simp) + apply(simp add: list_all_length, auto) + done +next + have "list_all (\(acn, s). s \ Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) + (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) div 2)) + (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) + (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) " + apply(rule_tac t_correct_termi, auto) + done + hence "list_all (\(acn, s). s \ Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) + (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) div 2) + (t_twice_len + 13)) + (abacus.tshift (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) + (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) (t_twice_len + 13))" + apply(rule_tac t_correct_shift, simp) + done + thus "list_all (\(acn, s). s \ (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2) + (abacus.tshift (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) + (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) (t_twice_len + 13))" + apply(simp add: t_twice_len_def t_twice_def) + using twice_len_even fourtimes_len_even + apply(auto simp: list_all_length) + done +qed + +lemma [intro]: "t_correct (t_wcode_prepare |+| t_wcode_main)" +apply(auto intro: t_correct_add) +done + +lemma prepare_mainpart_lemma: + "args \ [] \ + \ stp ln rn. steps (Suc 0, [], ) (t_wcode_prepare |+| t_wcode_main) stp + = (0, Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin ()\<^esup> @ Bk\<^bsup>rn\<^esup>)" +proof - + let ?P1 = "\ (l, r). l = [] \ r = " + let ?Q1 = "\ (l, r). wprepare_stop m args (l, r)" + let ?P2 = ?Q1 + let ?Q2 = "\ (l, r). (\ ln rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin ()\<^esup> @ Bk\<^bsup>rn\<^esup>)" + let ?P3 = "\ tp. False" + assume h: "args \ []" + have "?P1 \-> \ tp. (\ stp tp'. steps (Suc 0, tp) + (t_wcode_prepare |+| t_wcode_main) stp = (0, tp') \ ?Q2 tp')" + proof(rule_tac turing_merge.t_merge_halt[of t_wcode_prepare t_wcode_main ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], + auto simp: turing_merge_def) + show "\stp. case steps (Suc 0, [], ) t_wcode_prepare stp of (st, tp') + \ st = 0 \ wprepare_stop m args tp'" + using wprepare_correctness[of args m] h + apply(simp, auto) + apply(rule_tac x = n in exI, simp add: wprepare_inv.simps) + done + next + fix a b + assume "wprepare_stop m args (a, b)" + thus "\stp. case steps (Suc 0, a, b) t_wcode_main stp of + (st, tp') \ (st = 0) \ (case tp' of (l, r) \ l = Bk # Oc\<^bsup>Suc m\<^esup> \ + (\ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin ()\<^esup> @ Bk\<^bsup>rn\<^esup>))" + proof(simp only: wprepare_stop.simps, erule_tac exE) + fix rn + assume "a = Bk # @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \ + b = Bk # Oc # Bk\<^bsup>rn\<^esup>" + thus "?thesis" + using t_wcode_main_lemma_pre[of "args" "" 0 "Oc\<^bsup>Suc m\<^esup>" 0 rn] h + apply(simp) + apply(erule_tac exE)+ + apply(rule_tac x = stp in exI, simp add: tape_of_nl_rev, auto) + done + qed + next + show "wprepare_stop m args \-> wprepare_stop m args" + by(simp add: t_imply_def) + qed + thus "\ stp ln rn. steps (Suc 0, [], ) (t_wcode_prepare |+| t_wcode_main) stp + = (0, Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin ()\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(simp add: t_imply_def) + apply(erule_tac exE)+ + apply(auto) + done +qed + + +lemma [simp]: "tinres r r' \ + fetch t ss (case r of [] \ Bk | x # xs \ x) = + fetch t ss (case r' of [] \ Bk | x # xs \ x)" +apply(simp add: fetch.simps, auto split: if_splits simp: tinres_def) +apply(case_tac [!] r', simp_all) +apply(case_tac [!] n, simp_all add: exp_ind_def) +apply(case_tac [!] r, simp_all) +done + +lemma [intro]: "\ n. (a::block)\<^bsup>n\<^esup> = []" +by auto + +lemma [simp]: "\tinres r r'; r \ []; r' \ []\ \ hd r = hd r'" +apply(auto simp: tinres_def) +done + +lemma [intro]: "hd (Bk\<^bsup>Suc n\<^esup>) = Bk" +apply(simp add: exp_ind_def) +done + +lemma [simp]: "\tinres r []; r \ []\ \ hd r = Bk" +apply(auto simp: tinres_def) +apply(case_tac n, auto) +done + +lemma [simp]: "\tinres [] r'; r' \ []\ \ hd r' = Bk" +apply(auto simp: tinres_def) +done + +lemma [intro]: "\na. tl r = tl (r @ Bk\<^bsup>n\<^esup>) @ Bk\<^bsup>na\<^esup> \ tl (r @ Bk\<^bsup>n\<^esup>) = tl r @ Bk\<^bsup>na\<^esup>" +apply(case_tac r, simp) +apply(case_tac n, simp) +apply(rule_tac x = 0 in exI, simp) +apply(rule_tac x = nat in exI, simp add: exp_ind_def) +apply(simp) +apply(rule_tac x = n in exI, simp) +done + +lemma [simp]: "tinres r r' \ tinres (tl r) (tl r')" +apply(auto simp: tinres_def) +apply(case_tac r', simp_all) +apply(case_tac n, simp_all add: exp_ind_def) +apply(rule_tac x = 0 in exI, simp) +apply(rule_tac x = nat in exI, simp_all) +apply(rule_tac x = n in exI, simp) +done + +lemma [simp]: "\tinres r []; r \ []\ \ tinres (tl r) []" +apply(case_tac r, auto simp: tinres_def) +apply(case_tac n, simp_all add: exp_ind_def) +apply(rule_tac x = nat in exI, simp) +done + +lemma [simp]: "\tinres [] r'\ \ tinres [] (tl r')" +apply(case_tac r', auto simp: tinres_def) +apply(case_tac n, simp_all add: exp_ind_def) +apply(rule_tac x = nat in exI, simp) +done + +lemma [simp]: "tinres r r' \ tinres (b # r) (b # r')" +apply(auto simp: tinres_def) +done + +lemma tinres_step2: + "\tinres r r'; tstep (ss, l, r) t = (sa, la, ra); tstep (ss, l, r') t = (sb, lb, rb)\ + \ la = lb \ tinres ra rb \ sa = sb" +apply(case_tac "ss = 0", simp add: tstep_0) +apply(simp add: tstep.simps [simp del]) +apply(case_tac "fetch t ss (case r of [] \ Bk | x # xs \ x)", simp) +apply(auto simp: new_tape.simps) +apply(simp_all split: taction.splits if_splits) +apply(auto) +done + + +lemma tinres_steps2: + "\tinres r r'; steps (ss, l, r) t stp = (sa, la, ra); steps (ss, l, r') t stp = (sb, lb, rb)\ + \ la = lb \ tinres ra rb \ sa = sb" +apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps) +apply(simp add: tstep_red) +apply(case_tac "(steps (ss, l, r) t stp)") +apply(case_tac "(steps (ss, l, r') t stp)") +proof - + fix stp sa la ra sb lb rb a b c aa ba ca + assume ind: "\sa la ra sb lb rb. \steps (ss, l, r) t stp = (sa, la, ra); + steps (ss, l, r') t stp = (sb, lb, rb)\ \ la = lb \ tinres ra rb \ sa = sb" + and h: " tinres r r'" "tstep (steps (ss, l, r) t stp) t = (sa, la, ra)" + "tstep (steps (ss, l, r') t stp) t = (sb, lb, rb)" "steps (ss, l, r) t stp = (a, b, c)" + "steps (ss, l, r') t stp = (aa, ba, ca)" + have "b = ba \ tinres c ca \ a = aa" + apply(rule_tac ind, simp_all add: h) + done + thus "la = lb \ tinres ra rb \ sa = sb" + apply(rule_tac l = b and r = c and ss = a and r' = ca + and t = t in tinres_step2) + using h + apply(simp, simp, simp) + done +qed + + +text{**************Begin: adjust***************************} +definition t_wcode_adjust :: "tprog" + where + "t_wcode_adjust = [(W1, 1), (R, 2), (Nop, 2), (R, 3), (R, 3), (R, 4), + (L, 8), (L, 5), (L, 6), (W0, 5), (L, 6), (R, 7), + (W1, 2), (Nop, 7), (L, 9), (W0, 8), (L, 9), (L, 10), + (L, 11), (L, 10), (R, 0), (L, 11)]" + +lemma [simp]: "fetch t_wcode_adjust (Suc 0) Bk = (W1, 1)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust (Suc 0) Oc = (R, 2)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust (Suc (Suc 0)) Oc = (R, 3)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust (Suc (Suc (Suc 0))) Oc = (R, 4)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust (Suc (Suc (Suc 0))) Bk = (R, 3)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 4 Bk = (L, 8)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 4 Oc = (L, 5)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 5 Oc = (W0, 5)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 5 Bk = (L, 6)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 6 Oc = (R, 7)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 6 Bk = (L, 6)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 7 Bk = (W1, 2)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 8 Bk = (L, 9)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 8 Oc = (W0, 8)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 9 Oc = (L, 10)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 9 Bk = (L, 9)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 10 Bk = (L, 11)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 10 Oc = (L, 10)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 11 Oc = (L, 11)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +lemma [simp]: "fetch t_wcode_adjust 11 Bk = (R, 0)" +apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) +done + +fun wadjust_start :: "nat \ nat \ tape \ bool" + where + "wadjust_start m rs (l, r) = + (\ ln rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \ + tl r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wadjust_loop_start :: "nat \ nat \ tape \ bool" + where + "wadjust_loop_start m rs (l, r) = + (\ ln rn ml mr. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr = Suc (Suc rs) \ mr > 0)" + +fun wadjust_loop_right_move :: "nat \ nat \ tape \ bool" + where + "wadjust_loop_right_move m rs (l, r) = + (\ ml mr nl nr rn. l = Bk\<^bsup>nl\<^esup> @ Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Bk\<^bsup>nr\<^esup> @ Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr = Suc (Suc rs) \ mr > 0 \ + nl + nr > 0)" + +fun wadjust_loop_check :: "nat \ nat \ tape \ bool" + where + "wadjust_loop_check m rs (l, r) = + (\ ml mr ln rn. l = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ ml + mr = (Suc rs))" + +fun wadjust_loop_erase :: "nat \ nat \ tape \ bool" + where + "wadjust_loop_erase m rs (l, r) = + (\ ml mr ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \ + tl r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ ml + mr = (Suc rs) \ mr > 0)" + +fun wadjust_loop_on_left_moving_O :: "nat \ nat \ tape \ bool" + where + "wadjust_loop_on_left_moving_O m rs (l, r) = + (\ ml mr ln rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m \<^esup>\ + r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr = Suc rs \ mr > 0)" + +fun wadjust_loop_on_left_moving_B :: "nat \ nat \ tape \ bool" + where + "wadjust_loop_on_left_moving_B m rs (l, r) = + (\ ml mr nl nr rn. l = Bk\<^bsup>nl\<^esup> @ Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Bk\<^bsup>nr\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr = Suc rs \ mr > 0)" + +fun wadjust_loop_on_left_moving :: "nat \ nat \ tape \ bool" + where + "wadjust_loop_on_left_moving m rs (l, r) = + (wadjust_loop_on_left_moving_O m rs (l, r) \ + wadjust_loop_on_left_moving_B m rs (l, r))" + +fun wadjust_loop_right_move2 :: "nat \ nat \ tape \ bool" + where + "wadjust_loop_right_move2 m rs (l, r) = + (\ ml mr ln rn. l = Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr = Suc rs \ mr > 0)" + +fun wadjust_erase2 :: "nat \ nat \ tape \ bool" + where + "wadjust_erase2 m rs (l, r) = + (\ ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \ + tl r = Bk\<^bsup>rn\<^esup>)" + +fun wadjust_on_left_moving_O :: "nat \ nat \ tape \ bool" + where + "wadjust_on_left_moving_O m rs (l, r) = + (\ rn. l = Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Oc # Bk\<^bsup>rn\<^esup>)" + +fun wadjust_on_left_moving_B :: "nat \ nat \ tape \ bool" + where + "wadjust_on_left_moving_B m rs (l, r) = + (\ ln rn. l = Bk\<^bsup>ln\<^esup> @ Oc # Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Bk\<^bsup>rn\<^esup>)" + +fun wadjust_on_left_moving :: "nat \ nat \ tape \ bool" + where + "wadjust_on_left_moving m rs (l, r) = + (wadjust_on_left_moving_O m rs (l, r) \ + wadjust_on_left_moving_B m rs (l, r))" + +fun wadjust_goon_left_moving_B :: "nat \ nat \ tape \ bool" + where + "wadjust_goon_left_moving_B m rs (l, r) = + (\ rn. l = Oc\<^bsup>Suc m\<^esup> \ + r = Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wadjust_goon_left_moving_O :: "nat \ nat \ tape \ bool" + where + "wadjust_goon_left_moving_O m rs (l, r) = + (\ ml mr rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \ + r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr = Suc (Suc rs) \ mr > 0)" + +fun wadjust_goon_left_moving :: "nat \ nat \ tape \ bool" + where + "wadjust_goon_left_moving m rs (l, r) = + (wadjust_goon_left_moving_B m rs (l, r) \ + wadjust_goon_left_moving_O m rs (l, r))" + +fun wadjust_backto_standard_pos_B :: "nat \ nat \ tape \ bool" + where + "wadjust_backto_standard_pos_B m rs (l, r) = + (\ rn. l = [] \ + r = Bk # Oc\<^bsup>Suc m \<^esup>@ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +fun wadjust_backto_standard_pos_O :: "nat \ nat \ tape \ bool" + where + "wadjust_backto_standard_pos_O m rs (l, r) = + (\ ml mr rn. l = Oc\<^bsup>ml\<^esup> \ + r = Oc\<^bsup>mr\<^esup> @ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup> \ + ml + mr = Suc m \ mr > 0)" + +fun wadjust_backto_standard_pos :: "nat \ nat \ tape \ bool" + where + "wadjust_backto_standard_pos m rs (l, r) = + (wadjust_backto_standard_pos_B m rs (l, r) \ + wadjust_backto_standard_pos_O m rs (l, r))" + +fun wadjust_stop :: "nat \ nat \ tape \ bool" +where + "wadjust_stop m rs (l, r) = + (\ rn. l = [Bk] \ + r = Oc\<^bsup>Suc m \<^esup>@ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" + +declare wadjust_start.simps[simp del] wadjust_loop_start.simps[simp del] + wadjust_loop_right_move.simps[simp del] wadjust_loop_check.simps[simp del] + wadjust_loop_erase.simps[simp del] wadjust_loop_on_left_moving.simps[simp del] + wadjust_loop_right_move2.simps[simp del] wadjust_erase2.simps[simp del] + wadjust_on_left_moving_O.simps[simp del] wadjust_on_left_moving_B.simps[simp del] + wadjust_on_left_moving.simps[simp del] wadjust_goon_left_moving_B.simps[simp del] + wadjust_goon_left_moving_O.simps[simp del] wadjust_goon_left_moving.simps[simp del] + wadjust_backto_standard_pos.simps[simp del] wadjust_backto_standard_pos_B.simps[simp del] + wadjust_backto_standard_pos_O.simps[simp del] wadjust_stop.simps[simp del] + +fun wadjust_inv :: "nat \ nat \ nat \ tape \ bool" + where + "wadjust_inv st m rs (l, r) = + (if st = Suc 0 then wadjust_start m rs (l, r) + else if st = Suc (Suc 0) then wadjust_loop_start m rs (l, r) + else if st = Suc (Suc (Suc 0)) then wadjust_loop_right_move m rs (l, r) + else if st = 4 then wadjust_loop_check m rs (l, r) + else if st = 5 then wadjust_loop_erase m rs (l, r) + else if st = 6 then wadjust_loop_on_left_moving m rs (l, r) + else if st = 7 then wadjust_loop_right_move2 m rs (l, r) + else if st = 8 then wadjust_erase2 m rs (l, r) + else if st = 9 then wadjust_on_left_moving m rs (l, r) + else if st = 10 then wadjust_goon_left_moving m rs (l, r) + else if st = 11 then wadjust_backto_standard_pos m rs (l, r) + else if st = 0 then wadjust_stop m rs (l, r) + else False +)" + +declare wadjust_inv.simps[simp del] + +fun wadjust_phase :: "nat \ t_conf \ nat" + where + "wadjust_phase rs (st, l, r) = + (if st = 1 then 3 + else if st \ 2 \ st \ 7 then 2 + else if st \ 8 \ st \ 11 then 1 + else 0)" + +thm dropWhile.simps + +fun wadjust_stage :: "nat \ t_conf \ nat" + where + "wadjust_stage rs (st, l, r) = + (if st \ 2 \ st \ 7 then + rs - length (takeWhile (\ a. a = Oc) + (tl (dropWhile (\ a. a = Oc) (rev l @ r)))) + else 0)" + +fun wadjust_state :: "nat \ t_conf \ nat" + where + "wadjust_state rs (st, l, r) = + (if st \ 2 \ st \ 7 then 8 - st + else if st \ 8 \ st \ 11 then 12 - st + else 0)" + +fun wadjust_step :: "nat \ t_conf \ nat" + where + "wadjust_step rs (st, l, r) = + (if st = 1 then (if hd r = Bk then 1 + else 0) + else if st = 3 then length r + else if st = 5 then (if hd r = Oc then 1 + else 0) + else if st = 6 then length l + else if st = 8 then (if hd r = Oc then 1 + else 0) + else if st = 9 then length l + else if st = 10 then length l + else if st = 11 then (if hd r = Bk then 0 + else Suc (length l)) + else 0)" + +fun wadjust_measure :: "(nat \ t_conf) \ nat \ nat \ nat \ nat" + where + "wadjust_measure (rs, (st, l, r)) = + (wadjust_phase rs (st, l, r), + wadjust_stage rs (st, l, r), + wadjust_state rs (st, l, r), + wadjust_step rs (st, l, r))" + +definition wadjust_le :: "((nat \ t_conf) \ nat \ t_conf) set" + where "wadjust_le \ (inv_image lex_square wadjust_measure)" + +lemma [intro]: "wf lex_square" +by(auto intro:wf_lex_prod simp: abacus.lex_pair_def lex_square_def + abacus.lex_triple_def) + +lemma wf_wadjust_le[intro]: "wf wadjust_le" +by(auto intro:wf_inv_image simp: wadjust_le_def + abacus.lex_triple_def abacus.lex_pair_def) + +lemma [simp]: "wadjust_start m rs (c, []) = False" +apply(auto simp: wadjust_start.simps) +done + +lemma [simp]: "wadjust_loop_right_move m rs (c, []) \ c \ []" +apply(auto simp: wadjust_loop_right_move.simps) +done + +lemma [simp]: "wadjust_loop_right_move m rs (c, []) + \ wadjust_loop_check m rs (Bk # c, [])" +apply(simp only: wadjust_loop_right_move.simps wadjust_loop_check.simps) +apply(auto) +apply(case_tac [!] mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "wadjust_loop_check m rs (c, []) \ c \ []" +apply(simp only: wadjust_loop_check.simps, auto) +done + +lemma [simp]: "wadjust_loop_start m rs (c, []) = False" +apply(simp add: wadjust_loop_start.simps) +done + +lemma [simp]: "wadjust_loop_right_move m rs (c, []) \ + wadjust_loop_right_move m rs (Bk # c, [])" +apply(simp only: wadjust_loop_right_move.simps) +apply(erule_tac exE)+ +apply(auto) +apply(case_tac [!] mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "wadjust_loop_check m rs (c, []) \ wadjust_erase2 m rs (tl c, [hd c])" +apply(simp only: wadjust_loop_check.simps wadjust_erase2.simps, auto) +apply(case_tac mr, simp_all add: exp_ind_def, auto) +done + +lemma [simp]: " wadjust_loop_erase m rs (c, []) + \ (c = [] \ wadjust_loop_on_left_moving m rs ([], [Bk])) \ + (c \ [] \ wadjust_loop_on_left_moving m rs (tl c, [hd c]))" +apply(simp add: wadjust_loop_erase.simps, auto) +apply(case_tac [!] mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "wadjust_loop_on_left_moving m rs (c, []) = False" +apply(auto simp: wadjust_loop_on_left_moving.simps) +done + + +lemma [simp]: "wadjust_loop_right_move2 m rs (c, []) = False" +apply(auto simp: wadjust_loop_right_move2.simps) +done + +lemma [simp]: "wadjust_erase2 m rs ([], []) = False" +apply(auto simp: wadjust_erase2.simps) +done + +lemma [simp]: "wadjust_on_left_moving_B m rs + (Oc # Oc # Oc\<^bsup>rs\<^esup> @ Bk # Oc # Oc\<^bsup>m\<^esup>, [Bk])" +apply(simp add: wadjust_on_left_moving_B.simps, auto) +apply(rule_tac x = 0 in exI, simp add: exp_ind_def) +done + +lemma [simp]: "wadjust_on_left_moving_B m rs + (Bk\<^bsup>n\<^esup> @ Bk # Oc # Oc # Oc\<^bsup>rs\<^esup> @ Bk # Oc # Oc\<^bsup>m\<^esup>, [Bk])" +apply(simp add: wadjust_on_left_moving_B.simps exp_ind_def, auto) +apply(rule_tac x = "Suc n" in exI, simp add: exp_ind) +done + +lemma [simp]: "\wadjust_erase2 m rs (c, []); c \ []\ \ + wadjust_on_left_moving m rs (tl c, [hd c])" +apply(simp only: wadjust_erase2.simps) +apply(erule_tac exE)+ +apply(case_tac ln, simp_all add: exp_ind_def wadjust_on_left_moving.simps) +done + +lemma [simp]: "wadjust_erase2 m rs (c, []) + \ (c = [] \ wadjust_on_left_moving m rs ([], [Bk])) \ + (c \ [] \ wadjust_on_left_moving m rs (tl c, [hd c]))" +apply(auto) +done + +lemma [simp]: "wadjust_on_left_moving m rs ([], []) = False" +apply(simp add: wadjust_on_left_moving.simps + wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps) +done + +lemma [simp]: "wadjust_on_left_moving_O m rs (c, []) = False" +apply(simp add: wadjust_on_left_moving_O.simps) +done + +lemma [simp]: " \wadjust_on_left_moving_B m rs (c, []); c \ []; hd c = Bk\ \ + wadjust_on_left_moving_B m rs (tl c, [Bk])" +apply(simp add: wadjust_on_left_moving_B.simps, auto) +apply(case_tac [!] ln, simp_all add: exp_ind_def, auto) +done + +lemma [simp]: "\wadjust_on_left_moving_B m rs (c, []); c \ []; hd c = Oc\ \ + wadjust_on_left_moving_O m rs (tl c, [Oc])" +apply(simp add: wadjust_on_left_moving_B.simps wadjust_on_left_moving_O.simps, auto) +apply(case_tac [!] ln, simp_all add: exp_ind_def) +done + +lemma [simp]: "\wadjust_on_left_moving m rs (c, []); c \ []\ \ + wadjust_on_left_moving m rs (tl c, [hd c])" +apply(simp add: wadjust_on_left_moving.simps) +apply(case_tac "hd c", simp_all) +done + +lemma [simp]: "wadjust_on_left_moving m rs (c, []) + \ (c = [] \ wadjust_on_left_moving m rs ([], [Bk])) \ + (c \ [] \ wadjust_on_left_moving m rs (tl c, [hd c]))" +apply(auto) +done + +lemma [simp]: "wadjust_goon_left_moving m rs (c, []) = False" +apply(auto simp: wadjust_goon_left_moving.simps wadjust_goon_left_moving_B.simps + wadjust_goon_left_moving_O.simps) +done + +lemma [simp]: "wadjust_backto_standard_pos m rs (c, []) = False" +apply(auto simp: wadjust_backto_standard_pos.simps + wadjust_backto_standard_pos_B.simps wadjust_backto_standard_pos_O.simps) +done + +lemma [simp]: + "wadjust_start m rs (c, Bk # list) \ + (c = [] \ wadjust_start m rs ([], Oc # list)) \ + (c \ [] \ wadjust_start m rs (c, Oc # list))" +apply(auto simp: wadjust_start.simps) +done + +lemma [simp]: "wadjust_loop_start m rs (c, Bk # list) = False" +apply(auto simp: wadjust_loop_start.simps) +done + +lemma [simp]: "wadjust_loop_right_move m rs (c, b) \ c \ []" +apply(simp only: wadjust_loop_right_move.simps, auto) +done + +lemma [simp]: "wadjust_loop_right_move m rs (c, Bk # list) + \ wadjust_loop_right_move m rs (Bk # c, list)" +apply(simp only: wadjust_loop_right_move.simps) +apply(erule_tac exE)+ +apply(rule_tac x = ml in exI, simp) +apply(rule_tac x = mr in exI, simp) +apply(rule_tac x = "Suc nl" in exI, simp add: exp_ind_def) +apply(case_tac nr, simp, case_tac mr, simp_all add: exp_ind_def) +apply(rule_tac x = nat in exI, auto) +done + +lemma [simp]: "wadjust_loop_check m rs (c, b) \ c \ []" +apply(simp only: wadjust_loop_check.simps, auto) +done + +lemma [simp]: "wadjust_loop_check m rs (c, Bk # list) + \ wadjust_erase2 m rs (tl c, hd c # Bk # list)" +apply(auto simp: wadjust_loop_check.simps wadjust_erase2.simps) +apply(case_tac [!] mr, simp_all add: exp_ind_def, auto) +done + +lemma [simp]: "wadjust_loop_erase m rs (c, b) \ c \ []" +apply(simp only: wadjust_loop_erase.simps, auto) +done + +declare wadjust_loop_on_left_moving_O.simps[simp del] + wadjust_loop_on_left_moving_B.simps[simp del] + +lemma [simp]: "\wadjust_loop_erase m rs (c, Bk # list); hd c = Bk\ + \ wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)" +apply(simp only: wadjust_loop_erase.simps + wadjust_loop_on_left_moving_B.simps) +apply(erule_tac exE)+ +apply(rule_tac x = ml in exI, rule_tac x = mr in exI, + rule_tac x = ln in exI, rule_tac x = 0 in exI, simp) +apply(case_tac ln, simp_all add: exp_ind_def, auto) +apply(simp add: exp_ind exp_ind_def[THEN sym]) +done + +lemma [simp]: "\wadjust_loop_erase m rs (c, Bk # list); c \ []; hd c = Oc\ \ + wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)" +apply(simp only: wadjust_loop_erase.simps wadjust_loop_on_left_moving_O.simps, + auto) +apply(case_tac [!] ln, simp_all add: exp_ind_def) +done + +lemma [simp]: "\wadjust_loop_erase m rs (c, Bk # list); c \ []\ \ + wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)" +apply(case_tac "hd c", simp_all add:wadjust_loop_on_left_moving.simps) +done + +lemma [simp]: "wadjust_loop_on_left_moving m rs (c, b) \ c \ []" +apply(simp add: wadjust_loop_on_left_moving.simps +wadjust_loop_on_left_moving_O.simps wadjust_loop_on_left_moving_B.simps, auto) +done + +lemma [simp]: "wadjust_loop_on_left_moving_O m rs (c, Bk # list) = False" +apply(simp add: wadjust_loop_on_left_moving_O.simps) +done + +lemma [simp]: "\wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Bk\ + \ wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)" +apply(simp only: wadjust_loop_on_left_moving_B.simps) +apply(erule_tac exE)+ +apply(rule_tac x = ml in exI, rule_tac x = mr in exI) +apply(case_tac nl, simp_all add: exp_ind_def, auto) +apply(rule_tac x = "Suc nr" in exI, auto simp: exp_ind_def) +done + +lemma [simp]: "\wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Oc\ + \ wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)" +apply(simp only: wadjust_loop_on_left_moving_O.simps + wadjust_loop_on_left_moving_B.simps) +apply(erule_tac exE)+ +apply(rule_tac x = ml in exI, rule_tac x = mr in exI) +apply(case_tac nl, simp_all add: exp_ind_def, auto) +done + +lemma [simp]: "wadjust_loop_on_left_moving m rs (c, Bk # list) + \ wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)" +apply(simp add: wadjust_loop_on_left_moving.simps) +apply(case_tac "hd c", simp_all) +done + +lemma [simp]: "wadjust_loop_right_move2 m rs (c, b) \ c \ []" +apply(simp only: wadjust_loop_right_move2.simps, auto) +done + +lemma [simp]: "wadjust_loop_right_move2 m rs (c, Bk # list) \ wadjust_loop_start m rs (c, Oc # list)" +apply(auto simp: wadjust_loop_right_move2.simps wadjust_loop_start.simps) +apply(case_tac ln, simp_all add: exp_ind_def) +apply(rule_tac x = 0 in exI, simp) +apply(rule_tac x = rn in exI, simp) +apply(rule_tac x = "Suc ml" in exI, simp add: exp_ind_def, auto) +apply(rule_tac x = "Suc nat" in exI, simp add: exp_ind) +apply(rule_tac x = rn in exI, auto) +apply(rule_tac x = "Suc ml" in exI, auto simp: exp_ind_def) +done + +lemma [simp]: "wadjust_erase2 m rs (c, Bk # list) \ c \ []" +apply(auto simp:wadjust_erase2.simps ) +done + +lemma [simp]: "wadjust_erase2 m rs (c, Bk # list) \ + wadjust_on_left_moving m rs (tl c, hd c # Bk # list)" +apply(auto simp: wadjust_erase2.simps) +apply(case_tac ln, simp_all add: exp_ind_def wadjust_on_left_moving.simps + wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps) +apply(auto) +apply(rule_tac x = "(Suc (Suc rn))" in exI, simp add: exp_ind_def) +apply(rule_tac x = "Suc nat" in exI, simp add: exp_ind) +apply(rule_tac x = "(Suc (Suc rn))" in exI, simp add: exp_ind_def) +done + +lemma [simp]: "wadjust_on_left_moving m rs (c,b) \ c \ []" +apply(simp only:wadjust_on_left_moving.simps + wadjust_on_left_moving_O.simps + wadjust_on_left_moving_B.simps + , auto) +done + +lemma [simp]: "wadjust_on_left_moving_O m rs (c, Bk # list) = False" +apply(simp add: wadjust_on_left_moving_O.simps) +done + +lemma [simp]: "\wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Bk\ + \ wadjust_on_left_moving_B m rs (tl c, Bk # Bk # list)" +apply(auto simp: wadjust_on_left_moving_B.simps) +apply(case_tac ln, simp_all add: exp_ind_def, auto) +done + +lemma [simp]: "\wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Oc\ + \ wadjust_on_left_moving_O m rs (tl c, Oc # Bk # list)" +apply(auto simp: wadjust_on_left_moving_O.simps + wadjust_on_left_moving_B.simps) +apply(case_tac ln, simp_all add: exp_ind_def) +done + +lemma [simp]: "wadjust_on_left_moving m rs (c, Bk # list) \ + wadjust_on_left_moving m rs (tl c, hd c # Bk # list)" +apply(simp add: wadjust_on_left_moving.simps) +apply(case_tac "hd c", simp_all) +done + +lemma [simp]: "wadjust_goon_left_moving m rs (c, b) \ c \ []" +apply(simp add: wadjust_goon_left_moving.simps + wadjust_goon_left_moving_B.simps + wadjust_goon_left_moving_O.simps exp_ind_def, auto) +done + +lemma [simp]: "wadjust_goon_left_moving_O m rs (c, Bk # list) = False" +apply(simp add: wadjust_goon_left_moving_O.simps, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "\wadjust_goon_left_moving_B m rs (c, Bk # list); hd c = Bk\ + \ wadjust_backto_standard_pos_B m rs (tl c, Bk # Bk # list)" +apply(auto simp: wadjust_goon_left_moving_B.simps + wadjust_backto_standard_pos_B.simps exp_ind_def) +done + +lemma [simp]: "\wadjust_goon_left_moving_B m rs (c, Bk # list); hd c = Oc\ + \ wadjust_backto_standard_pos_O m rs (tl c, Oc # Bk # list)" +apply(auto simp: wadjust_goon_left_moving_B.simps + wadjust_backto_standard_pos_O.simps exp_ind_def) +apply(rule_tac x = m in exI, simp, auto) +done + +lemma [simp]: "wadjust_goon_left_moving m rs (c, Bk # list) \ + wadjust_backto_standard_pos m rs (tl c, hd c # Bk # list)" +apply(case_tac "hd c", simp_all add: wadjust_backto_standard_pos.simps + wadjust_goon_left_moving.simps) +done + +lemma [simp]: "wadjust_backto_standard_pos m rs (c, Bk # list) \ + (c = [] \ wadjust_stop m rs ([Bk], list)) \ (c \ [] \ wadjust_stop m rs (Bk # c, list))" +apply(auto simp: wadjust_backto_standard_pos.simps + wadjust_backto_standard_pos_B.simps + wadjust_backto_standard_pos_O.simps wadjust_stop.simps) +apply(case_tac [!] mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "wadjust_start m rs (c, Oc # list) + \ (c = [] \ wadjust_loop_start m rs ([Oc], list)) \ + (c \ [] \ wadjust_loop_start m rs (Oc # c, list))" +apply(auto simp:wadjust_loop_start.simps wadjust_start.simps ) +apply(rule_tac x = ln in exI, rule_tac x = rn in exI, + rule_tac x = "Suc 0" in exI, simp) +done + +lemma [simp]: "wadjust_loop_start m rs (c, b) \ c \ []" +apply(simp add: wadjust_loop_start.simps, auto) +done + +lemma [simp]: "wadjust_loop_start m rs (c, Oc # list) + \ wadjust_loop_right_move m rs (Oc # c, list)" +apply(simp add: wadjust_loop_start.simps wadjust_loop_right_move.simps, auto) +apply(rule_tac x = ml in exI, rule_tac x = mr in exI, + rule_tac x = 0 in exI, simp) +apply(rule_tac x = "Suc ln" in exI, simp add: exp_ind, auto) +done + +lemma [simp]: "wadjust_loop_right_move m rs (c, Oc # list) \ + wadjust_loop_check m rs (Oc # c, list)" +apply(simp add: wadjust_loop_right_move.simps + wadjust_loop_check.simps, auto) +apply(rule_tac [!] x = ml in exI, simp_all, auto) +apply(case_tac nl, auto simp: exp_ind_def) +apply(rule_tac x = "mr - 1" in exI, case_tac mr, simp_all add: exp_ind_def) +apply(case_tac [!] nr, simp_all add: exp_ind_def, auto) +done + +lemma [simp]: "wadjust_loop_check m rs (c, Oc # list) \ + wadjust_loop_erase m rs (tl c, hd c # Oc # list)" +apply(simp only: wadjust_loop_check.simps wadjust_loop_erase.simps) +apply(erule_tac exE)+ +apply(rule_tac x = ml in exI, rule_tac x = mr in exI, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +apply(case_tac rn, simp_all add: exp_ind_def) +done + +lemma [simp]: "wadjust_loop_erase m rs (c, Oc # list) \ + wadjust_loop_erase m rs (c, Bk # list)" +apply(auto simp: wadjust_loop_erase.simps) +done + +lemma [simp]: "wadjust_loop_on_left_moving_B m rs (c, Oc # list) = False" +apply(auto simp: wadjust_loop_on_left_moving_B.simps) +apply(case_tac nr, simp_all add: exp_ind_def) +done + +lemma [simp]: "wadjust_loop_on_left_moving m rs (c, Oc # list) + \ wadjust_loop_right_move2 m rs (Oc # c, list)" +apply(simp add:wadjust_loop_on_left_moving.simps) +apply(auto simp: wadjust_loop_on_left_moving_O.simps + wadjust_loop_right_move2.simps) +done + +lemma [simp]: "wadjust_loop_right_move2 m rs (c, Oc # list) = False" +apply(auto simp: wadjust_loop_right_move2.simps ) +apply(case_tac ln, simp_all add: exp_ind_def) +done + +lemma [simp]: "wadjust_erase2 m rs (c, Oc # list) + \ (c = [] \ wadjust_erase2 m rs ([], Bk # list)) + \ (c \ [] \ wadjust_erase2 m rs (c, Bk # list))" +apply(auto simp: wadjust_erase2.simps ) +done + +lemma [simp]: "wadjust_on_left_moving_B m rs (c, Oc # list) = False" +apply(auto simp: wadjust_on_left_moving_B.simps) +done + +lemma [simp]: "\wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Bk\ \ + wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)" +apply(auto simp: wadjust_on_left_moving_O.simps + wadjust_goon_left_moving_B.simps exp_ind_def) +done + +lemma [simp]: "\wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Oc\ + \ wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)" +apply(auto simp: wadjust_on_left_moving_O.simps + wadjust_goon_left_moving_O.simps exp_ind_def) +apply(rule_tac x = rs in exI, simp) +apply(auto simp: exp_ind_def numeral_2_eq_2) +done + + +lemma [simp]: "wadjust_on_left_moving m rs (c, Oc # list) \ + wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)" +apply(simp add: wadjust_on_left_moving.simps + wadjust_goon_left_moving.simps) +apply(case_tac "hd c", simp_all) +done + +lemma [simp]: "wadjust_on_left_moving m rs (c, Oc # list) \ + wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)" +apply(simp add: wadjust_on_left_moving.simps + wadjust_goon_left_moving.simps) +apply(case_tac "hd c", simp_all) +done + +lemma [simp]: "wadjust_goon_left_moving_B m rs (c, Oc # list) = False" +apply(auto simp: wadjust_goon_left_moving_B.simps) +done + +lemma [simp]: "\wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Bk\ + \ wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)" +apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps) +apply(case_tac [!] ml, auto simp: exp_ind_def) +done + +lemma [simp]: "\wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Oc\ \ + wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)" +apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps) +apply(rule_tac x = "ml - 1" in exI, simp) +apply(case_tac ml, simp_all add: exp_ind_def) +apply(rule_tac x = "Suc mr" in exI, auto simp: exp_ind_def) +done + +lemma [simp]: "wadjust_goon_left_moving m rs (c, Oc # list) \ + wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)" +apply(simp add: wadjust_goon_left_moving.simps) +apply(case_tac "hd c", simp_all) +done + +lemma [simp]: "wadjust_backto_standard_pos_B m rs (c, Oc # list) = False" +apply(simp add: wadjust_backto_standard_pos_B.simps) +done + +lemma [simp]: "wadjust_backto_standard_pos_O m rs (c, Bk # xs) = False" +apply(simp add: wadjust_backto_standard_pos_O.simps, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +done + + + +lemma [simp]: "wadjust_backto_standard_pos_O m rs ([], Oc # list) \ + wadjust_backto_standard_pos_B m rs ([], Bk # Oc # list)" +apply(auto simp: wadjust_backto_standard_pos_O.simps + wadjust_backto_standard_pos_B.simps) +apply(rule_tac x = rn in exI, simp) +apply(case_tac ml, simp_all add: exp_ind_def) +done + + +lemma [simp]: + "\wadjust_backto_standard_pos_O m rs (c, Oc # list); c \ []; hd c = Bk\ + \ wadjust_backto_standard_pos_B m rs (tl c, Bk # Oc # list)" +apply(simp add:wadjust_backto_standard_pos_O.simps + wadjust_backto_standard_pos_B.simps, auto) +apply(case_tac [!] ml, simp_all add: exp_ind_def) +done + +lemma [simp]: "\wadjust_backto_standard_pos_O m rs (c, Oc # list); c \ []; hd c = Oc\ + \ wadjust_backto_standard_pos_O m rs (tl c, Oc # Oc # list)" +apply(simp add: wadjust_backto_standard_pos_O.simps, auto) +apply(case_tac ml, simp_all add: exp_ind_def, auto) +apply(rule_tac x = nat in exI, auto simp: exp_ind_def) +done + +lemma [simp]: "wadjust_backto_standard_pos m rs (c, Oc # list) + \ (c = [] \ wadjust_backto_standard_pos m rs ([], Bk # Oc # list)) \ + (c \ [] \ wadjust_backto_standard_pos m rs (tl c, hd c # Oc # list))" +apply(auto simp: wadjust_backto_standard_pos.simps) +apply(case_tac "hd c", simp_all) +done +thm wadjust_loop_right_move.simps + +lemma [simp]: "wadjust_loop_right_move m rs (c, []) = False" +apply(simp only: wadjust_loop_right_move.simps) +apply(rule_tac iffI) +apply(erule_tac exE)+ +apply(case_tac nr, simp_all add: exp_ind_def) +apply(case_tac mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "wadjust_loop_erase m rs (c, []) = False" +apply(simp only: wadjust_loop_erase.simps, auto) +apply(case_tac mr, simp_all add: exp_ind_def) +done + +lemma [simp]: "\Suc (Suc rs) = a; wadjust_loop_erase m rs (c, Bk # list)\ + \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) + < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) \ + a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) = + a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list))))" +apply(simp only: wadjust_loop_erase.simps) +apply(rule_tac disjI2) +apply(case_tac c, simp, simp) +done + +lemma [simp]: + "\Suc (Suc rs) = a; wadjust_loop_on_left_moving m rs (c, Bk # list)\ + \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) + < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) \ + a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) = + a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list))))" +apply(subgoal_tac "c \ []") +apply(case_tac c, simp_all) +done + +lemma dropWhile_exp1: "dropWhile (\a. a = Oc) (Oc\<^bsup>n\<^esup> @ xs) = dropWhile (\a. a = Oc) xs" +apply(induct n, simp_all add: exp_ind_def) +done +lemma takeWhile_exp1: "takeWhile (\a. a = Oc) (Oc\<^bsup>n\<^esup> @ xs) = Oc\<^bsup>n\<^esup> @ takeWhile (\a. a = Oc) xs" +apply(induct n, simp_all add: exp_ind_def) +done + +lemma [simp]: "\Suc (Suc rs) = a; wadjust_loop_right_move2 m rs (c, Bk # list)\ + \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list)))) + < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list))))" +apply(simp add: wadjust_loop_right_move2.simps, auto) +apply(simp add: dropWhile_exp1 takeWhile_exp1) +apply(case_tac ln, simp, simp add: exp_ind_def) +done + +lemma [simp]: "wadjust_loop_check m rs ([], b) = False" +apply(simp add: wadjust_loop_check.simps) +done + +lemma [simp]: "\Suc (Suc rs) = a; wadjust_loop_check m rs (c, Oc # list)\ + \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Oc # list)))) + < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list)))) \ + a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Oc # list)))) = + a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list))))" +apply(case_tac "c", simp_all) +done + +lemma [simp]: + "\Suc (Suc rs) = a; wadjust_loop_erase m rs (c, Oc # list)\ + \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) + < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list)))) \ + a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) = + a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list))))" +apply(simp add: wadjust_loop_erase.simps) +apply(rule_tac disjI2) +apply(auto) +apply(simp add: dropWhile_exp1 takeWhile_exp1) +done + +declare numeral_2_eq_2[simp del] + +lemma wadjust_correctness: + shows "let P = (\ (len, st, l, r). st = 0) in + let Q = (\ (len, st, l, r). wadjust_inv st m rs (l, r)) in + let f = (\ stp. (Suc (Suc rs), steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, + Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp)) in + \ n .P (f n) \ Q (f n)" +proof - + let ?P = "(\ (len, st, l, r). st = 0)" + let ?Q = "\ (len, st, l, r). wadjust_inv st m rs (l, r)" + let ?f = "\ stp. (Suc (Suc rs), steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, + Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp)" + have "\ n. ?P (?f n) \ ?Q (?f n)" + proof(rule_tac halt_lemma2) + show "wf wadjust_le" by auto + next + show "\ n. \ ?P (?f n) \ ?Q (?f n) \ + ?Q (?f (Suc n)) \ (?f (Suc n), ?f n) \ wadjust_le" + proof(rule_tac allI, rule_tac impI, case_tac "?f n", + simp add: tstep_red tstep.simps, rule_tac conjI, erule_tac conjE, + erule_tac conjE) + fix n a b c d + assume "0 < b" "wadjust_inv b m rs (c, d)" "Suc (Suc rs) = a" + thus "case case fetch t_wcode_adjust b (case d of [] \ Bk | x # xs \ x) + of (ac, ns) \ (ns, new_tape ac (c, d)) of (st, x) \ wadjust_inv st m rs x" + apply(case_tac d, simp, case_tac [2] aa) + apply(simp_all add: wadjust_inv.simps wadjust_le_def new_tape.simps + abacus.lex_triple_def abacus.lex_pair_def lex_square_def + split: if_splits) + done + next + fix n a b c d + assume "0 < b \ wadjust_inv b m rs (c, d)" + "Suc (Suc rs) = a \ steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, + Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust n = (b, c, d)" + thus "((a, case fetch t_wcode_adjust b (case d of [] \ Bk | x # xs \ x) + of (ac, ns) \ (ns, new_tape ac (c, d))), a, b, c, d) \ wadjust_le" + proof(erule_tac conjE, erule_tac conjE, erule_tac conjE) + assume "0 < b" "wadjust_inv b m rs (c, d)" "Suc (Suc rs) = a" + thus "?thesis" + apply(case_tac d, case_tac [2] aa) + apply(simp_all add: wadjust_inv.simps wadjust_le_def new_tape.simps + abacus.lex_triple_def abacus.lex_pair_def lex_square_def + split: if_splits) + done + qed + qed + next + show "?Q (?f 0)" + apply(simp add: steps.simps wadjust_inv.simps wadjust_start.simps) + apply(rule_tac x = ln in exI,auto) + done + next + show "\ ?P (?f 0)" + apply(simp add: steps.simps) + done + qed + thus "?thesis" + apply(auto) + done +qed + +lemma [intro]: "t_correct t_wcode_adjust" +apply(auto simp: t_wcode_adjust_def t_correct.simps iseven_def) +apply(rule_tac x = 11 in exI, simp) +done + +lemma wcode_lemma_pre': + "args \ [] \ + \ stp rn. steps (Suc 0, [], ) + ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp + = (0, [Bk], Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin ())\<^esup> @ Bk\<^bsup>rn\<^esup>)" +proof - + let ?P1 = "\ (l, r). l = [] \ r = " + let ?Q1 = "\(l, r). l = Bk # Oc\<^bsup>Suc m\<^esup> \ + (\ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin ()\<^esup> @ Bk\<^bsup>rn\<^esup>)" + let ?P2 = ?Q1 + let ?Q2 = "\ (l, r). (wadjust_stop m (bl_bin () - 1) (l, r))" + let ?P3 = "\ tp. False" + assume h: "args \ []" + have "?P1 \-> \ tp. (\ stp tp'. steps (Suc 0, tp) + ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp = (0, tp') \ ?Q2 tp')" + proof(rule_tac turing_merge.t_merge_halt[of "t_wcode_prepare |+| t_wcode_main" + t_wcode_adjust ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], + auto simp: turing_merge_def) + + show "\stp. case steps (Suc 0, [], ) (t_wcode_prepare |+| t_wcode_main) stp of + (st, tp') \ st = 0 \ (case tp' of (l, r) \ l = Bk # Oc\<^bsup>Suc m\<^esup> \ + (\ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin ()\<^esup> @ Bk\<^bsup>rn\<^esup>))" + using h prepare_mainpart_lemma[of args m] + apply(auto) + apply(rule_tac x = stp in exI, simp) + apply(rule_tac x = ln in exI, auto) + done + next + fix ln rn + show "\stp. case steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # + Oc\<^bsup>bl_bin ()\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp of + (st, tp') \ st = 0 \ wadjust_stop m (bl_bin () - Suc 0) tp'" + using wadjust_correctness[of m "bl_bin () - 1" "Suc ln" rn] + apply(subgoal_tac "bl_bin () > 0", auto simp: wadjust_inv.simps) + apply(rule_tac x = n in exI, simp add: exp_ind) + using h + apply(case_tac args, simp_all, case_tac list, + simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def + bl_bin.simps) + done + next + show "?Q1 \-> ?P2" + by(simp add: t_imply_def) + qed + thus "\stp rn. steps (Suc 0, [], ) ((t_wcode_prepare |+| t_wcode_main) |+| + t_wcode_adjust) stp = (0, [Bk], Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin ())\<^esup> @ Bk\<^bsup>rn\<^esup>)" + apply(simp add: t_imply_def) + apply(erule_tac exE)+ + apply(subgoal_tac "bl_bin () > 0", auto simp: wadjust_stop.simps) + using h + apply(case_tac args, simp_all, case_tac list, + simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def + bl_bin.simps) + done +qed + +text {* + The initialization TM @{text "t_wcode"}. + *} +definition t_wcode :: "tprog" + where + "t_wcode = (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust" + + +text {* + The correctness of @{text "t_wcode"}. + *} +lemma wcode_lemma_1: + "args \ [] \ + \ stp ln rn. steps (Suc 0, [], ) (t_wcode) stp = + (0, [Bk], Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin ())\<^esup> @ Bk\<^bsup>rn\<^esup>)" +apply(simp add: wcode_lemma_pre' t_wcode_def) +done + +lemma wcode_lemma: + "args \ [] \ + \ stp ln rn. steps (Suc 0, [], ) (t_wcode) stp = + (0, [Bk], <[m ,bl_bin ()]> @ Bk\<^bsup>rn\<^esup>)" +using wcode_lemma_1[of args m] +apply(simp add: t_wcode_def tape_of_nl_abv tape_of_nat_list.simps) +done + +section {* The universal TM @{text "UTM"} *} + +text {* + This section gives the explicit construction of {\em Universal Turing Machine}, defined as @{text "UTM"} and proves its + correctness. It is pretty easy by composing the partial results we have got so far. + *} + + +definition UTM :: "tprog" + where + "UTM = (let (aprog, rs_pos, a_md) = rec_ci rec_F in + let abc_F = aprog [+] dummy_abc (Suc (Suc 0)) in + (t_wcode |+| (tm_of abc_F @ tMp (Suc (Suc 0)) (start_of (layout_of abc_F) + (length abc_F) - Suc 0))))" + +definition F_aprog :: "abc_prog" + where + "F_aprog \ (let (aprog, rs_pos, a_md) = rec_ci rec_F in + aprog [+] dummy_abc (Suc (Suc 0)))" + +definition F_tprog :: "tprog" + where + "F_tprog = tm_of (F_aprog)" + +definition t_utm :: "tprog" + where + "t_utm \ + (F_tprog) @ tMp (Suc (Suc 0)) (start_of (layout_of (F_aprog)) + (length (F_aprog)) - Suc 0)" + +definition UTM_pre :: "tprog" + where + "UTM_pre = t_wcode |+| t_utm" + +lemma F_abc_halt_eq: + "\turing_basic.t_correct tp; + length lm = k; + steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>); + rs > 0\ + \ \ stp m. abc_steps_l (0, [code tp, bl2wc ()]) (F_aprog) stp = + (length (F_aprog), code tp # bl2wc () # (rs - 1) # 0\<^bsup>m\<^esup>)" +apply(drule_tac F_t_halt_eq, simp, simp, simp) +apply(case_tac "rec_ci rec_F") +apply(frule_tac abc_append_dummy_complie, simp, simp, erule_tac exE, + erule_tac exE) +apply(rule_tac x = stp in exI, rule_tac x = m in exI) +apply(simp add: F_aprog_def dummy_abc_def) +done + +lemma F_abc_utm_halt_eq: + "\rs > 0; + abc_steps_l (0, [code tp, bl2wc ()]) F_aprog stp = + (length F_aprog, code tp # bl2wc () # (rs - 1) # 0\<^bsup>m\<^esup>)\ + \ \stp m n.(steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) t_utm stp = + (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>))" + thm abacus_turing_eq_halt + using abacus_turing_eq_halt + [of "layout_of F_aprog" "F_aprog" "F_tprog" "length (F_aprog)" + "[code tp, bl2wc ()]" stp "code tp # bl2wc () # (rs - 1) # 0\<^bsup>m\<^esup>" "Suc (Suc 0)" + "start_of (layout_of (F_aprog)) (length (F_aprog))" "[]" 0] +apply(simp add: F_tprog_def t_utm_def abc_lm_v.simps nth_append) +apply(erule_tac exE)+ +apply(rule_tac x = stpa in exI, rule_tac x = "Suc (Suc ma)" in exI, + rule_tac x = l in exI, simp add: exp_ind) +done + +declare tape_of_nl_abv_cons[simp del] + +lemma t_utm_halt_eq': + "\turing_basic.t_correct tp; + 0 < rs; + steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\ + \ \stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) t_utm stp = + (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" +apply(drule_tac l = l in F_abc_halt_eq, simp, simp, simp) +apply(erule_tac exE, erule_tac exE) +apply(rule_tac F_abc_utm_halt_eq, simp_all) +done + +lemma [simp]: "tinres xs (xs @ Bk\<^bsup>i\<^esup>)" +apply(auto simp: tinres_def) +done + +lemma [elim]: "\rs > 0; Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup> = c @ Bk\<^bsup>n\<^esup>\ + \ \n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" +apply(case_tac "na > n") +apply(subgoal_tac "\ d. na = d + n", auto simp: exp_add) +apply(rule_tac x = "na - n" in exI, simp) +apply(subgoal_tac "\ d. n = d + na", auto simp: exp_add) +apply(case_tac rs, simp_all add: exp_ind, case_tac d, + simp_all add: exp_ind) +apply(rule_tac x = "n - na" in exI, simp) +done + + +lemma t_utm_halt_eq'': + "\turing_basic.t_correct tp; + 0 < rs; + steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\ + \ \stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>) t_utm stp = + (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" +apply(drule_tac t_utm_halt_eq', simp_all) +apply(erule_tac exE)+ +proof - + fix stpa ma na + assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)" + and gr: "rs > 0" + thus "\stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>) t_utm stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" + apply(rule_tac x = stpa in exI, rule_tac x = ma in exI, simp) + proof(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>) t_utm stpa", simp) + fix a b c + assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)" + "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (a, b, c)" + thus " a = 0 \ b = Bk\<^bsup>ma\<^esup> \ (\n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" + using tinres_steps2[of "<[code tp, bl2wc ()]>" "<[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>" + "Suc 0" " [Bk, Bk]" t_utm stpa 0 "Bk\<^bsup>ma\<^esup>" "Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>" a b c] + apply(simp) + using gr + apply(simp only: tinres_def, auto) + apply(rule_tac x = "na + n" in exI, simp add: exp_add) + done + qed +qed + +lemma [simp]: "tinres [Bk, Bk] [Bk]" +apply(auto simp: tinres_def) +done + +lemma [elim]: "Bk\<^bsup>ma\<^esup> = b @ Bk\<^bsup>n\<^esup> \ \m. b = Bk\<^bsup>m\<^esup>" +apply(subgoal_tac "ma = length b + n") +apply(rule_tac x = "ma - n" in exI, simp add: exp_add) +apply(drule_tac length_equal) +apply(simp) +done + +lemma t_utm_halt_eq: + "\turing_basic.t_correct tp; + 0 < rs; + steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\ + \ \stp m n. steps (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>) t_utm stp = + (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" +apply(drule_tac i = i in t_utm_halt_eq'', simp_all) +apply(erule_tac exE)+ +proof - + fix stpa ma na + assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)" + and gr: "rs > 0" + thus "\stp m n. steps (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>) t_utm stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" + apply(rule_tac x = stpa in exI) + proof(case_tac "steps (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>) t_utm stpa", simp) + fix a b c + assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)" + "steps (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (a, b, c)" + thus "a = 0 \ (\m. b = Bk\<^bsup>m\<^esup>) \ (\n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" + using tinres_steps[of "[Bk, Bk]" "[Bk]" "Suc 0" "<[code tp, bl2wc ()]> @ Bk\<^bsup>i\<^esup>" t_utm stpa 0 + "Bk\<^bsup>ma\<^esup>" "Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>" a b c] + apply(simp) + apply(auto simp: tinres_def) + apply(rule_tac x = "ma + n" in exI, simp add: exp_add) + done + qed +qed + +lemma [intro]: "t_correct t_wcode" +apply(simp add: t_wcode_def) +apply(auto) +done + +lemma [intro]: "t_correct t_utm" +apply(simp add: t_utm_def F_tprog_def) +apply(rule_tac t_compiled_correct, auto) +done + +lemma UTM_halt_lemma_pre: + "\turing_basic.t_correct tp; + 0 < rs; + args \ []; + steps (Suc 0, Bk\<^bsup>i\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)\ + \ \stp m n. steps (Suc 0, [], ) UTM_pre stp = + (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" +proof - + let ?Q2 = "\ (l, r). (\ ln rn. l = Bk\<^bsup>ln\<^esup> \ r = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" + term ?Q2 + let ?P1 = "\ (l, r). l = [] \ r = " + let ?Q1 = "\ (l, r). (l = [Bk] \ + (\ rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin ())\<^esup> @ Bk\<^bsup>rn\<^esup>))" + let ?P2 = ?Q1 + let ?P3 = "\ (l, r). False" + assume h: "turing_basic.t_correct tp" "0 < rs" + "args \ []" "steps (Suc 0, Bk\<^bsup>i\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)" + have "?P1 \-> \ tp. (\ stp tp'. steps (Suc 0, tp) + (t_wcode |+| t_utm) stp = (0, tp') \ ?Q2 tp')" + proof(rule_tac turing_merge.t_merge_halt [of "t_wcode" "t_utm" + ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], auto simp: turing_merge_def) + show "\stp. case steps (Suc 0, [], ) t_wcode stp of (st, tp') \ + st = 0 \ (case tp' of (l, r) \ l = [Bk] \ + (\rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin ())\<^esup> @ Bk\<^bsup>rn\<^esup>))" + using wcode_lemma_1[of args "code tp"] h + apply(simp, auto) + apply(rule_tac x = stpa in exI, auto) + done + next + fix rn + show "\stp. case steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ + Bk # Oc\<^bsup>Suc (bl_bin ())\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp of + (st, tp') \ st = 0 \ (case tp' of (l, r) \ + (\ln. l = Bk\<^bsup>ln\<^esup>) \ (\rn. r = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>rn\<^esup>))" + using t_utm_halt_eq[of tp rs i args stp m k rn] h + apply(auto) + apply(rule_tac x = stpa in exI, simp add: bin_wc_eq + tape_of_nat_list.simps tape_of_nl_abv) + apply(auto) + done + next + show "?Q1 \-> ?P2" + apply(simp add: t_imply_def) + done + qed + thus "?thesis" + apply(simp add: t_imply_def) + apply(auto simp: UTM_pre_def) + done +qed + +text {* + The correctness of @{text "UTM"}, the halt case. +*} +theorem UTM_halt_lemma: + "\turing_basic.t_correct tp; + 0 < rs; + args \ []; + steps (Suc 0, Bk\<^bsup>i\<^esup>, ) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)\ + \ \stp m n. steps (Suc 0, [], ) UTM stp = + (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" +using UTM_halt_lemma_pre[of tp rs args i stp m k] +apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def) +apply(case_tac "rec_ci rec_F", simp) +done + +definition TSTD:: "t_conf \ bool" + where + "TSTD c = (let (st, l, r) = c in + st = 0 \ (\ m. l = Bk\<^bsup>m\<^esup>) \ (\ rs n. r = Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>))" + +thm abacus_turing_eq_uhalt + +lemma nstd_case1: "0 < a \ NSTD (trpl_code (a, b, c))" +apply(simp add: NSTD.simps trpl_code.simps) +done + +lemma [simp]: "\m. b \ Bk\<^bsup>m\<^esup> \ 0 < bl2wc b" +apply(rule classical, simp) +apply(induct b, erule_tac x = 0 in allE, simp) +apply(simp add: bl2wc.simps, case_tac a, simp_all + add: bl2nat.simps bl2nat_double) +apply(case_tac "\ m. b = Bk\<^bsup>m\<^esup>", erule exE) +apply(erule_tac x = "Suc m" in allE, simp add: exp_ind_def, simp) +done +lemma nstd_case2: "\m. b \ Bk\<^bsup>m\<^esup> \ NSTD (trpl_code (a, b, c))" +apply(simp add: NSTD.simps trpl_code.simps) +done + +thm lg.simps +thm lgR.simps + +lemma [elim]: "Suc (2 * x) = 2 * y \ RR" +apply(induct x arbitrary: y, simp, simp) +apply(case_tac y, simp, simp) +done + +lemma bl2nat_zero_eq[simp]: "(bl2nat c 0 = 0) = (\n. c = Bk\<^bsup>n\<^esup>)" +apply(auto) +apply(induct c, simp add: bl2nat.simps) +apply(rule_tac x = 0 in exI, simp) +apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) +done + +lemma bl2wc_exp_ex: + "\Suc (bl2wc c) = 2 ^ m\ \ \ rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" +apply(induct c arbitrary: m, simp add: bl2wc.simps bl2nat.simps) +apply(case_tac a, auto) +apply(case_tac m, simp_all add: bl2wc.simps, auto) +apply(rule_tac x = 0 in exI, rule_tac x = "Suc n" in exI, + simp add: exp_ind_def) +apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double) +apply(case_tac m, simp, simp) +proof - + fix c m nat + assume ind: + "\m. Suc (bl2nat c 0) = 2 ^ m \ \rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" + and h: + "Suc (Suc (2 * bl2nat c 0)) = 2 * 2 ^ nat" + have "\rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" + apply(rule_tac m = nat in ind) + using h + apply(simp) + done + from this obtain rs n where " c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" by blast + thus "\rs n. Oc # c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" + apply(rule_tac x = "Suc rs" in exI, simp add: exp_ind_def) + apply(rule_tac x = n in exI, simp) + done +qed + +lemma [elim]: + "\\rs n. c \ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>; + bl2wc c = 2 ^ lg (Suc (bl2wc c)) 2 - Suc 0\ \ bl2wc c = 0" +apply(subgoal_tac "\ m. Suc (bl2wc c) = 2^m", erule_tac exE) +apply(drule_tac bl2wc_exp_ex, simp, erule_tac exE, erule_tac exE) +apply(case_tac rs, simp, simp, erule_tac x = nat in allE, + erule_tac x = n in allE, simp) +using bl2wc_exp_ex[of c "lg (Suc (bl2wc c)) 2"] +apply(case_tac "(2::nat) ^ lg (Suc (bl2wc c)) 2", + simp, simp, erule_tac exE, erule_tac exE, simp) +apply(simp add: bl2wc.simps) +apply(rule_tac x = rs in exI) +apply(case_tac "(2::nat)^rs", simp, simp) +done + +lemma nstd_case3: + "\rs n. c \ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup> \ NSTD (trpl_code (a, b, c))" +apply(simp add: NSTD.simps trpl_code.simps) +apply(rule_tac impI) +apply(rule_tac disjI2, rule_tac disjI2, auto) +done + +lemma NSTD_1: "\ TSTD (a, b, c) + \ rec_exec rec_NSTD [trpl_code (a, b, c)] = Suc 0" + using NSTD_lemma1[of "trpl_code (a, b, c)"] + NSTD_lemma2[of "trpl_code (a, b, c)"] + apply(simp add: TSTD_def) + apply(erule_tac disjE, erule_tac nstd_case1) + apply(erule_tac disjE, erule_tac nstd_case2) + apply(erule_tac nstd_case3) + done + +lemma nonstop_t_uhalt_eq: + "\turing_basic.t_correct tp; + steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp = (a, b, c); + \ TSTD (a, b, c)\ + \ rec_exec rec_nonstop [code tp, bl2wc (), stp] = Suc 0" +apply(simp add: rec_nonstop_def rec_exec.simps) +apply(subgoal_tac + "rec_exec rec_conf [code tp, bl2wc (), stp] = + trpl_code (a, b, c)", simp) +apply(erule_tac NSTD_1) +using rec_t_eq_steps[of tp l lm stp] +apply(simp) +done + +lemma nonstop_true: + "\turing_basic.t_correct tp; + \ stp. (\ TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp))\ + \ \y. rec_calc_rel rec_nonstop + ([code tp, bl2wc (), y]) (Suc 0)" +apply(rule_tac allI, erule_tac x = y in allE) +apply(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp y", simp) +apply(rule_tac nonstop_t_uhalt_eq, simp_all) +done + +(* +lemma [simp]: + "\jturing_basic.t_correct tp; + \ stp. (\ TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp)); + rec_ci rec_F = (F_ap, rs_pos, a_md)\ + \ \ stp. case abc_steps_l (0, [code tp, bl2wc ()] @ 0\<^bsup>a_md - rs_pos \<^esup> + @ suflm) (F_ap) stp of (ss, e) \ ss < length (F_ap)" +apply(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf + ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])])") +apply(simp only: rec_F_def, rule_tac i = 0 and ga = a and gb = b and + gc = c in cn_gi_uhalt, simp, simp, simp, simp, simp, simp, simp) +apply(simp add: ci_cn_para_eq) +apply(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_conf + ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt]))") +apply(rule_tac rf = "(Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf + ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])])" + and n = "Suc (Suc 0)" and f = rec_right and + gs = "[Cn (Suc (Suc 0)) rec_conf + ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])]" + and i = 0 and ga = aa and gb = ba and gc = ca in + cn_gi_uhalt) +apply(simp, simp, simp, simp, simp, simp, simp, + simp add: ci_cn_para_eq) +apply(case_tac "rec_ci rec_halt") +apply(rule_tac rf = "(Cn (Suc (Suc 0)) rec_conf + ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt]))" + and n = "Suc (Suc 0)" and f = "rec_conf" and + gs = "([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])" and + i = "Suc (Suc 0)" and gi = "rec_halt" and ga = ab and gb = bb and + gc = cb in cn_gi_uhalt) +apply(simp, simp, simp, simp, simp add: nth_append, simp, + simp add: nth_append, simp add: rec_halt_def) +apply(simp only: rec_halt_def) +apply(case_tac [!] "rec_ci ((rec_nonstop))") +apply(rule_tac allI, rule_tac impI, simp) +apply(case_tac j, simp) +apply(rule_tac x = "code tp" in exI, rule_tac calc_id, simp, simp, simp, simp) +apply(rule_tac x = "bl2wc ()" in exI, rule_tac calc_id, simp, simp, simp) +apply(rule_tac rf = "Mn (Suc (Suc 0)) (rec_nonstop)" + and f = "(rec_nonstop)" and n = "Suc (Suc 0)" + and aprog' = ac and rs_pos' = bc and a_md' = cc in Mn_unhalt) +apply(simp, simp add: rec_halt_def , simp, simp) +apply(drule_tac nonstop_true, simp_all) +apply(rule_tac allI) +apply(erule_tac x = y in allE)+ +apply(simp) +done + +thm abc_list_crsp_steps + +lemma uabc_uhalt': + "\turing_basic.t_correct tp; + \ stp. (\ TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp)); + rec_ci rec_F = (ap, pos, md)\ + \ \ stp. case abc_steps_l (0, [code tp, bl2wc ()]) ap stp of (ss, e) + \ ss < length ap" +proof(frule_tac F_ap = ap and rs_pos = pos and a_md = md + and suflm = "[]" in F_aprog_uhalt, auto) + fix stp a b + assume h: + "\stp. case abc_steps_l (0, code tp # bl2wc () # 0\<^bsup>md - pos\<^esup>) ap stp of + (ss, e) \ ss < length ap" + "abc_steps_l (0, [code tp, bl2wc ()]) ap stp = (a, b)" + "turing_basic.t_correct tp" + "rec_ci rec_F = (ap, pos, md)" + moreover have "ap \ []" + using h apply(rule_tac rec_ci_not_null, simp) + done + ultimately show "a < length ap" + proof(erule_tac x = stp in allE, + case_tac "abc_steps_l (0, code tp # bl2wc () # 0\<^bsup>md - pos\<^esup>) ap stp", simp) + fix aa ba + assume g: "aa < length ap" + "abc_steps_l (0, code tp # bl2wc () # 0\<^bsup>md - pos\<^esup>) ap stp = (aa, ba)" + "ap \ []" + thus "?thesis" + using abc_list_crsp_steps[of "[code tp, bl2wc ()]" + "md - pos" ap stp aa ba] h + apply(simp) + done + qed +qed + +lemma uabc_uhalt: + "\turing_basic.t_correct tp; + \ stp. (\ TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp))\ + \ \ stp. case abc_steps_l (0, [code tp, bl2wc ()]) F_aprog + stp of (ss, e) \ ss < length F_aprog" +apply(case_tac "rec_ci rec_F", simp add: F_aprog_def) +thm uabc_uhalt' +apply(drule_tac ap = a and pos = b and md = c in uabc_uhalt', simp_all) +proof - + fix a b c + assume + "\stp. case abc_steps_l (0, [code tp, bl2wc ()]) a stp of (ss, e) + \ ss < length a" + "rec_ci rec_F = (a, b, c)" + thus + "\stp. case abc_steps_l (0, [code tp, bl2wc ()]) + (a [+] dummy_abc (Suc (Suc 0))) stp of (ss, e) \ + ss < Suc (Suc (Suc (length a)))" + using abc_append_uhalt1[of a "[code tp, bl2wc ()]" + "a [+] dummy_abc (Suc (Suc 0))" "[]" "dummy_abc (Suc (Suc 0))"] + apply(simp) + done +qed + +thm abacus_turing_eq_uhalt +lemma tutm_uhalt': + "\turing_basic.t_correct tp; + \ stp. (\ TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp))\ + \ \ stp. \ isS0 (steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) t_utm stp)" + using abacus_turing_eq_uhalt[of "layout_of (F_aprog)" + "F_aprog" "F_tprog" "[code tp, bl2wc ()]" + "start_of (layout_of (F_aprog )) (length (F_aprog))" + "Suc (Suc 0)"] +apply(simp add: F_tprog_def) +apply(subgoal_tac "\stp. case abc_steps_l (0, [code tp, bl2wc ()]) + (F_aprog) stp of (as, am) \ as < length (F_aprog)", simp) +thm abacus_turing_eq_uhalt +apply(simp add: t_utm_def F_tprog_def) +apply(rule_tac uabc_uhalt, simp_all) +done + +lemma tinres_commute: "tinres r r' \ tinres r' r" +apply(auto simp: tinres_def) +done + +lemma inres_tape: + "\steps (st, l, r) tp stp = (a, b, c); steps (st, l', r') tp stp = (a', b', c'); + tinres l l'; tinres r r'\ + \ a = a' \ tinres b b' \ tinres c c'" +proof(case_tac "steps (st, l', r) tp stp") + fix aa ba ca + assume h: "steps (st, l, r) tp stp = (a, b, c)" + "steps (st, l', r') tp stp = (a', b', c')" + "tinres l l'" "tinres r r'" + "steps (st, l', r) tp stp = (aa, ba, ca)" + have "tinres b ba \ c = ca \ a = aa" + using h + apply(rule_tac tinres_steps, auto) + done + + thm tinres_steps2 + moreover have "b' = ba \ tinres c' ca \ a' = aa" + using h + apply(rule_tac tinres_steps2, auto intro: tinres_commute) + done + ultimately show "?thesis" + apply(auto intro: tinres_commute) + done +qed + +lemma tape_normalize: "\ stp. \ isS0 (steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) t_utm stp) + \ \ stp. \ isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, <[code tp, bl2wc ()]> @ Bk\<^bsup>n\<^esup>) t_utm stp)" +apply(rule_tac allI, case_tac "(steps (Suc 0, Bk\<^bsup>m\<^esup>, + <[code tp, bl2wc ()]> @ Bk\<^bsup>n\<^esup>) t_utm stp)", simp add: isS0_def) +apply(erule_tac x = stp in allE) +apply(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) t_utm stp", simp) +apply(drule_tac inres_tape, auto) +apply(auto simp: tinres_def) +apply(case_tac "m > Suc (Suc 0)") +apply(rule_tac x = "m - Suc (Suc 0)" in exI) +apply(case_tac m, simp_all add: exp_ind_def, case_tac nat, simp_all add: exp_ind_def) +apply(rule_tac x = "2 - m" in exI, simp add: exp_ind_def[THEN sym] exp_add[THEN sym]) +apply(simp only: numeral_2_eq_2, simp add: exp_ind_def) +done + +lemma tutm_uhalt: + "\turing_basic.t_correct tp; + \ stp. (\ TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp))\ + \ \ stp. \ isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, <[code tp, bl2wc ()]> @ Bk\<^bsup>n\<^esup>) t_utm stp)" +apply(rule_tac tape_normalize) +apply(rule_tac tutm_uhalt', simp_all) +done + +lemma UTM_uhalt_lemma_pre: + "\turing_basic.t_correct tp; + \ stp. (\ TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp)); + args \ []\ + \ \ stp. \ isS0 (steps (Suc 0, [], ) UTM_pre stp)" +proof - + let ?P1 = "\ (l, r). l = [] \ r = " + let ?Q1 = "\ (l, r). (l = [Bk] \ + (\ rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin ())\<^esup> @ Bk\<^bsup>rn\<^esup>))" + let ?P4 = ?Q1 + let ?P3 = "\ (l, r). False" + assume h: "turing_basic.t_correct tp" "\stp. \ TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp)" + "args \ []" + have "?P1 \-> \ tp. \ (\ stp. isS0 (steps (Suc 0, tp) (t_wcode |+| t_utm) stp))" + proof(rule_tac turing_merge.t_merge_uhalt [of "t_wcode" "t_utm" + ?P1 ?P3 ?P3 ?P4 ?Q1 ?P3], auto simp: turing_merge_def) + show "\stp. case steps (Suc 0, [], ) t_wcode stp of (st, tp') \ + st = 0 \ (case tp' of (l, r) \ l = [Bk] \ + (\rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin ())\<^esup> @ Bk\<^bsup>rn\<^esup>))" + using wcode_lemma_1[of args "code tp"] h + apply(simp, auto) + apply(rule_tac x = stp in exI, auto) + done + next + fix rn stp + show " isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin ())\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp) + \ False" + using tutm_uhalt[of tp l args "Suc 0" rn] h + apply(simp) + apply(erule_tac x = stp in allE) + apply(simp add: tape_of_nl_abv tape_of_nat_list.simps bin_wc_eq) + done + next + fix rn stp + show "isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin ())\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp) \ + isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin ())\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp)" + by simp + next + show "?Q1 \-> ?P4" + apply(simp add: t_imply_def) + done + qed + thus "?thesis" + apply(simp add: t_imply_def UTM_pre_def) + done +qed + +text {* + The correctness of @{text "UTM"}, the unhalt case. + *} + +theorem UTM_uhalt_lemma: + "\turing_basic.t_correct tp; + \ stp. (\ TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, ) tp stp)); + args \ []\ + \ \ stp. \ isS0 (steps (Suc 0, [], ) UTM stp)" +using UTM_uhalt_lemma_pre[of tp l args] +apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def) +apply(case_tac "rec_ci rec_F", simp) +done + +end \ No newline at end of file diff -r cbb4ac6c8081 -r 1ce04eb1c8ad utm/abacus.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/utm/abacus.thy Sat Sep 29 12:38:12 2012 +0000 @@ -0,0 +1,7005 @@ +header {* + {\em Abacus} (a kind of register machine) +*} + +theory abacus +imports Main turing_basic +begin + +text {* + {\em Abacus} instructions: +*} + +datatype abc_inst = + -- {* @{text "Inc n"} increments the memory cell (or register) with address @{text "n"} by one. + *} + Inc nat + -- {* + @{text "Dec n label"} decrements the memory cell with address @{text "n"} by one. + If cell @{text "n"} is already zero, no decrements happens and the executio jumps to + the instruction labeled by @{text "label"}. + *} + | Dec nat nat + -- {* + @{text "Goto label"} unconditionally jumps to the instruction labeled by @{text "label"}. + *} + | Goto nat + + +text {* + Abacus programs are defined as lists of Abacus instructions. +*} +type_synonym abc_prog = "abc_inst list" + +section {* + Sample Abacus programs + *} + +text {* + Abacus for addition and clearance. +*} +fun plus_clear :: "nat \ nat \ nat \ abc_prog" + where + "plus_clear m n e = [Dec m e, Inc n, Goto 0]" + +text {* + Abacus for clearing memory untis. +*} +fun clear :: "nat \ nat \ abc_prog" + where + "clear n e = [Dec n e, Goto 0]" + +fun plus:: "nat \ nat \ nat \ nat \ abc_prog" + where + "plus m n p e = [Dec m 4, Inc n, Inc p, + Goto 0, Dec p e, Inc m, Goto 4]" + +fun mult :: "nat \ nat \ nat \ nat \ nat \ abc_prog" + where + "mult m1 m2 n p e = [Dec m1 e]@ plus m1 m2 p 1" + +fun expo :: "nat \ nat \ nat \ nat \ nat \ abc_prog" + where + "expo n m1 m2 p e = [Inc n, Dec m1 e] @ mult m2 n n p 2" + + +text {* + The state of Abacus machine. + *} +type_synonym abc_state = nat + +(* text {* + The memory of Abacus machine is defined as a function from address to contents. +*} +type_synonym abc_mem = "nat \ nat" *) + +text {* + The memory of Abacus machine is defined as a list of contents, with + every units addressed by index into the list. + *} +type_synonym abc_lm = "nat list" + +text {* + Fetching contents out of memory. Units not represented by list elements are considered + as having content @{text "0"}. +*} +fun abc_lm_v :: "abc_lm \ nat \ nat" + where + "abc_lm_v lm n = (if (n < length lm) then (lm!n) else 0)" + +(* +fun abc_l2m :: "abc_lm \ abc_mem" + where + "abc_l2m lm = abc_lm_v lm" +*) + +text {* + Set the content of memory unit @{text "n"} to value @{text "v"}. + @{text "am"} is the Abacus memory before setting. + If address @{text "n"} is outside to scope of @{text "am"}, @{text "am"} + is extended so that @{text "n"} becomes in scope. +*} +fun abc_lm_s :: "abc_lm \ nat \ nat \ abc_lm" + where + "abc_lm_s am n v = (if (n < length am) then (am[n:=v]) else + am@ (replicate (n - length am) 0) @ [v])" + + +text {* + The configuration of Abaucs machines consists of its current state and its + current memory: +*} +type_synonym abc_conf_l = "abc_state \ abc_lm" + +text {* + Fetch instruction out of Abacus program: +*} + +fun abc_fetch :: "nat \ abc_prog \ abc_inst option" + where + "abc_fetch s p = (if (s < length p) then Some (p ! s) + else None)" + +text {* + Single step execution of Abacus machine. If no instruction is feteched, + configuration does not change. +*} +fun abc_step_l :: "abc_conf_l \ abc_inst option \ abc_conf_l" + where + "abc_step_l (s, lm) a = (case a of + None \ (s, lm) | + Some (Inc n) \ (let nv = abc_lm_v lm n in + (s + 1, abc_lm_s lm n (nv + 1))) | + Some (Dec n e) \ (let nv = abc_lm_v lm n in + if (nv = 0) then (e, abc_lm_s lm n 0) + else (s + 1, abc_lm_s lm n (nv - 1))) | + Some (Goto n) \ (n, lm) + )" + +text {* + Multi-step execution of Abacus machine. +*} +fun abc_steps_l :: "abc_conf_l \ abc_prog \ nat \ abc_conf_l" + where + "abc_steps_l (s, lm) p 0 = (s, lm)" | + "abc_steps_l (s, lm) p (Suc n) = abc_steps_l (abc_step_l (s, lm) (abc_fetch s p)) p n" + +section {* + Compiling Abacus machines into Truing machines +*} + + +subsection {* + Compiling functions +*} + +text {* + @{text "findnth n"} returns the TM which locates the represention of + memory cell @{text "n"} on the tape and changes representation of zero + on the way. +*} + +fun findnth :: "nat \ tprog" + where + "findnth 0 = []" | + "findnth (Suc n) = (findnth n @ [(W1, 2 * n + 1), + (R, 2 * n + 2), (R, 2 * n + 3), (R, 2 * n + 2)])" + +text {* + @{text "tinc_b"} returns the TM which increments the representation + of the memory cell under rw-head by one and move the representation + of cells afterwards to the right accordingly. + *} + +definition tinc_b :: "tprog" + where + "tinc_b \ [(W1, 1), (R, 2), (W1, 3), (R, 2), (W1, 3), (R, 4), + (L, 7), (W0, 5), (R, 6), (W0, 5), (W1, 3), (R, 6), + (L, 8), (L, 7), (R, 9), (L, 7), (R, 10), (W0, 9)]" + +text {* + @{text "tshift tm off"} shifts @{text "tm"} by offset @{text "off"}, leaving + instructions concerning state @{text "0"} unchanged, because state @{text "0"} + is the end state, which needs not be changed with shift operation. + *} + +fun tshift :: "tprog \ nat \ tprog" + where + "tshift tp off = (map (\ (action, state). + (action, (if state = 0 then 0 + else state + off))) tp)" + +text {* + @{text "tinc ss n"} returns the TM which simulates the execution of + Abacus instruction @{text "Inc n"}, assuming that TM is located at + location @{text "ss"} in the final TM complied from the whole + Abacus program. +*} + +fun tinc :: "nat \ nat \ tprog" + where + "tinc ss n = tshift (findnth n @ tshift tinc_b (2 * n)) (ss - 1)" + +text {* + @{text "tinc_b"} returns the TM which decrements the representation + of the memory cell under rw-head by one and move the representation + of cells afterwards to the left accordingly. + *} + +definition tdec_b :: "tprog" + where + "tdec_b \ [(W1, 1), (R, 2), (L, 14), (R, 3), (L, 4), (R, 3), + (R, 5), (W0, 4), (R, 6), (W0, 5), (L, 7), (L, 8), + (L, 11), (W0, 7), (W1, 8), (R, 9), (L, 10), (R, 9), + (R, 5), (W0, 10), (L, 12), (L, 11), (R, 13), (L, 11), + (R, 17), (W0, 13), (L, 15), (L, 14), (R, 16), (L, 14), + (R, 0), (W0, 16)]" + +text {* + @{text "sete tp e"} attaches the termination edges (edges leading to state @{text "0"}) + of TM @{text "tp"} to the intruction labelled by @{text "e"}. + *} + +fun sete :: "tprog \ nat \ tprog" + where + "sete tp e = map (\ (action, state). (action, if state = 0 then e else state)) tp" + +text {* + @{text "tdec ss n label"} returns the TM which simulates the execution of + Abacus instruction @{text "Dec n label"}, assuming that TM is located at + location @{text "ss"} in the final TM complied from the whole + Abacus program. +*} + +fun tdec :: "nat \ nat \ nat \ tprog" + where + "tdec ss n e = sete (tshift (findnth n @ tshift tdec_b (2 * n)) + (ss - 1)) e" + +text {* + @{text "tgoto f(label)"} returns the TM simulating the execution of Abacus instruction + @{text "Goto label"}, where @{text "f(label)"} is the corresponding location of + @{text "label"} in the final TM compiled from the overall Abacus program. +*} + +fun tgoto :: "nat \ tprog" + where + "tgoto n = [(Nop, n), (Nop, n)]" + +text {* + The layout of the final TM compiled from an Abacus program is represented + as a list of natural numbers, where the list element at index @{text "n"} represents the + starting state of the TM simulating the execution of @{text "n"}-th instruction + in the Abacus program. +*} + +type_synonym layout = "nat list" + +text {* + @{text "length_of i"} is the length of the + TM simulating the Abacus instruction @{text "i"}. +*} +fun length_of :: "abc_inst \ nat" + where + "length_of i = (case i of + Inc n \ 2 * n + 9 | + Dec n e \ 2 * n + 16 | + Goto n \ 1)" + +text {* + @{text "layout_of ap"} returns the layout of Abacus program @{text "ap"}. +*} +fun layout_of :: "abc_prog \ layout" + where "layout_of ap = map length_of ap" + + +text {* + @{text "start_of layout n"} looks out the starting state of @{text "n"}-th + TM in the finall TM. +*} + +fun start_of :: "nat list \ nat \ nat" + where + "start_of ly 0 = Suc 0" | + "start_of ly (Suc as) = + (if as < length ly then start_of ly as + (ly ! as) + else start_of ly as)" + +text {* + @{text "ci lo ss i"} complies Abacus instruction @{text "i"} + assuming the TM of @{text "i"} starts from state @{text "ss"} + within the overal layout @{text "lo"}. +*} + +fun ci :: "layout \ nat \ abc_inst \ tprog" + where + "ci ly ss i = (case i of + Inc n \ tinc ss n | + Dec n e \ tdec ss n (start_of ly e) | + Goto n \ tgoto (start_of ly n))" + +text {* + @{text "tpairs_of ap"} transfroms Abacus program @{text "ap"} pairing + every instruction with its starting state. +*} +fun tpairs_of :: "abc_prog \ (nat \ abc_inst) list" + where "tpairs_of ap = (zip (map (start_of (layout_of ap)) + [0..<(length ap)]) ap)" + + +text {* + @{text "tms_of ap"} returns the list of TMs, where every one of them simulates + the corresponding Abacus intruction in @{text "ap"}. +*} + +fun tms_of :: "abc_prog \ tprog list" + where "tms_of ap = map (\ (n, tm). ci (layout_of ap) n tm) + (tpairs_of ap)" + +text {* + @{text "tm_of ap"} returns the final TM machine compiled from Abacus program @{text "ap"}. +*} +fun tm_of :: "abc_prog \ tprog" + where "tm_of ap = concat (tms_of ap)" + +text {* + The following two functions specify the well-formedness of complied TM. +*} +fun t_ncorrect :: "tprog \ bool" + where + "t_ncorrect tp = (length tp mod 2 = 0)" + +fun abc2t_correct :: "abc_prog \ bool" + where + "abc2t_correct ap = list_all (\ (n, tm). + t_ncorrect (ci (layout_of ap) n tm)) (tpairs_of ap)" + +lemma findnth_length: "length (findnth n) div 2 = 2 * n" +apply(induct n, simp, simp) +done + +lemma ci_length : "length (ci ns n ai) div 2 = length_of ai" +apply(auto simp: ci.simps tinc_b_def tdec_b_def findnth_length + split: abc_inst.splits) +done + +subsection {* + Representation of Abacus memory by TM tape +*} + +consts tape_of :: "'a \ block list" ("<_>" 100) + +text {* + @{text "tape_of_nat_list am"} returns the TM tape representing + Abacus memory @{text "am"}. + *} + +fun tape_of_nat_list :: "nat list \ block list" + where + "tape_of_nat_list [] = []" | + "tape_of_nat_list [n] = Oc\<^bsup>n+1\<^esup>" | + "tape_of_nat_list (n#ns) = (Oc\<^bsup>n+1\<^esup>) @ [Bk] @ (tape_of_nat_list ns)" + +defs (overloaded) + tape_of_nl_abv: " \ tape_of_nat_list am" + tape_of_nat_abv : "<(n::nat)> \ Oc\<^bsup>n+1\<^esup>" + +text {* + @{text "crsp_l acf tcf"} meams the abacus configuration @{text "acf"} + is corretly represented by the TM configuration @{text "tcf"}. +*} + +fun crsp_l :: "layout \ abc_conf_l \ t_conf \ block list \ bool" + where + "crsp_l ly (as, lm) (ts, (l, r)) inres = + (ts = start_of ly as \ (\ rn. r = @ Bk\<^bsup>rn\<^esup>) + \ l = Bk # Bk # inres)" + +declare crsp_l.simps[simp del] + +subsection {* + A more general definition of TM execution. +*} + +(* +fun nnth_of :: "(taction \ nat) list \ nat \ nat \ (taction \ nat)" + where + "nnth_of p s b = (if 2*s < length p + then (p ! (2*s + b)) else (Nop, 0))" + +thm nth_of.simps + +fun nfetch :: "tprog \ nat \ block \ taction \ nat" + where + "nfetch p 0 b = (Nop, 0)" | + "nfetch p (Suc s) b = + (case b of + Bk \ nnth_of p s 0 | + Oc \ nnth_of p s 1)" +*) + +text {* + @{text "t_step tcf (tp, ss)"} returns the result of one step exection of TM @{text "tp"} + assuming @{text "tp"} starts from instial state @{text "ss"}. +*} + +fun t_step :: "t_conf \ (tprog \ nat) \ t_conf" + where + "t_step c (p, off) = + (let (state, leftn, rightn) = c in + let (action, next_state) = fetch p (state-off) + (case rightn of + [] \ Bk | + Bk # xs \ Bk | + Oc # xs \ Oc + ) + in + (next_state, new_tape action (leftn, rightn)))" + + +text {* + @{text "t_steps tcf (tp, ss) n"} returns the result of @{text "n"}-step exection + of TM @{text "tp"} assuming @{text "tp"} starts from instial state @{text "ss"}. +*} + +fun t_steps :: "t_conf \ (tprog \ nat) \ nat \ t_conf" + where + "t_steps c (p, off) 0 = c" | + "t_steps c (p, off) (Suc n) = t_steps + (t_step c (p, off)) (p, off) n" + +lemma stepn: "t_steps c (p, off) (Suc n) = + t_step (t_steps c (p, off) n) (p, off)" +apply(induct n arbitrary: c, simp add: t_steps.simps) +apply(simp add: t_steps.simps) +done + +text {* + The type of invarints expressing correspondence between + Abacus configuration and TM configuration. +*} + +type_synonym inc_inv_t = "abc_conf_l \ t_conf \ block list \ bool" + +declare tms_of.simps[simp del] tm_of.simps[simp del] + layout_of.simps[simp del] abc_fetch.simps [simp del] + t_step.simps[simp del] t_steps.simps[simp del] + tpairs_of.simps[simp del] start_of.simps[simp del] + fetch.simps [simp del] t_ncorrect.simps[simp del] + new_tape.simps [simp del] ci.simps [simp del] length_of.simps[simp del] + layout_of.simps[simp del] crsp_l.simps[simp del] + abc2t_correct.simps[simp del] + +lemma tct_div2: "t_ncorrect tp \ (length tp) mod 2 = 0" +apply(simp add: t_ncorrect.simps) +done + +lemma t_shift_fetch: + "\t_ncorrect tp1; t_ncorrect tp; + length tp1 div 2 < a \ a \ length tp1 div 2 + length tp div 2\ + \ fetch tp (a - length tp1 div 2) b = + fetch (tp1 @ tp @ tp2) a b" +apply(subgoal_tac "\ x. a = length tp1 div 2 + x", erule exE, simp) +apply(case_tac x, simp) +apply(subgoal_tac "length tp1 div 2 + Suc nat = + Suc (length tp1 div 2 + nat)") +apply(simp only: fetch.simps nth_of.simps, auto) +apply(case_tac b, simp) +apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp) +apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, simp) +apply(simp add: t_ncorrect.simps, auto) +apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp) +apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, auto) +apply(simp add: t_ncorrect.simps, auto) +apply(rule_tac x = "a - length tp1 div 2" in exI, simp) +done + +lemma t_shift_in_step: + "\t_step (a, aa, ba) (tp, length tp1 div 2) = (s, l, r); + t_ncorrect tp1; t_ncorrect tp; + length tp1 div 2 < a \ a \ length tp1 div 2 + length tp div 2\ + \ t_step (a, aa, ba) (tp1 @ tp @ tp2, 0) = (s, l, r)" +apply(simp add: t_step.simps) +apply(subgoal_tac "fetch tp (a - length tp1 div 2) (case ba of [] \ + Bk | x # xs \ x) + = fetch (tp1 @ tp @ tp2) a (case ba of [] \ Bk | x # xs + \ x)") +apply(case_tac "fetch tp (a - length tp1 div 2) (case ba of [] \ Bk + | x # xs \ x)") +apply(auto intro: t_shift_fetch) +apply(case_tac ba, simp, simp) +apply(case_tac aaa, simp, simp) +done + +declare add_Suc_right[simp del] +lemma t_step_add: "t_steps c (p, off) (m + n) = + t_steps (t_steps c (p, off) m) (p, off) n" +apply(induct m arbitrary: n, simp add: t_steps.simps, simp) +apply(subgoal_tac "t_steps c (p, off) (Suc (m + n)) = + t_steps c (p, off) (m + Suc n)", simp) +apply(subgoal_tac "t_steps (t_steps c (p, off) m) (p, off) (Suc n) = + t_steps (t_step (t_steps c (p, off) m) (p, off)) + (p, off) n") +apply(simp, simp add: stepn) +apply(simp only: t_steps.simps) +apply(simp only: add_Suc_right) +done +declare add_Suc_right[simp] + +lemma s_out_fetch: "\t_ncorrect tp; + \ (length tp1 div 2 < a \ a \ length tp1 div 2 + + length tp div 2)\ + \ fetch tp (a - length tp1 div 2) b = (Nop, 0)" +apply(auto) +apply(simp add: fetch.simps) +apply(subgoal_tac "\ x. a - length tp1 div 2 = length tp div 2 + x") +apply(erule exE, simp) +apply(case_tac x, simp) +apply(auto simp add: fetch.simps) +apply(subgoal_tac "2 * (length tp div 2) = length tp") +apply(auto simp: t_ncorrect.simps split: block.splits) +apply(rule_tac x = "a - length tp1 div 2 - length tp div 2" in exI + , simp) +done + +lemma conf_keep_step: + "\t_ncorrect tp; + \ (length tp1 div 2 < a \ a \ length tp1 div 2 + + length tp div 2)\ + \ t_step (a, aa, ba) (tp, length tp1 div 2) = (0, aa, ba)" +apply(simp add: t_step.simps) +apply(subgoal_tac "fetch tp (a - length tp1 div 2) (case ba of [] \ + Bk | Bk # xs \ Bk | Oc # xs \ Oc) = (Nop, 0)") +apply(simp add: new_tape.simps) +apply(rule s_out_fetch, simp, simp) +done + +lemma conf_keep: + "\t_ncorrect tp; + \ (length tp1 div 2 < a \ + a \ length tp1 div 2 + length tp div 2); n > 0\ + \ t_steps (a, aa, ba) (tp, length tp1 div 2) n = (0, aa, ba)" +apply(induct n, simp) +apply(case_tac n, simp add: t_steps.simps) +apply(rule_tac conf_keep_step, simp+) +apply(subgoal_tac " t_steps (a, aa, ba) + (tp, length tp1 div 2) (Suc (Suc nat)) + = t_step (t_steps (a, aa, ba) + (tp, length tp1 div 2) (Suc nat)) (tp, length tp1 div 2)") +apply(simp) +apply(rule_tac conf_keep_step, simp, simp) +apply(rule stepn) +done + +lemma state_bef_inside: + "\t_ncorrect tp1; t_ncorrect tp; + t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r); + length tp1 div 2 < s0 \ + s0 \ length tp1 div 2 + length tp div 2; + length tp1 div 2 < s \ s \ length tp1 div 2 + length tp div 2; + n < stp; t_steps (s0, l0, r0) (tp, length tp1 div 2) n = + (a, aa, ba)\ + \ length tp1 div 2 < a \ + a \ length tp1 div 2 + length tp div 2" +apply(subgoal_tac "\ x. stp = n + x", erule exE) +apply(simp only: t_step_add) +apply(rule classical) +apply(subgoal_tac "t_steps (a, aa, ba) + (tp, length tp1 div 2) x = (0, aa, ba)") +apply(simp) +apply(rule conf_keep, simp, simp, simp) +apply(rule_tac x = "stp - n" in exI, simp) +done + +lemma turing_shift_inside: + "\t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r); + length tp1 div 2 < s0 \ + s0 \ length tp1 div 2 + length tp div 2; + t_ncorrect tp1; t_ncorrect tp; + length tp1 div 2 < s \ + s \ length tp1 div 2 + length tp div 2\ + \ t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp = (s, l, r)" +apply(induct stp arbitrary: s l r) +apply(simp add: t_steps.simps) +apply(subgoal_tac " t_steps (s0, l0, r0) + (tp, length tp1 div 2) (Suc stp) + = t_step (t_steps (s0, l0, r0) + (tp, length tp1 div 2) stp) (tp, length tp1 div 2)") +apply(case_tac "t_steps (s0, l0, r0) (tp, length tp1 div 2) stp") +apply(subgoal_tac "length tp1 div 2 < a \ + a \ length tp1 div 2 + length tp div 2") +apply(subgoal_tac "t_steps (s0, l0, r0) + (tp1 @ tp @ tp2, 0) stp = (a, b, c)") +apply(simp only: stepn, simp) +apply(rule_tac t_shift_in_step, simp+) +defer +apply(rule stepn) +apply(rule_tac n = stp and stp = "Suc stp" and a = a + and aa = b and ba = c in state_bef_inside, simp+) +done + +lemma take_Suc_last[elim]: "Suc as \ length xs \ + take (Suc as) xs = take as xs @ [xs ! as]" +apply(induct xs arbitrary: as, simp, simp) +apply(case_tac as, simp, simp) +done + +lemma concat_suc: "Suc as \ length xs \ + concat (take (Suc as) xs) = concat (take as xs) @ xs! as" +apply(subgoal_tac "take (Suc as) xs = take as xs @ [xs ! as]", simp) +by auto + +lemma concat_take_suc_iff: "Suc n \ length tps \ + concat (take n tps) @ (tps ! n) = concat (take (Suc n) tps)" +apply(drule_tac concat_suc, simp) +done + +lemma concat_drop_suc_iff: + "Suc n < length tps \ concat (drop (Suc n) tps) = + tps ! Suc n @ concat (drop (Suc (Suc n)) tps)" +apply(induct tps arbitrary: n, simp, simp) +apply(case_tac tps, simp, simp) +apply(case_tac n, simp, simp) +done + +declare append_assoc[simp del] + +lemma tm_append: "\n < length tps; tp = tps ! n\ \ + \ tp1 tp2. concat tps = tp1 @ tp @ tp2 \ tp1 = + concat (take n tps) \ tp2 = concat (drop (Suc n) tps)" +apply(rule_tac x = "concat (take n tps)" in exI) +apply(rule_tac x = "concat (drop (Suc n) tps)" in exI) +apply(auto) +apply(induct n, simp) +apply(case_tac tps, simp, simp, simp) +apply(subgoal_tac "concat (take n tps) @ (tps ! n) = + concat (take (Suc n) tps)") +apply(simp only: append_assoc[THEN sym], simp only: append_assoc) +apply(subgoal_tac " concat (drop (Suc n) tps) = tps ! Suc n @ + concat (drop (Suc (Suc n)) tps)", simp) +apply(rule_tac concat_drop_suc_iff, simp) +apply(rule_tac concat_take_suc_iff, simp) +done + +declare append_assoc[simp] + +lemma map_of: "n < length xs \ (map f xs) ! n = f (xs ! n)" +by(auto) + +lemma [simp]: "length (tms_of aprog) = length aprog" +apply(auto simp: tms_of.simps tpairs_of.simps) +done + +lemma ci_nth: "\ly = layout_of aprog; as < length aprog; + abc_fetch as aprog = Some ins\ + \ ci ly (start_of ly as) ins = tms_of aprog ! as" +apply(simp add: tms_of.simps tpairs_of.simps + abc_fetch.simps map_of del: map_append) +done + +lemma t_split:"\ + ly = layout_of aprog; + as < length aprog; abc_fetch as aprog = Some ins\ + \ \ tp1 tp2. concat (tms_of aprog) = + tp1 @ (ci ly (start_of ly as) ins) @ tp2 + \ tp1 = concat (take as (tms_of aprog)) \ + tp2 = concat (drop (Suc as) (tms_of aprog))" +apply(insert tm_append[of "as" "tms_of aprog" + "ci ly (start_of ly as) ins"], simp) +apply(subgoal_tac "ci ly (start_of ly as) ins = (tms_of aprog) ! as") +apply(subgoal_tac "length (tms_of aprog) = length aprog", simp, simp) +apply(rule_tac ci_nth, auto) +done + +lemma math_sub: "\x >= Suc 0; x - 1 = z\ \ x + y - Suc 0 = z + y" +by auto + +lemma start_more_one: "as \ 0 \ start_of ly as >= Suc 0" +apply(induct as, simp add: start_of.simps) +apply(case_tac as, auto simp: start_of.simps) +done + +lemma tm_ct: "\abc2t_correct aprog; tp \ set (tms_of aprog)\ \ + t_ncorrect tp" +apply(simp add: abc2t_correct.simps tms_of.simps) +apply(auto) +apply(simp add:list_all_iff, auto) +done + +lemma div_apart: "\x mod (2::nat) = 0; y mod 2 = 0\ + \ (x + y) div 2 = x div 2 + y div 2" +apply(drule mod_eqD)+ +apply(auto) +done + +lemma div_apart_iff: "\x mod (2::nat) = 0; y mod 2 = 0\ \ + (x + y) mod 2 = 0" +apply(auto) +done + +lemma tms_ct: "\abc2t_correct aprog; n < length aprog\ \ + t_ncorrect (concat (take n (tms_of aprog)))" +apply(induct n, simp add: t_ncorrect.simps, simp) +apply(subgoal_tac "concat (take (Suc n) (tms_of aprog)) = + concat (take n (tms_of aprog)) @ (tms_of aprog ! n)", simp) +apply(simp add: t_ncorrect.simps) +apply(rule_tac div_apart_iff, simp) +apply(subgoal_tac "t_ncorrect (tms_of aprog ! n)", + simp add: t_ncorrect.simps) +apply(rule_tac tm_ct, simp) +apply(rule_tac nth_mem, simp add: tms_of.simps tpairs_of.simps) +apply(rule_tac concat_suc, simp add: tms_of.simps tpairs_of.simps) +done + +lemma tcorrect_div2: "\abc2t_correct aprog; Suc as < length aprog\ + \ (length (concat (take as (tms_of aprog))) + length (tms_of aprog + ! as)) div 2 = length (concat (take as (tms_of aprog))) div 2 + + length (tms_of aprog ! as) div 2" +apply(subgoal_tac "t_ncorrect (tms_of aprog ! as)") +apply(subgoal_tac "t_ncorrect (concat (take as (tms_of aprog)))") +apply(rule_tac div_apart) +apply(rule tct_div2, simp)+ +apply(erule_tac tms_ct, simp) +apply(rule_tac tm_ct, simp) +apply(rule_tac nth_mem) +apply(simp add: tms_of.simps tpairs_of.simps) +done + +lemma [simp]: "length (layout_of aprog) = length aprog" +apply(auto simp: layout_of.simps) +done + +lemma start_of_ind: "\as < length aprog; ly = layout_of aprog\ \ + start_of ly (Suc as) = start_of ly as + + length ((tms_of aprog) ! as) div 2" +apply(simp only: start_of.simps, simp) +apply(auto simp: start_of.simps tms_of.simps layout_of.simps + tpairs_of.simps) +apply(simp add: ci_length) +done + +lemma concat_take_suc: "Suc n \ length xs \ + concat (take (Suc n) xs) = concat (take n xs) @ (xs ! n)" +apply(subgoal_tac "take (Suc n) xs = + take n xs @ [xs ! n]") +apply(auto) +done + +lemma ci_length_not0: "Suc 0 <= length (ci ly as i) div 2" +apply(subgoal_tac "length (ci ly as i) div 2 = length_of i") +apply(simp add: length_of.simps split: abc_inst.splits) +apply(rule ci_length) +done + +lemma findnth_length2: "length (findnth n) = 4 * n" +apply(induct n, simp) +apply(simp) +done + +lemma ci_length2: "length (ci ly as i) = 2 * (length_of i)" +apply(simp add: ci.simps length_of.simps tinc_b_def tdec_b_def + split: abc_inst.splits, auto) +apply(simp add: findnth_length2)+ +done + +lemma tm_mod2: "as < length aprog \ + length (tms_of aprog ! as) mod 2 = 0" +apply(simp add: tms_of.simps) +apply(subgoal_tac "map (\(x, y). ci (layout_of aprog) x y) + (tpairs_of aprog) ! as + = (\(x, y). ci (layout_of aprog) x y) + ((tpairs_of aprog) ! as)", simp) +apply(case_tac "(tpairs_of aprog ! as)", simp) +apply(subgoal_tac "length (ci (layout_of aprog) a b) = + 2 * (length_of b)", simp) +apply(rule ci_length2) +apply(rule map_of, simp add: tms_of.simps tpairs_of.simps) +done + +lemma tms_mod2: "as \ length aprog \ + length (concat (take as (tms_of aprog))) mod 2 = 0" +apply(induct as, simp, simp) +apply(subgoal_tac "concat (take (Suc as) (tms_of aprog)) + = concat (take as (tms_of aprog)) @ + (tms_of aprog ! as)", auto) +apply(rule div_apart_iff, simp, rule tm_mod2, simp) +apply(rule concat_take_suc, simp add: tms_of.simps tpairs_of.simps) +done + +lemma [simp]: "\as < length aprog; (abc_fetch as aprog) = Some ins\ + \ ci (layout_of aprog) + (start_of (layout_of aprog) as) (ins) \ set (tms_of aprog)" +apply(insert ci_nth[of "layout_of aprog" aprog as], simp) +done + +lemma startof_not0: "start_of ly as > 0" +apply(induct as, simp add: start_of.simps) +apply(case_tac as, auto simp: start_of.simps) +done + +declare abc_step_l.simps[simp del] +lemma pre_lheq: "\tp = concat (take as (tms_of aprog)); + abc2t_correct aprog; as \ length aprog\ \ + start_of (layout_of aprog) as - Suc 0 = length tp div 2" +apply(induct as arbitrary: tp, simp add: start_of.simps, simp) +proof - + fix as tp + assume h1: "\tp. tp = concat (take as (tms_of aprog)) \ + start_of (layout_of aprog) as - Suc 0 = + length (concat (take as (tms_of aprog))) div 2" + and h2: " abc2t_correct aprog" "Suc as \ length aprog" + from h2 show "start_of (layout_of aprog) (Suc as) - Suc 0 = + length (concat (take (Suc as) (tms_of aprog))) div 2" + apply(insert h1[of "concat (take as (tms_of aprog))"], simp) + apply(insert start_of_ind[of as aprog "layout_of aprog"], simp) + apply(subgoal_tac "(take (Suc as) (tms_of aprog)) = + take as (tms_of aprog) @ [(tms_of aprog) ! as]", simp) + apply(subgoal_tac "(length (concat (take as (tms_of aprog))) + + length (tms_of aprog ! as)) div 2 + = length (concat (take as (tms_of aprog))) div 2 + + length (tms_of aprog ! as) div 2", simp) + apply(subgoal_tac "start_of (layout_of aprog) as = + length (concat (take as (tms_of aprog))) div 2 + Suc 0", simp) + apply(subgoal_tac "start_of (layout_of aprog) as > 0", simp, + rule_tac startof_not0) + apply(insert tm_mod2[of as aprog], simp) + apply(insert tms_mod2[of as aprog], simp, arith) + apply(rule take_Suc_last, simp) + done +qed + +lemma crsp2stateq: + "\as < length aprog; abc2t_correct aprog; + crsp_l (layout_of aprog) (as, am) (a, aa, ba) inres\ \ + a = length (concat (take as (tms_of aprog))) div 2 + 1" +apply(simp add: crsp_l.simps) +apply(insert pre_lheq[of "(concat (take as (tms_of aprog)))" as aprog] +, simp) +apply(subgoal_tac "start_of (layout_of aprog) as > 0", + auto intro: startof_not0) +done + +lemma turing_shift_outside: + "\t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r); + s \ 0; stp > 0; + length tp1 div 2 < s0 \ + s0 \ length tp1 div 2 + length tp div 2; + t_ncorrect tp1; t_ncorrect tp; + \ (length tp1 div 2 < s \ + s \ length tp1 div 2 + length tp div 2)\ + \ \stp' > 0. t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp' + = (s, l, r)" +apply(rule_tac x = stp in exI) +apply(case_tac stp, simp add: t_steps.simps) +apply(simp only: stepn) +apply(case_tac "t_steps (s0, l0, r0) (tp, length tp1 div 2) nat") +apply(subgoal_tac "length tp1 div 2 < a \ + a \ length tp1 div 2 + length tp div 2") +apply(subgoal_tac "t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) nat + = (a, b, c)", simp) +apply(rule_tac t_shift_in_step, simp+) +apply(rule_tac turing_shift_inside, simp+) +apply(rule classical) +apply(subgoal_tac "t_step (a,b,c) + (tp, length tp1 div 2) = (0, b, c)", simp) +apply(rule_tac conf_keep_step, simp+) +done + +lemma turing_shift: + "\t_steps (s0, (l0, r0)) (tp, (length tp1 div 2)) stp + = (s, (l, r)); s \ 0; stp > 0; + (length tp1 div 2 < s0 \ s0 <= length tp1 div 2 + length tp div 2); + t_ncorrect tp1; t_ncorrect tp\ \ + \ stp' > 0. t_steps (s0, (l0, r0)) (tp1 @ tp @ tp2, 0) stp' = + (s, (l, r))" +apply(case_tac "s > length tp1 div 2 \ + s <= length tp1 div 2 + length tp div 2") +apply(subgoal_tac " t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp = + (s, l, r)") +apply(rule_tac x = stp in exI, simp) +apply(rule_tac turing_shift_inside, simp+) +apply(rule_tac turing_shift_outside, simp+) +done + +lemma inc_startof_not0: "start_of ly as \ Suc 0" +apply(induct as, simp add: start_of.simps) +apply(simp add: start_of.simps) +done + +lemma s_crsp: + "\as < length aprog; abc_fetch as aprog = Some ins; + abc2t_correct aprog; + crsp_l (layout_of aprog) (as, am) (a, aa, ba) inres\ \ + length (concat (take as (tms_of aprog))) div 2 < a + \ a \ length (concat (take as (tms_of aprog))) div 2 + + length (ci (layout_of aprog) (start_of (layout_of aprog) as) + ins) div 2" +apply(subgoal_tac "a = length (concat (take as (tms_of aprog))) div + 2 + 1", simp) +apply(rule_tac ci_length_not0) +apply(rule crsp2stateq, simp+) +done + +lemma tms_out_ex: + "\ly = layout_of aprog; tprog = tm_of aprog; + abc2t_correct aprog; + crsp_l ly (as, am) tc inres; as < length aprog; + abc_fetch as aprog = Some ins; + t_steps tc (ci ly (start_of ly as) ins, + (start_of ly as) - 1) n = (s, l, r); + n > 0; + abc_step_l (as, am) (abc_fetch as aprog) = (as', am'); + s = start_of ly as' + \ + \ \ stp > 0. (t_steps tc (tprog, 0) stp = (s, (l, r)))" +apply(simp only: tm_of.simps) +apply(subgoal_tac "\ tp1 tp2. concat (tms_of aprog) = + tp1 @ (ci ly (start_of ly as) ins) @ tp2 + \ tp1 = concat (take as (tms_of aprog)) \ + tp2 = concat (drop (Suc as) (tms_of aprog))") +apply(erule exE, erule exE, erule conjE, erule conjE, + case_tac tc, simp) +apply(rule turing_shift) +apply(subgoal_tac "start_of (layout_of aprog) as - Suc 0 + = length tp1 div 2", simp) +apply(rule_tac pre_lheq, simp, simp, simp) +apply(simp add: startof_not0, simp) +apply(rule_tac s_crsp, simp, simp, simp, simp) +apply(rule tms_ct, simp, simp) +apply(rule tm_ct, simp) +apply(subgoal_tac "ci (layout_of aprog) + (start_of (layout_of aprog) as) ins + = (tms_of aprog ! as)", simp) +apply(simp add: tms_of.simps tpairs_of.simps) +apply(simp add: tms_of.simps tpairs_of.simps abc_fetch.simps) +apply(erule_tac t_split, auto simp: tm_of.simps) +done + +subsubsection {* The compilation of @{text "Inc n"} *} + +text {* + The lemmas in this section lead to the correctness of + the compilation of @{text "Inc n"} instruction. +*} + +(*****Begin: inc crsp*******) +fun at_begin_fst_bwtn :: "inc_inv_t" + where + "at_begin_fst_bwtn (as, lm) (s, l, r) ires = + (\ lm1 tn rn. lm1 = (lm @ (0\<^bsup>tn\<^esup>)) \ length lm1 = s \ + (if lm1 = [] then l = Bk # Bk # ires + else l = [Bk]@@Bk#Bk#ires) \ r = (Bk\<^bsup>rn\<^esup>))" + + +fun at_begin_fst_awtn :: "inc_inv_t" + where + "at_begin_fst_awtn (as, lm) (s, l, r) ires = + (\ lm1 tn rn. lm1 = (lm @ (0\<^bsup>tn\<^esup>)) \ length lm1 = s \ + (if lm1 = [] then l = Bk # Bk # ires + else l = [Bk]@@Bk#Bk#ires) \ r = [Oc]@Bk\<^bsup>rn\<^esup> + )" + +fun at_begin_norm :: "inc_inv_t" + where + "at_begin_norm (as, lm) (s, l, r) ires= + (\ lm1 lm2 rn. lm = lm1 @ lm2 \ length lm1 = s \ + (if lm1 = [] then l = Bk # Bk # ires + else l = Bk # @ Bk# Bk # ires ) \ r = @ (Bk\<^bsup>rn\<^esup>))" + +fun in_middle :: "inc_inv_t" + where + "in_middle (as, lm) (s, l, r) ires = + (\ lm1 lm2 tn m ml mr rn. lm @ 0\<^bsup>tn\<^esup> = lm1 @ [m] @ lm2 + \ length lm1 = s \ m + 1 = ml + mr \ + ml \ 0 \ tn = s + 1 - length lm \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml\<^esup>)@[Bk]@@ + Bk # Bk # ires) \ (r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ @ (Bk\<^bsup>rn\<^esup>) \ + (lm2 = [] \ r = (Oc\<^bsup>mr\<^esup>))) + )" + +fun inv_locate_a :: "inc_inv_t" + where "inv_locate_a (as, lm) (s, l, r) ires = + (at_begin_norm (as, lm) (s, l, r) ires \ + at_begin_fst_bwtn (as, lm) (s, l, r) ires \ + at_begin_fst_awtn (as, lm) (s, l, r) ires + )" + +fun inv_locate_b :: "inc_inv_t" + where "inv_locate_b (as, lm) (s, l, r) ires = + (in_middle (as, lm) (s, l, r)) ires " + +fun inv_after_write :: "inc_inv_t" + where "inv_after_write (as, lm) (s, l, r) ires = + (\ rn m lm1 lm2. lm = lm1 @ m # lm2 \ + (if lm1 = [] then l = Oc\<^bsup>m\<^esup> @ Bk # Bk # ires + else Oc # l = Oc\<^bsup>Suc m \<^esup>@ Bk # @ + Bk # Bk # ires) \ r = [Oc] @ @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_after_move :: "inc_inv_t" + where "inv_after_move (as, lm) (s, l, r) ires = + (\ rn m lm1 lm2. lm = lm1 @ m # lm2 \ + (if lm1 = [] then l = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires + else l = Oc\<^bsup>Suc m\<^esup>@ Bk # @ Bk # Bk # ires) \ + r = @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_after_clear :: "inc_inv_t" + where "inv_after_clear (as, lm) (s, l, r) ires = + (\ rn m lm1 lm2 r'. lm = lm1 @ m # lm2 \ + (if lm1 = [] then l = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires + else l = Oc\<^bsup>Suc m\<^esup>@ Bk # @ Bk # Bk # ires) \ + r = Bk # r' \ Oc # r' = @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_on_right_moving :: "inc_inv_t" + where "inv_on_right_moving (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = m \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ @ Bk # Bk # ires) \ + ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ @ (Bk\<^bsup>rn\<^esup>)) \ + (r = (Oc\<^bsup>mr\<^esup>) \ lm2 = [])))" + +fun inv_on_left_moving_norm :: "inc_inv_t" + where "inv_on_left_moving_norm (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = Suc m \ mr > 0 \ (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml\<^esup>) @ Bk # @ Bk # Bk # ires) + \ (r = (Oc\<^bsup>mr\<^esup>) @ Bk # @ (Bk\<^bsup>rn\<^esup>) \ + (lm2 = [] \ r = Oc\<^bsup>mr\<^esup>)))" + +fun inv_on_left_moving_in_middle_B:: "inc_inv_t" + where "inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires = + (\ lm1 lm2 rn. lm = lm1 @ lm2 \ + (if lm1 = [] then l = Bk # ires + else l = @ Bk # Bk # ires) \ + r = Bk # @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_on_left_moving :: "inc_inv_t" + where "inv_on_left_moving (as, lm) (s, l, r) ires = + (inv_on_left_moving_norm (as, lm) (s, l, r) ires \ + inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires)" + + +fun inv_check_left_moving_on_leftmost :: "inc_inv_t" + where "inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires = + (\ rn. l = ires \ r = [Bk, Bk] @ @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_check_left_moving_in_middle :: "inc_inv_t" + where "inv_check_left_moving_in_middle (as, lm) (s, l, r) ires = + + (\ lm1 lm2 r' rn. lm = lm1 @ lm2 \ + (Oc # l = @ Bk # Bk # ires) \ r = Oc # Bk # r' \ + r' = @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_check_left_moving :: "inc_inv_t" + where "inv_check_left_moving (as, lm) (s, l, r) ires = + (inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires \ + inv_check_left_moving_in_middle (as, lm) (s, l, r) ires)" + +fun inv_after_left_moving :: "inc_inv_t" + where "inv_after_left_moving (as, lm) (s, l, r) ires= + (\ rn. l = Bk # ires \ r = Bk # @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_stop :: "inc_inv_t" + where "inv_stop (as, lm) (s, l, r) ires= + (\ rn. l = Bk # Bk # ires \ r = @ (Bk\<^bsup>rn\<^esup>))" + + +fun inc_inv :: "layout \ nat \ inc_inv_t" + where + "inc_inv ly n (as, lm) (s, l, r) ires = + (let ss = start_of ly as in + let lm' = abc_lm_s lm n ((abc_lm_v lm n)+1) in + if s = 0 then False + else if s < ss then False + else if s < ss + 2 * n then + if (s - ss) mod 2 = 0 then + inv_locate_a (as, lm) ((s - ss) div 2, l, r) ires + else inv_locate_b (as, lm) ((s - ss) div 2, l, r) ires + else if s = ss + 2 * n then + inv_locate_a (as, lm) (n, l, r) ires + else if s = ss + 2 * n + 1 then + inv_locate_b (as, lm) (n, l, r) ires + else if s = ss + 2 * n + 2 then + inv_after_write (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 3 then + inv_after_move (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 4 then + inv_after_clear (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 5 then + inv_on_right_moving (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 6 then + inv_on_left_moving (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 7 then + inv_check_left_moving (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 8 then + inv_after_left_moving (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 9 then + inv_stop (as, lm') (s - ss, l, r) ires + else False) " + +lemma fetch_intro: + "\\xs.\ba = Oc # xs\ \ P (fetch prog i Oc); + \xs.\ba = Bk # xs\ \ P (fetch prog i Bk); + ba = [] \ P (fetch prog i Bk) + \ \ P (fetch prog i + (case ba of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc))" +by (auto split:list.splits block.splits) + +lemma length_findnth[simp]: "length (findnth n) = 4 * n" +apply(induct n, simp) +apply(simp) +done + +declare tshift.simps[simp del] +declare findnth.simps[simp del] + +lemma findnth_nth: + "\n > q; x < 4\ \ + (findnth n) ! (4 * q + x) = (findnth (Suc q) ! (4 * q + x))" +apply(induct n, simp) +apply(case_tac "q < n", simp add: findnth.simps, auto) +apply(simp add: nth_append) +apply(subgoal_tac "q = n", simp) +apply(arith) +done + +lemma Suc_pre[simp]: "\ a < start_of ly as \ + (Suc a - start_of ly as) = Suc (a - start_of ly as)" +apply(arith) +done + +lemma fetch_locate_a_o: " +\a q xs. + \\ a < start_of (layout_of aprog) as; + a < start_of (layout_of aprog) as + 2 * n; + a - start_of (layout_of aprog) as = 2 * q; + start_of (layout_of aprog) as > 0\ + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (2 * q)) Oc) = (R, a+1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append Suc_pre) +apply(subgoal_tac "(findnth n ! Suc (4 * q)) = + findnth (Suc q) ! (4 * q + 1)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n !(4 * q + 1) = + findnth (Suc q) ! (4 * q + 1)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma fetch_locate_a_b: " +\a q xs. + \abc_fetch as aprog = Some (Inc n); + \ a < start_of (layout_of aprog) as; + a < start_of (layout_of aprog) as + 2 * n; + a - start_of (layout_of aprog) as = 2 * q; + start_of (layout_of aprog) as > 0\ + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * q)) Bk) + = (W1, a)" +apply(auto simp: ci.simps findnth.simps fetch.simps + tshift.simps nth_append) +apply(subgoal_tac "(findnth n ! (4 * q)) = + findnth (Suc q) ! (4 * q )") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n !(4 * q + 0) = + findnth (Suc q) ! (4 * q + 0)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma [intro]: "x mod 2 = Suc 0 \ \ q. x = Suc (2 * q)" +apply(drule mod_eqD, auto) +done + +lemma add3_Suc: "x + 3 = Suc (Suc (Suc x))" +apply(arith) +done + +declare start_of.simps[simp] +(* +lemma layout_not0: "start_of ly as > 0" +by(induct as, auto) +*) +lemma [simp]: + "\\ a < start_of (layout_of aprog) as; + a - start_of (layout_of aprog) as = Suc (2 * q); + abc_fetch as aprog = Some (Inc n); + start_of (layout_of aprog) as > 0\ + \ Suc (Suc (2 * q + start_of (layout_of aprog) as - Suc 0)) = a" +apply(subgoal_tac +"Suc (Suc (2 * q + start_of (layout_of aprog) as - Suc 0)) + = 2 + 2 * q + start_of (layout_of aprog) as - Suc 0", + simp, simp add: inc_startof_not0) +done + +lemma fetch_locate_b_o: " +\a xs. + \0 < a; \ a < start_of (layout_of aprog) as; + a < start_of (layout_of aprog) as + 2 * n; + (a - start_of (layout_of aprog) as) mod 2 = Suc 0; + start_of (layout_of aprog) as > 0\ + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (a - start_of (layout_of aprog) as)) Oc) = (R, a)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append) +apply(subgoal_tac "\ q. (a - start_of (layout_of aprog) as) = + 2 * q + 1", auto) +apply(subgoal_tac "(findnth n ! Suc (Suc (Suc (4 * q)))) + = findnth (Suc q) ! (4 * q + 3)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n ! (4 * q + 3) = + findnth (Suc q) ! (4 * q + 3)", simp add: add3_Suc) +apply(rule_tac findnth_nth, auto) +done + +lemma fetch_locate_b_b: " +\a xs. + \0 < a; \ a < start_of (layout_of aprog) as; + a < start_of (layout_of aprog) as + 2 * n; + (a - start_of (layout_of aprog) as) mod 2 = Suc 0; + start_of (layout_of aprog) as > 0\ + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (a - start_of (layout_of aprog) as)) Bk) + = (R, a + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append) +apply(subgoal_tac "\ q. (a - start_of (layout_of aprog) as) = + 2 * q + 1", auto) +apply(subgoal_tac "(findnth n ! Suc ((Suc (4 * q)))) = + findnth (Suc q) ! (4 * q + 2)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n ! (4 * q + 2) = + findnth (Suc q) ! (4 * q + 2)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma fetch_locate_n_a_o: + "start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * n)) Oc) = + (R, start_of (layout_of aprog) as + 2 * n + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_locate_n_a_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * n)) Bk) + = (W1, start_of (layout_of aprog) as + 2 * n)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_locate_n_b_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (Suc (2 * n))) Oc) = + (R, start_of (layout_of aprog) as + 2 * n + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_locate_n_b_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (Suc (2 * n))) Bk) = + (W1, start_of (layout_of aprog) as + 2 * n + 2)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_after_write_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (Suc (Suc (2 * n)))) Oc) = + (R, start_of (layout_of aprog) as + 2*n + 3)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_after_move_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (4 + 2 * n) Oc) + = (W0, start_of (layout_of aprog) as + 2 * n + 4)" +apply(auto simp: ci.simps findnth.simps tshift.simps + tinc_b_def add3_Suc) +apply(subgoal_tac "4 + 2*n = Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_after_move_b: " + start_of (layout_of aprog) as > 0 + \(fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (4 + 2 * n) Bk) + = (L, start_of (layout_of aprog) as + 2 * n + 6)" +apply(auto simp: ci.simps findnth.simps tshift.simps + tinc_b_def add3_Suc) +apply(subgoal_tac "4 + 2*n = Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_clear_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (5 + 2 * n) Bk) + = (R, start_of (layout_of aprog) as + 2 * n + 5)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tinc_b_def add3_Suc) +apply(subgoal_tac "5 + 2*n = Suc (2*n + 4)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_right_move_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (6 + 2*n) Oc) + = (R, start_of (layout_of aprog) as + 2 * n + 5)" +apply(auto simp: ci.simps findnth.simps tshift.simps + tinc_b_def add3_Suc) +apply(subgoal_tac "6 + 2*n = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_right_move_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (6 + 2*n) Bk) + = (W1, start_of (layout_of aprog) as + 2 * n + 2)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tinc_b_def add3_Suc) +apply(subgoal_tac "6 + 2*n = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_left_move_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (7 + 2*n) Oc) + = (L, start_of (layout_of aprog) as + 2 * n + 6)" +apply(auto simp: ci.simps findnth.simps tshift.simps + tinc_b_def add3_Suc) +apply(subgoal_tac "7 + 2*n = Suc (2*n + 6)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_left_move_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (7 + 2*n) Bk) + = (L, start_of (layout_of aprog) as + 2 * n + 7)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tinc_b_def add3_Suc) +apply(subgoal_tac "7 + 2*n = Suc (2*n + 6)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_check_left_move_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (8 + 2*n) Oc) + = (L, start_of (layout_of aprog) as + 2 * n + 6)" +apply(auto simp: ci.simps findnth.simps tshift.simps tinc_b_def) +apply(subgoal_tac "8 + 2 * n = Suc (2 * n + 7)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_check_left_move_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (8 + 2*n) Bk) + = (R, start_of (layout_of aprog) as + 2 * n + 8) " +apply(auto simp: ci.simps findnth.simps + tshift.simps tinc_b_def add3_Suc) +apply(subgoal_tac "8 + 2*n= Suc (2*n + 7)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_after_left_move: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (9 + 2*n) Bk) + = (R, start_of (layout_of aprog) as + 2 * n + 9)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_stop: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (10 + 2 *n) b) + = (Nop, 0)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def + split: block.splits) +done + +lemma fetch_state0: " + (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) 0 b) + = (Nop, 0)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemmas fetch_simps = + fetch_locate_a_o fetch_locate_a_b fetch_locate_b_o fetch_locate_b_b + fetch_locate_n_a_b fetch_locate_n_a_o fetch_locate_n_b_o + fetch_locate_n_b_b fetch_after_write_o fetch_after_move_o + fetch_after_move_b fetch_clear_b fetch_right_move_o + fetch_right_move_b fetch_left_move_o fetch_left_move_b + fetch_after_left_move fetch_check_left_move_o fetch_stop + fetch_state0 fetch_check_left_move_b + +text {* *} +declare exponent_def[simp del] tape_of_nat_list.simps[simp del] + at_begin_norm.simps[simp del] at_begin_fst_bwtn.simps[simp del] + at_begin_fst_awtn.simps[simp del] in_middle.simps[simp del] + abc_lm_s.simps[simp del] abc_lm_v.simps[simp del] + ci.simps[simp del] t_step.simps[simp del] + inv_after_move.simps[simp del] + inv_on_left_moving_norm.simps[simp del] + inv_on_left_moving_in_middle_B.simps[simp del] + inv_after_clear.simps[simp del] + inv_after_write.simps[simp del] inv_on_left_moving.simps[simp del] + inv_on_right_moving.simps[simp del] + inv_check_left_moving.simps[simp del] + inv_check_left_moving_in_middle.simps[simp del] + inv_check_left_moving_on_leftmost.simps[simp del] + inv_after_left_moving.simps[simp del] + inv_stop.simps[simp del] inv_locate_a.simps[simp del] + inv_locate_b.simps[simp del] +declare tms_of.simps[simp del] tm_of.simps[simp del] + layout_of.simps[simp del] abc_fetch.simps [simp del] + t_step.simps[simp del] t_steps.simps[simp del] + tpairs_of.simps[simp del] start_of.simps[simp del] + fetch.simps [simp del] new_tape.simps [simp del] + nth_of.simps [simp del] ci.simps [simp del] + length_of.simps[simp del] + +(*! Start point *) +lemma [simp]: "Suc (2 * q) mod 2 = Suc 0" +by arith + +lemma [simp]: "Suc (2 * q) div 2 = q" +by arith + +lemma [simp]: "\ \ a < start_of ly as; + a < start_of ly as + 2 * n; a - start_of ly as = 2 * q\ + \ Suc a < start_of ly as + 2 * n" +apply(arith) +done + +lemma [simp]: "x mod 2 = Suc 0 \ (Suc x) mod 2 = 0" +by arith + +lemma [simp]: "x mod 2 = Suc 0 \ (Suc x) div 2 = Suc (x div 2)" +by arith +lemma exp_def[simp]: "a\<^bsup>Suc n \<^esup>= a # a\<^bsup>n\<^esup>" +by(simp add: exponent_def) +lemma [intro]: "Bk # r = Oc\<^bsup>mr\<^esup> @ r' \ mr = 0" +by(case_tac mr, auto simp: exponent_def) + +lemma [intro]: "Bk # r = replicate mr Oc \ mr = 0" +by(case_tac mr, auto) +lemma tape_of_nl_abv_cons[simp]: "xs \ [] \ + = Oc\<^bsup>Suc x\<^esup>@ Bk # " +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac xs, simp, simp add: tape_of_nat_list.simps) +done + +lemma [simp]: "<[]::nat list> = []" +by(auto simp: tape_of_nl_abv tape_of_nat_list.simps) +lemma [simp]: "Oc # r = <(lm::nat list)> @ Bk\<^bsup>rn\<^esup>\ lm \ []" +apply(auto simp: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac rn, auto simp: exponent_def) +done +lemma BkCons_nil: "Bk # xs = @ Bk\<^bsup>rn\<^esup>\ lm = []" +apply(case_tac lm, simp) +apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps) +done +lemma BkCons_nil': "Bk # xs = @ Bk\<^bsup>ln\<^esup>\ lm = []" +by(auto intro: BkCons_nil) + +lemma hd_tl_tape_of_nat_list: + "tl (lm::nat list) \ [] \ = @ Bk # " +apply(frule tape_of_nl_abv_cons[of "tl lm" "hd lm"]) +apply(simp add: tape_of_nat_abv Bk_def del: tape_of_nl_abv_cons) +apply(subgoal_tac "lm = hd lm # tl lm", auto) +apply(case_tac lm, auto) +done +lemma [simp]: "Oc # xs = Oc\<^bsup>mr\<^esup> @ Bk # @ Bk\<^bsup>rn\<^esup>\ mr > 0" +apply(case_tac mr, auto simp: exponent_def) +done + +lemma tape_of_nat_list_cons: "xs \ [] \ tape_of_nat_list (x # xs) = + replicate (Suc x) Oc @ Bk # tape_of_nat_list xs" +apply(drule tape_of_nl_abv_cons[of xs x]) +apply(auto simp: tape_of_nl_abv tape_of_nat_abv Oc_def Bk_def exponent_def) +done + +lemma rev_eq: "rev xs = rev ys \ xs = ys" +by simp + +lemma tape_of_nat_list_eq: " xs = ys \ + tape_of_nat_list xs = tape_of_nat_list ys" +by simp + +lemma tape_of_nl_nil_eq: "<(lm::nat list)> = [] = (lm = [])" +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac lm, simp add: tape_of_nat_list.simps) +apply(case_tac "list") +apply(auto simp: tape_of_nat_list.simps) +done + +lemma rep_ind: "replicate (Suc n) a = replicate n a @ [a]" +apply(induct n, simp, simp) +done + +lemma [simp]: "Oc # r = @ replicate rn Bk \ Suc 0 \ length lm" +apply(rule_tac classical, auto) +apply(case_tac lm, simp, case_tac rn, auto) +done +lemma Oc_Bk_Cons: "Oc # Bk # list = @ Bk\<^bsup>ln\<^esup> \ + lm \ [] \ hd lm = 0" +apply(case_tac lm, simp, case_tac ln, simp add: exponent_def, simp add: exponent_def, simp) +apply(case_tac lista, auto simp: tape_of_nl_abv tape_of_nat_list.simps) +done +(*lemma Oc_Oc_Cons: "Oc # Oc # list = @ Bk\<^bsup>ln\<^esup> \ + lm \ [] \ hd lm > 0" +apply(case_tac lm, simp add: exponent_def, case_tac ln, simp, simp) +apply(case_tac lista, + auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def) +apply(case_tac [!] a, auto) +apply(case_tac ln, auto) +done +*) +lemma Oc_nil_zero[simp]: "[Oc] = @ Bk\<^bsup>ln\<^esup> + \ lm = [0] \ ln = 0" +apply(case_tac lm, simp) +apply(case_tac ln, auto simp: exponent_def) +apply(case_tac [!] list, + auto simp: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma [simp]: "Oc # r = @ replicate rn Bk \ + (\rn. r = replicate (hd lm2) Oc @ Bk # @ + replicate rn Bk) \ + tl lm2 = [] \ r = replicate (hd lm2) Oc" +apply(rule_tac disjCI, simp) +apply(case_tac "tl lm2 = []", simp) +apply(case_tac lm2, simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac rn, simp, simp, simp) +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exponent_def) +apply(case_tac rn, simp, simp) +apply(rule_tac x = rn in exI) +apply(simp add: hd_tl_tape_of_nat_list) +apply(simp add: tape_of_nat_abv Oc_def exponent_def) +done + +(*inv: from locate_a to locate_b*) +lemma [simp]: + "inv_locate_a (as, lm) (q, l, Oc # r) ires + \ inv_locate_b (as, lm) (q, Oc # l, r) ires" +apply(simp only: inv_locate_a.simps inv_locate_b.simps in_middle.simps + at_begin_norm.simps at_begin_fst_bwtn.simps + at_begin_fst_awtn.simps) +apply(erule disjE, erule exE, erule exE, erule exE) +apply(rule_tac x = lm1 in exI, rule_tac x = "tl lm2" in exI, simp) +apply(rule_tac x = "0" in exI, rule_tac x = "hd lm2" in exI, + auto simp: exponent_def) +apply(rule_tac x = "Suc 0" in exI, simp add:exponent_def) +apply(rule_tac x = "lm @ replicate tn 0" in exI, + rule_tac x = "[]" in exI, + rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI) +apply(simp only: rep_ind, simp) +apply(rule_tac x = "Suc 0" in exI, auto) +apply(case_tac [1-3] rn, simp_all ) +apply(rule_tac x = "lm @ replicate tn 0" in exI, + rule_tac x = "[]" in exI, + rule_tac x = "Suc tn" in exI, + rule_tac x = 0 in exI, simp add: rep_ind del: replicate_Suc split:if_splits) +apply(rule_tac x = "Suc 0" in exI, auto) +apply(case_tac rn, simp, simp) +apply(rule_tac [!] x = "Suc 0" in exI, auto) +apply(case_tac [!] rn, simp_all) +done + +(*inv: from locate_a to _locate_a*) +lemma locate_a_2_locate_a[simp]: "inv_locate_a (as, am) (q, aaa, Bk # xs) ires + \ inv_locate_a (as, am) (q, aaa, Oc # xs) ires" +apply(simp only: inv_locate_a.simps at_begin_norm.simps + at_begin_fst_bwtn.simps at_begin_fst_awtn.simps) +apply(erule_tac disjE, erule exE, erule exE, erule exE, + rule disjI2, rule disjI2) +defer +apply(erule_tac disjE, erule exE, erule exE, + erule exE, rule disjI2, rule disjI2) +prefer 2 +apply(simp) +proof- + fix lm1 tn rn + assume k: "lm1 = am @ 0\<^bsup>tn\<^esup> \ length lm1 = q \ (if lm1 = [] then aaa = Bk # Bk # + ires else aaa = [Bk] @ @ Bk # Bk # ires) \ Bk # xs = Bk\<^bsup>rn\<^esup>" + thus "\lm1 tn rn. lm1 = am @ 0\<^bsup>tn\<^esup> \ length lm1 = q \ (if lm1 = [] then + aaa = Bk # Bk # ires else aaa = [Bk] @ @ Bk # Bk # ires) \ Oc # xs = [Oc] @ Bk\<^bsup>rn\<^esup>" + (is "\lm1 tn rn. ?P lm1 tn rn") + proof - + from k have "?P lm1 tn (rn - 1)" + apply(auto simp: Oc_def) + by(case_tac [!] "rn::nat", auto simp: exponent_def) + thus ?thesis by blast + qed +next + fix lm1 lm2 rn + assume h1: "am = lm1 @ lm2 \ length lm1 = q \ (if lm1 = [] + then aaa = Bk # Bk # ires else aaa = Bk # @ Bk # Bk # ires) \ + Bk # xs = @ Bk\<^bsup>rn\<^esup>" + from h1 have h2: "lm2 = []" + proof(rule_tac xs = xs and rn = rn in BkCons_nil, simp) + qed + from h1 and h2 show "\lm1 tn rn. lm1 = am @ 0\<^bsup>tn\<^esup> \ length lm1 = q \ + (if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ @ Bk # Bk # ires) \ + Oc # xs = [Oc] @ Bk\<^bsup>rn\<^esup>" + (is "\lm1 tn rn. ?P lm1 tn rn") + proof - + from h1 and h2 have "?P lm1 0 (rn - 1)" + apply(auto simp: Oc_def exponent_def + tape_of_nl_abv tape_of_nat_list.simps) + by(case_tac "rn::nat", simp, simp) + thus ?thesis by blast + qed +qed + +lemma [intro]: "\rn. [a] = a\<^bsup>rn\<^esup>" +by(rule_tac x = "Suc 0" in exI, simp add: exponent_def) + +lemma [intro]: "\tn. [] = a\<^bsup>tn\<^esup>" +apply(rule_tac x = 0 in exI, simp add: exponent_def) +done + +lemma [intro]: "at_begin_norm (as, am) (q, aaa, []) ires + \ at_begin_norm (as, am) (q, aaa, [Bk]) ires" +apply(simp add: at_begin_norm.simps, erule_tac exE, erule_tac exE) +apply(rule_tac x = lm1 in exI, simp, auto) +done + +lemma [intro]: "at_begin_fst_bwtn (as, am) (q, aaa, []) ires + \ at_begin_fst_bwtn (as, am) (q, aaa, [Bk]) ires" +apply(simp only: at_begin_fst_bwtn.simps, erule_tac exE, erule_tac exE, erule_tac exE) +apply(rule_tac x = "am @ 0\<^bsup>tn\<^esup>" in exI, auto) +done + +lemma [intro]: "at_begin_fst_awtn (as, am) (q, aaa, []) ires + \ at_begin_fst_awtn (as, am) (q, aaa, [Bk]) ires" +apply(auto simp: at_begin_fst_awtn.simps) +done + +lemma [intro]: "inv_locate_a (as, am) (q, aaa, []) ires + \ inv_locate_a (as, am) (q, aaa, [Bk]) ires" +apply(simp only: inv_locate_a.simps) +apply(erule disj_forward) +defer +apply(erule disj_forward, auto) +done + +lemma [simp]: "inv_locate_a (as, am) (q, aaa, []) ires \ + inv_locate_a (as, am) (q, aaa, [Oc]) ires" +apply(insert locate_a_2_locate_a [of as am q aaa "[]"]) +apply(subgoal_tac "inv_locate_a (as, am) (q, aaa, [Bk]) ires", auto) +done + +(*inv: from locate_b to locate_b*) +lemma [simp]: "inv_locate_b (as, am) (q, aaa, Oc # xs) ires + \ inv_locate_b (as, am) (q, Oc # aaa, xs) ires" +apply(simp only: inv_locate_b.simps in_middle.simps) +apply(erule exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = tn in exI, rule_tac x = m in exI) +apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - 1" in exI, + rule_tac x = rn in exI) +apply(case_tac mr, simp_all add: exponent_def, auto) +done +lemma zero_and_nil[intro]: "(Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup> @ Bk # @ + Bk\<^bsup>rn \<^esup>) \ (lm2 = [] \ Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup>) + \ mr = 0 \ lm = []" +apply(rule context_conjI) +apply(case_tac mr, auto simp:exponent_def) +apply(insert BkCons_nil[of "replicate (n - 1) Bk" lm rn]) +apply(case_tac n, auto simp: exponent_def Bk_def tape_of_nl_nil_eq) +done + +lemma tape_of_nat_def: "<[m::nat]> = Oc # Oc\<^bsup>m\<^esup>" +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) +done +lemma [simp]: "\inv_locate_b (as, am) (q, aaa, Bk # xs) ires; \n. xs = Bk\<^bsup>n\<^esup>\ + \ inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" +apply(simp add: inv_locate_b.simps inv_locate_a.simps) +apply(rule_tac disjI2, rule_tac disjI1) +apply(simp only: in_middle.simps at_begin_fst_bwtn.simps) +apply(erule_tac exE)+ +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = tn in exI, simp) +apply(subgoal_tac "mr = 0 \ lm2 = []") +defer +apply(rule_tac n = n and mr = mr and lm = "lm2" + and rn = rn and n = n in zero_and_nil) +apply(auto simp: exponent_def) +apply(case_tac "lm1 = []", auto simp: tape_of_nat_def) +done + +lemma length_equal: "xs = ys \ length xs = length ys" +by auto +lemma [simp]: "a\<^bsup>0\<^esup> = []" +by(simp add: exp_zero) +(*inv: from locate_b to locate_a*) +lemma [simp]: "length (a\<^bsup>b\<^esup>) = b" +apply(simp add: exponent_def) +done + +lemma [simp]: "\inv_locate_b (as, am) (q, aaa, Bk # xs) ires; + \ (\n. xs = Bk\<^bsup>n\<^esup>)\ + \ inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" +apply(simp add: inv_locate_b.simps inv_locate_a.simps) +apply(rule_tac disjI1) +apply(simp only: in_middle.simps at_begin_norm.simps) +apply(erule_tac exE)+ +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = lm2 in exI, simp) +apply(subgoal_tac "tn = 0", simp add: exponent_def , auto split: if_splits) +apply(case_tac [!] mr, simp_all add: tape_of_nat_def, auto) +apply(case_tac lm2, simp, erule_tac x = rn in allE, simp) +apply(case_tac am, simp, simp) +apply(case_tac lm2, simp, erule_tac x = rn in allE, simp) +apply(drule_tac length_equal, simp) +done + +lemma locate_b_2_a[intro]: + "inv_locate_b (as, am) (q, aaa, Bk # xs) ires + \ inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" +apply(case_tac "\ n. xs = Bk\<^bsup>n\<^esup>", simp, simp) +done + +lemma locate_b_2_locate_a[simp]: + "\\ a < start_of ly as; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, Bk # xs) ires\ + \ (Suc a < start_of ly as + 2 * n \ inv_locate_a (as, am) + (Suc ((a - start_of ly as) div 2), Bk # aaa, xs) ires) \ + (\ Suc a < start_of ly as + 2 * n \ + inv_locate_a (as, am) (n, Bk # aaa, xs) ires)" +apply(auto) +apply(subgoal_tac "n > 0") +apply(subgoal_tac "(a - start_of ly as) div 2 = n - 1") +apply(insert locate_b_2_a [of as am "n - 1" aaa xs], simp) +apply(arith) +apply(case_tac n, simp, simp) +done + +lemma [simp]: "inv_locate_b (as, am) (q, l, []) ires + \ inv_locate_b (as, am) (q, l, [Bk]) ires" +apply(simp only: inv_locate_b.simps in_middle.simps) +apply(erule exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = tn in exI, rule_tac x = m in exI, + rule_tac x = ml in exI, rule_tac x = mr in exI) +apply(auto) +done + +lemma locate_b_2_locate_a_B[simp]: + "\\ a < start_of ly as; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, []) ires\ + \ (Suc a < start_of ly as + 2 * n \ + inv_locate_a (as, am) + (Suc ((a - start_of ly as) div 2), Bk # aaa, []) ires) + \ (\ Suc a < start_of ly as + 2 * n \ + inv_locate_a (as, am) (n, Bk # aaa, []) ires)" +apply(insert locate_b_2_locate_a [of a ly as n am aaa "[]"], simp) +done + +(*inv: from locate_b to after_write*) +lemma inv_locate_b_2_after_write[simp]: + "inv_locate_b (as, am) (n, aaa, Bk # xs) ires + \ inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) + (Suc (Suc (2 * n)), aaa, Oc # xs) ires" +apply(auto simp: in_middle.simps inv_after_write.simps + abc_lm_v.simps abc_lm_s.simps inv_locate_b.simps) +apply(subgoal_tac [!] "mr = 0", auto simp: exponent_def split: if_splits) +apply(subgoal_tac "lm2 = []", simp) +apply(rule_tac x = rn in exI, rule_tac x = "Suc m" in exI, + rule_tac x = "lm1" in exI, simp, rule_tac x = "[]" in exI, simp) +apply(case_tac "Suc (length lm1) - length am", simp, simp only: rep_ind, simp) +apply(subgoal_tac "length lm1 - length am = nat", simp, arith) +apply(drule_tac length_equal, simp) +done + +lemma [simp]: "inv_locate_b (as, am) (n, aaa, []) ires \ + inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) + (Suc (Suc (2 * n)), aaa, [Oc]) ires" +apply(insert inv_locate_b_2_after_write [of as am n aaa "[]"]) +by(simp) + +(*inv: from after_write to after_move*) +lemma [simp]: "inv_after_write (as, lm) (Suc (Suc (2 * n)), l, Oc # r) ires + \ inv_after_move (as, lm) (2 * n + 3, Oc # l, r) ires" +apply(auto simp:inv_after_move.simps inv_after_write.simps split: if_splits) +done + +lemma [simp]: "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n) + )) (Suc (Suc (2 * n)), aaa, Bk # xs) ires = False" +apply(simp add: inv_after_write.simps ) +done + +lemma [simp]: + "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) + (Suc (Suc (2 * n)), aaa, []) ires = False" +apply(simp add: inv_after_write.simps ) +done + +(*inv: from after_move to after_clear*) +lemma [simp]: "inv_after_move (as, lm) (s, l, Oc # r) ires + \ inv_after_clear (as, lm) (s', l, Bk # r) ires" +apply(auto simp: inv_after_move.simps inv_after_clear.simps split: if_splits) +done + +(*inv: from after_move to on_leftmoving*) +lemma inv_after_move_2_inv_on_left_moving[simp]: + "inv_after_move (as, lm) (s, l, Bk # r) ires + \ (l = [] \ + inv_on_left_moving (as, lm) (s', [], Bk # Bk # r) ires) \ + (l \ [] \ + inv_on_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)" +apply(simp only: inv_after_move.simps inv_on_left_moving.simps) +apply(subgoal_tac "l \ []", rule conjI, simp, rule impI, + rule disjI1, simp only: inv_on_left_moving_norm.simps) +apply(erule exE)+ +apply(subgoal_tac "lm2 = []") +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, rule_tac x = m in exI, + rule_tac x = 1 in exI, + rule_tac x = "rn - 1" in exI, simp, case_tac rn) +apply(auto simp: exponent_def intro: BkCons_nil split: if_splits) +done + +lemma [elim]: "[] = \ lm = []" +using tape_of_nl_nil_eq[of lm] +by simp + +lemma inv_after_move_2_inv_on_left_moving_B[simp]: + "inv_after_move (as, lm) (s, l, []) ires + \ (l = [] \ inv_on_left_moving (as, lm) (s', [], [Bk]) ires) \ + (l \ [] \ inv_on_left_moving (as, lm) (s', tl l, [hd l]) ires)" +apply(simp only: inv_after_move.simps inv_on_left_moving.simps) +apply(subgoal_tac "l \ []", rule conjI, simp, rule impI, rule disjI1, + simp only: inv_on_left_moving_norm.simps) +apply(erule exE)+ +apply(subgoal_tac "lm2 = []") +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, rule_tac x = m in exI, + rule_tac x = 1 in exI, rule_tac x = "rn - 1" in exI, simp, case_tac rn) +apply(auto simp: exponent_def tape_of_nl_nil_eq intro: BkCons_nil split: if_splits) +done + +(*inv: from after_clear to on_right_moving*) +lemma [simp]: "Oc # r = replicate rn Bk = False" +apply(case_tac rn, simp, simp) +done + +lemma inv_after_clear_2_inv_on_right_moving[simp]: + "inv_after_clear (as, lm) (2 * n + 4, l, Bk # r) ires + \ inv_on_right_moving (as, lm) (2 * n + 5, Bk # l, r) ires" +apply(auto simp: inv_after_clear.simps inv_on_right_moving.simps ) +apply(subgoal_tac "lm2 \ []") +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, + rule_tac x = "hd lm2" in exI, simp) +apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI) +apply(simp add: exponent_def, rule conjI) +apply(case_tac [!] "lm2::nat list", auto simp: exponent_def) +apply(case_tac rn, auto split: if_splits simp: tape_of_nat_def) +apply(case_tac list, + simp add: tape_of_nl_abv tape_of_nat_list.simps exponent_def) +apply(erule_tac x = "rn - 1" in allE, + case_tac rn, auto simp: exponent_def) +apply(case_tac list, + simp add: tape_of_nl_abv tape_of_nat_list.simps exponent_def) +apply(erule_tac x = "rn - 1" in allE, + case_tac rn, auto simp: exponent_def) +done + + +lemma [simp]: "inv_after_clear (as, lm) (2 * n + 4, l, []) ires\ + inv_after_clear (as, lm) (2 * n + 4, l, [Bk]) ires" +by(auto simp: inv_after_clear.simps) + +lemma [simp]: "inv_after_clear (as, lm) (2 * n + 4, l, []) ires + \ inv_on_right_moving (as, lm) (2 * n + 5, Bk # l, []) ires" +by(insert + inv_after_clear_2_inv_on_right_moving[of as lm n l "[]"], simp) + +(*inv: from on_right_moving to on_right_movign*) +lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, Oc # r) ires + \ inv_on_right_moving (as, lm) (2 * n + 5, Oc # l, r) ires" +apply(auto simp: inv_on_right_moving.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = "ml + mr" in exI, simp) +apply(rule_tac x = "Suc ml" in exI, + rule_tac x = "mr - 1" in exI, simp) +apply(case_tac mr, auto simp: exponent_def ) +apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, + rule_tac x = "ml + mr" in exI, simp) +apply(rule_tac x = "Suc ml" in exI, + rule_tac x = "mr - 1" in exI, simp) +apply(case_tac mr, auto split: if_splits simp: exponent_def) +done + +lemma inv_on_right_moving_2_inv_on_right_moving[simp]: + "inv_on_right_moving (as, lm) (2 * n + 5, l, Bk # r) ires + \ inv_after_write (as, lm) (Suc (Suc (2 * n)), l, Oc # r) ires" +apply(auto simp: inv_on_right_moving.simps inv_after_write.simps ) +apply(case_tac mr, auto simp: exponent_def split: if_splits) +apply(case_tac [!] mr, simp_all) +done + +lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, []) ires\ + inv_on_right_moving (as, lm) (2 * n + 5, l, [Bk]) ires" +apply(auto simp: inv_on_right_moving.simps exponent_def) +apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, simp) +apply (rule_tac x = m in exI, auto split: if_splits simp: exponent_def) +done + +(*inv: from on_right_moving to after_write*) +lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, []) ires + \ inv_after_write (as, lm) (Suc (Suc (2 * n)), l, [Oc]) ires" +apply(rule_tac inv_on_right_moving_2_inv_on_right_moving, simp) +done + +(*inv: from on_left_moving to on_left_moving*) +lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) + (s, l, Oc # r) ires = False" +apply(auto simp: inv_on_left_moving_in_middle_B.simps ) +done + +lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, Bk # r) ires + = False" +apply(auto simp: inv_on_left_moving_norm.simps) +apply(case_tac [!] mr, auto simp: ) +done + +lemma [intro]: "\rna. Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk\<^bsup>rn\<^esup> = @ Bk\<^bsup>rna\<^esup>" +apply(case_tac lm, simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(rule_tac x = "Suc rn" in exI, simp) +apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps, auto) +done + + +lemma [simp]: + "\inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires; + hd l = Bk; l \ []\ \ + inv_on_left_moving_in_middle_B (as, lm) (s, tl l, Bk # Oc # r) ires" +apply(case_tac l, simp, simp) +apply(simp only: inv_on_left_moving_norm.simps + inv_on_left_moving_in_middle_B.simps) +apply(erule_tac exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = "m # lm2" in exI, auto) +apply(case_tac [!] ml, auto) +apply(rule_tac [!] x = 0 in exI, simp_all add: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma [simp]: "\inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires; + hd l = Oc; l \ []\ + \ inv_on_left_moving_norm (as, lm) + (s, tl l, Oc # Oc # r) ires" +apply(simp only: inv_on_left_moving_norm.simps) +apply(erule exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, rule_tac x = "ml - 1" in exI, + rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, simp) +apply(case_tac ml, auto simp: exponent_def split: if_splits) +done + +lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, [], Oc # r) ires + \ inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # Oc # r) ires" +apply(auto simp: inv_on_left_moving_norm.simps + inv_on_left_moving_in_middle_B.simps split: if_splits) +done + +lemma [simp]:"inv_on_left_moving (as, lm) (s, l, Oc # r) ires + \ (l = [] \ inv_on_left_moving (as, lm) (s, [], Bk # Oc # r) ires) + \ (l \ [] \ inv_on_left_moving (as, lm) (s, tl l, hd l # Oc # r) ires)" +apply(simp add: inv_on_left_moving.simps) +apply(case_tac "l \ []", rule conjI, simp, simp) +apply(case_tac "hd l", simp, simp, simp) +done + +(*inv: from on_left_moving to check_left_moving*) +lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) + (s, Bk # list, Bk # r) ires + \ inv_check_left_moving_on_leftmost (as, lm) + (s', list, Bk # Bk # r) ires" +apply(auto simp: inv_on_left_moving_in_middle_B.simps + inv_check_left_moving_on_leftmost.simps split: if_splits) +apply(case_tac [!] "rev lm1", simp_all) +apply(case_tac [!] lista, simp_all add: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma [simp]: + "inv_check_left_moving_in_middle (as, lm) (s, l, Bk # r) ires= False" +by(auto simp: inv_check_left_moving_in_middle.simps ) + +lemma [simp]: + "inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # r) ires\ + inv_check_left_moving_on_leftmost (as, lm) (s', [], Bk # Bk # r) ires" +apply(auto simp: inv_on_left_moving_in_middle_B.simps + inv_check_left_moving_on_leftmost.simps split: if_splits) +done + + +lemma [simp]: "inv_check_left_moving_on_leftmost (as, lm) + (s, list, Oc # r) ires= False" +by(auto simp: inv_check_left_moving_on_leftmost.simps split: if_splits) + +lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) + (s, Oc # list, Bk # r) ires + \ inv_check_left_moving_in_middle (as, lm) (s', list, Oc # Bk # r) ires" +apply(auto simp: inv_on_left_moving_in_middle_B.simps + inv_check_left_moving_in_middle.simps split: if_splits) +done + +lemma inv_on_left_moving_2_check_left_moving[simp]: + "inv_on_left_moving (as, lm) (s, l, Bk # r) ires + \ (l = [] \ inv_check_left_moving (as, lm) (s', [], Bk # Bk # r) ires) + \ (l \ [] \ + inv_check_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)" +apply(simp add: inv_on_left_moving.simps inv_check_left_moving.simps) +apply(case_tac l, simp, simp) +apply(case_tac a, simp, simp) +done + +lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, []) ires = False" +apply(auto simp: inv_on_left_moving_norm.simps) +by(case_tac [!] mr, auto) + +lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires\ + inv_on_left_moving (as, lm) (6 + 2 * n, l, [Bk]) ires" +apply(simp add: inv_on_left_moving.simps) +apply(auto simp: inv_on_left_moving_in_middle_B.simps) +done + +lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires = False" +apply(simp add: inv_on_left_moving.simps) +apply(simp add: inv_on_left_moving_in_middle_B.simps) +done + +lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires + \ (l = [] \ inv_check_left_moving (as, lm) (s', [], [Bk]) ires) \ + (l \ [] \ inv_check_left_moving (as, lm) (s', tl l, [hd l]) ires)" +by simp + +lemma Oc_Bk_Cons_ex[simp]: + "Oc # Bk # list = @ Bk\<^bsup>ln\<^esup> \ + \ln. list = @ Bk\<^bsup>ln\<^esup>" +apply(case_tac "lm", simp) +apply(case_tac ln, simp_all add: exponent_def) +apply(case_tac lista, + auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def) +apply(case_tac [!] a, auto simp: ) +apply(case_tac ln, simp, rule_tac x = nat in exI, simp) +done + +lemma [simp]: + "Oc # Bk # list = @ Bk\<^bsup>ln\<^esup> \ + \rna. Oc # Bk # @ Bk\<^bsup>rn\<^esup> = @ Bk\<^bsup>rna\<^esup>" +apply(frule Oc_Bk_Cons, simp) +apply(case_tac lm2, + auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def ) +apply(rule_tac x = "Suc rn" in exI, simp) +done + +(*inv: from check_left_moving to on_left_moving*) +lemma [intro]: "\rna. a # a\<^bsup>rn\<^esup> = a\<^bsup>rna\<^esup>" +apply(rule_tac x = "Suc rn" in exI, simp) +done + +lemma +inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[simp]: +"inv_check_left_moving_in_middle (as, lm) (s, Bk # list, Oc # r) ires + \ inv_on_left_moving_in_middle_B (as, lm) (s', list, Bk # Oc # r) ires" +apply(simp only: inv_check_left_moving_in_middle.simps + inv_on_left_moving_in_middle_B.simps) +apply(erule_tac exE)+ +apply(rule_tac x = "rev (tl (rev lm1))" in exI, + rule_tac x = "[hd (rev lm1)] @ lm2" in exI, auto) +apply(case_tac [!] "rev lm1",simp_all add: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac [!] a, simp_all) +apply(case_tac [1] lm2, simp_all add: tape_of_nat_list.simps, auto) +apply(case_tac [3] lm2, simp_all add: tape_of_nat_list.simps, auto) +apply(case_tac [!] lista, simp_all add: tape_of_nat_list.simps) +done + +lemma [simp]: + "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires\ + inv_check_left_moving_in_middle (as, lm) (s', [Bk], Oc # r) ires" +apply(auto simp: inv_check_left_moving_in_middle.simps ) +done + +lemma [simp]: + "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires + \ inv_on_left_moving_in_middle_B (as, lm) (s', [], Bk # Oc # r) ires" +apply(insert +inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[of + as lm n "[]" r], simp) +done + +lemma [simp]: "a\<^bsup>0\<^esup> = []" +apply(simp add: exponent_def) +done + +lemma [simp]: "inv_check_left_moving_in_middle (as, lm) + (s, Oc # list, Oc # r) ires + \ inv_on_left_moving_norm (as, lm) (s', list, Oc # Oc # r) ires" +apply(auto simp: inv_check_left_moving_in_middle.simps + inv_on_left_moving_norm.simps) +apply(rule_tac x = "rev (tl (rev lm1))" in exI, + rule_tac x = lm2 in exI, rule_tac x = "hd (rev lm1)" in exI) +apply(rule_tac conjI) +apply(case_tac "rev lm1", simp, simp) +apply(rule_tac x = "hd (rev lm1) - 1" in exI, auto) +apply(rule_tac [!] x = "Suc (Suc 0)" in exI, simp) +apply(case_tac [!] "rev lm1", simp_all) +apply(case_tac [!] a, simp_all add: tape_of_nl_abv tape_of_nat_list.simps, auto) +done + +lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Oc # r) ires +\ (l = [] \ inv_on_left_moving (as, lm) (s', [], Bk # Oc # r) ires) \ + (l \ [] \ inv_on_left_moving (as, lm) (s', tl l, hd l # Oc # r) ires)" +apply(case_tac l, + auto simp: inv_check_left_moving.simps inv_on_left_moving.simps) +apply(case_tac a, simp, simp) +done + +(*inv: check_left_moving to after_left_moving*) +lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Bk # r) ires + \ inv_after_left_moving (as, lm) (s', Bk # l, r) ires" +apply(auto simp: inv_check_left_moving.simps + inv_check_left_moving_on_leftmost.simps inv_after_left_moving.simps) +done + + +lemma [simp]:"inv_check_left_moving (as, lm) (s, l, []) ires + \ inv_after_left_moving (as, lm) (s', Bk # l, []) ires" +by(simp add: inv_check_left_moving.simps +inv_check_left_moving_in_middle.simps +inv_check_left_moving_on_leftmost.simps) + +(*inv: after_left_moving to inv_stop*) +lemma [simp]: "inv_after_left_moving (as, lm) (s, l, Bk # r) ires + \ inv_stop (as, lm) (s', Bk # l, r) ires" +apply(auto simp: inv_after_left_moving.simps inv_stop.simps) +done + +lemma [simp]: "inv_after_left_moving (as, lm) (s, l, []) ires + \ inv_stop (as, lm) (s', Bk # l, []) ires" +by(auto simp: inv_after_left_moving.simps) + +(*inv: stop to stop*) +lemma [simp]: "inv_stop (as, lm) (x, l, r) ires \ + inv_stop (as, lm) (y, l, r) ires" +apply(simp add: inv_stop.simps) +done + +lemma [simp]: "inv_after_clear (as, lm) (s, aaa, Oc # xs) ires= False" +apply(auto simp: inv_after_clear.simps ) +done + +lemma [simp]: + "inv_after_left_moving (as, lm) (s, aaa, Oc # xs) ires = False" +by(auto simp: inv_after_left_moving.simps ) + +lemma start_of_not0: "as \ 0 \ start_of ly as > 0" +apply(rule startof_not0) +done + +text {* + The single step currectness of the TM complied from Abacus instruction @{text "Inc n"}. + It shows every single step execution of this TM keeps the invariant. +*} + +lemma inc_inv_step: + assumes + -- {* Invariant holds on the start *} + h11: "inc_inv ly n (as, am) tc ires" + -- {* The layout of Abacus program @{text "aprog"} is @{text "ly"} *} + and h12: "ly = layout_of aprog" + -- {* The instruction at position @{text "as"} is @{text "Inc n"} *} + and h21: "abc_fetch as aprog = Some (Inc n)" + -- {* TM not yet reach the final state, where @{text "start_of ly as + 2*n + 9"} is the state + where the current TM stops and the next TM starts. *} + and h22: "(\ (s, l, r). s \ start_of ly as + 2*n + 9) tc" + shows + -- {* + Single step execution of the TM keeps the invaraint, where + the TM compiled from @{text "Inc n"} is @{text "(ci ly (start_of ly as) (Inc n))"} + @{text "start_of ly as - Suc 0)"} is the offset used to execute this {\em shifted} + TM. + *} + "inc_inv ly n (as, am) (t_step tc (ci ly (start_of ly as) (Inc n), start_of ly as - Suc 0)) ires" +proof - + from h21 h22 have h3 : "start_of (layout_of aprog) as > 0" + apply(case_tac as, simp add: start_of.simps abc_fetch.simps) + apply(insert start_of_not0[of as "layout_of aprog"], simp) + done + from h11 h12 and h21 h22 and this show ?thesis + apply(case_tac tc, simp) + apply(case_tac "a = 0", + auto split:if_splits simp add:t_step.simps, + tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *}) + apply (simp_all add:fetch_simps new_tape.simps) + done +qed + + +lemma t_steps_ind: "t_steps tc (p, off) (Suc n) + = t_step (t_steps tc (p, off) n) (p, off)" +apply(induct n arbitrary: tc) +apply(simp add: t_steps.simps) +apply(simp add: t_steps.simps) +done + +definition lex_pair :: "((nat \ nat) \ (nat \ nat)) set" + where + "lex_pair \ less_than <*lex*> less_than" + +definition lex_triple :: + "((nat \ (nat \ nat)) \ (nat \ (nat \ nat))) set" + where "lex_triple \ less_than <*lex*> lex_pair" + +definition lex_square :: + "((nat \ nat \ nat \ nat) \ (nat \ nat \ nat \ nat)) set" + where "lex_square \ less_than <*lex*> lex_triple" + +fun abc_inc_stage1 :: "t_conf \ nat \ nat \ nat" + where + "abc_inc_stage1 (s, l, r) ss n = + (if s = 0 then 0 + else if s \ ss+2*n+1 then 5 + else if s\ ss+2*n+5 then 4 + else if s \ ss+2*n+7 then 3 + else if s = ss+2*n+8 then 2 + else 1)" + +fun abc_inc_stage2 :: "t_conf \ nat \ nat \ nat" + where + "abc_inc_stage2 (s, l, r) ss n = + (if s \ ss + 2*n + 1 then 0 + else if s = ss + 2*n + 2 then length r + else if s = ss + 2*n + 3 then length r + else if s = ss + 2*n + 4 then length r + else if s = ss + 2*n + 5 then + if r \ [] then length r + else 1 + else if s = ss+2*n+6 then length l + else if s = ss+2*n+7 then length l + else 0)" + +fun abc_inc_stage3 :: "t_conf \ nat \ nat \ block list \ nat" + where + "abc_inc_stage3 (s, l, r) ss n ires = ( + if s = ss + 2*n + 3 then 4 + else if s = ss + 2*n + 4 then 3 + else if s = ss + 2*n + 5 then + if r \ [] \ hd r = Oc then 2 + else 1 + else if s = ss + 2*n + 2 then 0 + else if s = ss + 2*n + 6 then + if l = Bk # ires \ r \ [] \ hd r = Oc then 2 + else 1 + else if s = ss + 2*n + 7 then + if r \ [] \ hd r = Oc then 3 + else 0 + else ss+2*n+9 - s)" + +fun abc_inc_stage4 :: "t_conf \ nat \ nat \ block list \ nat" + where + "abc_inc_stage4 (s, l, r) ss n ires = + (if s \ ss+2*n+1 \ (s - ss) mod 2 = 0 then + if (r\[] \ hd r = Oc) then 0 + else 1 + else if (s \ ss+2*n+1 \ (s - ss) mod 2 = Suc 0) + then length r + else if s = ss + 2*n + 6 then + if l = Bk # ires \ hd r = Bk then 0 + else Suc (length l) + else 0)" + +fun abc_inc_measure :: "(t_conf \ nat \ nat \ block list) \ + (nat \ nat \ nat \ nat)" + where + "abc_inc_measure (c, ss, n, ires) = + (abc_inc_stage1 c ss n, abc_inc_stage2 c ss n, + abc_inc_stage3 c ss n ires, abc_inc_stage4 c ss n ires)" + +definition abc_inc_LE :: "(((nat \ block list \ block list) \ nat \ + nat \ block list) \ ((nat \ block list \ block list) \ nat \ nat \ block list)) set" + where "abc_inc_LE \ (inv_image lex_square abc_inc_measure)" + +lemma wf_lex_triple: "wf lex_triple" +by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def) + +lemma wf_lex_square: "wf lex_square" +by (auto intro:wf_lex_triple simp:lex_triple_def lex_square_def lex_pair_def) + +lemma wf_abc_inc_le[intro]: "wf abc_inc_LE" +by(auto intro:wf_inv_image wf_lex_square simp:abc_inc_LE_def) + +(********************************************************************) +declare inc_inv.simps[simp del] + +lemma halt_lemma2': + "\wf LE; \ n. ((\ P (f n) \ Q (f n)) \ + (Q (f (Suc n)) \ (f (Suc n), (f n)) \ LE)); Q (f 0)\ + \ \ n. P (f n)" +apply(intro exCI, simp) +apply(subgoal_tac "\ n. Q (f n)", simp) +apply(drule_tac f = f in wf_inv_image) +apply(simp add: inv_image_def) +apply(erule wf_induct, simp) +apply(erule_tac x = x in allE) +apply(erule_tac x = n in allE, erule_tac x = n in allE) +apply(erule_tac x = "Suc x" in allE, simp) +apply(rule_tac allI) +apply(induct_tac n, simp) +apply(erule_tac x = na in allE, simp) +done + +lemma halt_lemma2'': + "\P (f n); \ P (f (0::nat))\ \ + \ n. (P (f n) \ (\ i < n. \ P (f i)))" +apply(induct n rule: nat_less_induct, auto) +done + +lemma halt_lemma2''': + "\\n. \ P (f n) \ Q (f n) \ Q (f (Suc n)) \ (f (Suc n), f n) \ LE; + Q (f 0); \i P (f i)\ \ Q (f na)" +apply(induct na, simp, simp) +done + +lemma halt_lemma2: + "\wf LE; + \ n. ((\ P (f n) \ Q (f n)) \ (Q (f (Suc n)) \ (f (Suc n), (f n)) \ LE)); + Q (f 0); \ P (f 0)\ + \ \ n. P (f n) \ Q (f n)" +apply(insert halt_lemma2' [of LE P f Q], simp, erule_tac exE) +apply(subgoal_tac "\ n. (P (f n) \ (\ i < n. \ P (f i)))") +apply(erule_tac exE)+ +apply(rule_tac x = na in exI, auto) +apply(rule halt_lemma2''', simp, simp, simp) +apply(erule_tac halt_lemma2'', simp) +done + +lemma [simp]: + "\ly = layout_of aprog; abc_fetch as aprog = Some (Inc n)\ + \ start_of ly (Suc as) = start_of ly as + 2*n +9" +apply(case_tac as, auto simp: abc_fetch.simps start_of.simps + layout_of.simps length_of.simps split: if_splits) +done + +lemma inc_inv_init: + "\abc_fetch as aprog = Some (Inc n); + crsp_l ly (as, am) (start_of ly as, l, r) ires; ly = layout_of aprog\ + \ inc_inv ly n (as, am) (start_of ly as, l, r) ires" +apply(auto simp: crsp_l.simps inc_inv.simps + inv_locate_a.simps at_begin_fst_bwtn.simps + at_begin_fst_awtn.simps at_begin_norm.simps ) +apply(auto intro: startof_not0) +done + +lemma inc_inv_stop_pre[simp]: + "\ly = layout_of aprog; inc_inv ly n (as, am) (s, l, r) ires; + s = start_of ly as; abc_fetch as aprog = Some (Inc n)\ + \ (\na. \ (\((s, l, r), ss, n', ires'). s = start_of ly (Suc as)) + (t_steps (s, l, r) (ci ly (start_of ly as) + (Inc n), start_of ly as - Suc 0) na, s, n, ires) \ + (\((s, l, r), ss, n', ires'). inc_inv ly n (as, am) (s, l, r) ires') + (t_steps (s, l, r) (ci ly (start_of ly as) + (Inc n), start_of ly as - Suc 0) na, s, n, ires) \ + (\((s, l, r), ss, n', ires'). inc_inv ly n (as, am) (s, l, r) ires') + (t_steps (s, l, r) (ci ly (start_of ly as) + (Inc n), start_of ly as - Suc 0) (Suc na), s, n, ires) \ + ((t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) (Suc na), s, n, ires), + t_steps (s, l, r) (ci ly (start_of ly as) + (Inc n), start_of ly as - Suc 0) na, s, n, ires) \ abc_inc_LE)" +apply(rule allI, rule impI, simp add: t_steps_ind, + rule conjI, erule_tac conjE) +apply(rule_tac inc_inv_step, simp, simp, simp) +apply(case_tac "t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na", simp) +proof - + fix na + assume h1: "abc_fetch as aprog = Some (Inc n)" + "\ (\(s, l, r) (ss, n', ires'). s = start_of (layout_of aprog) as + 2 * n + 9) + (t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) + (start_of (layout_of aprog) as, n, ires) \ + inc_inv (layout_of aprog) n (as, am) (t_steps (start_of (layout_of aprog) as, l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) ires" + from h1 have h2: "start_of (layout_of aprog) as > 0" + apply(rule_tac startof_not0) + done + from h1 and h2 show "((t_step (t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0), + start_of (layout_of aprog) as, n, ires), + t_steps (start_of (layout_of aprog) as, l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na, + start_of (layout_of aprog) as, n, ires) + \ abc_inc_LE" + apply(case_tac "(t_steps (start_of (layout_of aprog) as, l, r) + (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), + start_of (layout_of aprog) as - Suc 0) na)", simp) + apply(case_tac "a = 0", + auto split:if_splits simp add:t_step.simps inc_inv.simps, + tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *}) + apply(simp_all add:fetch_simps new_tape.simps) + apply(auto simp add: abc_inc_LE_def + lex_square_def lex_triple_def lex_pair_def + inv_after_write.simps inv_after_move.simps inv_after_clear.simps + inv_on_left_moving.simps inv_on_left_moving_norm.simps split: if_splits) + done +qed + +lemma inc_inv_stop_pre1: + "\ + ly = layout_of aprog; + abc_fetch as aprog = Some (Inc n); + s = start_of ly as; + inc_inv ly n (as, am) (s, l, r) ires + \ \ + (\ stp > 0. (\ (s', l', r'). + s' = start_of ly (Suc as) \ + inc_inv ly n (as, am) (s', l', r') ires) + (t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) stp))" +apply(insert halt_lemma2[of abc_inc_LE + "\ ((s, l, r), ss, n', ires'). s = start_of ly (Suc as)" + "(\ stp. (t_steps (s, l, r) + (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) stp, s, n, ires))" + "\ ((s, l, r), ss, n'). inc_inv ly n (as, am) (s, l, r) ires"]) +apply(insert wf_abc_inc_le) +apply(insert inc_inv_stop_pre[of ly aprog n as am s l r ires], simp) +apply(simp only: t_steps.simps, auto) +apply(rule_tac x = na in exI) +apply(case_tac "(t_steps (start_of (layout_of aprog) as, l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n), start_of (layout_of aprog) as - Suc 0) na)", simp) +apply(case_tac na, simp add: t_steps.simps, simp) +done + +lemma inc_inv_stop: + assumes program_and_layout: + -- {* There is an Abacus program @{text "aprog"} and its layout is @{text "ly"}: *} + "ly = layout_of aprog" + and an_instruction: + -- {* There is an instruction @{text "Inc n"} at postion @{text "as"} of @{text "aprog"} *} + "abc_fetch as aprog = Some (Inc n)" + and the_start_state: + -- {* According to @{text "ly"} and @{text "as"}, + the start state of the TM compiled from this + @{text "Inc n"} instruction should be @{text "s"}: + *} + "s = start_of ly as" + and inv: + -- {* Invariant holds on configuration @{text "(s, l, r)"} *} + "inc_inv ly n (as, am) (s, l, r) ires" + shows -- {* After @{text "stp"} steps of execution, the compiled + TM reaches the start state of next compiled TM and the invariant + still holds. + *} + "(\ stp > 0. (\ (s', l', r'). + s' = start_of ly (Suc as) \ + inc_inv ly n (as, am) (s', l', r') ires) + (t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) stp))" +proof - + from inc_inv_stop_pre1 [OF program_and_layout an_instruction the_start_state inv] + show ?thesis . +qed + +lemma inc_inv_stop_cond: + "\ly = layout_of aprog; + s' = start_of ly (as + 1); + inc_inv ly n (as, lm) (s', (l', r')) ires; + abc_fetch as aprog = Some (Inc n)\ \ + crsp_l ly (Suc as, abc_lm_s lm n (Suc (abc_lm_v lm n))) + (s', l', r') ires" +apply(subgoal_tac "s' = start_of ly as + 2*n + 9", simp) +apply(auto simp: inc_inv.simps inv_stop.simps crsp_l.simps ) +done + +lemma inc_crsp_ex_pre: + "\ly = layout_of aprog; + crsp_l ly (as, am) tc ires; + abc_fetch as aprog = Some (Inc n)\ + \ \stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Inc n))) + (t_steps tc (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) stp) ires" +proof(case_tac tc, simp add: abc_step_l.simps) + fix a b c + assume h1: "ly = layout_of aprog" + "crsp_l (layout_of aprog) (as, am) (a, b, c) ires" + "abc_fetch as aprog = Some (Inc n)" + hence h2: "a = start_of ly as" + by(auto simp: crsp_l.simps) + from h1 and h2 have h3: + "inc_inv ly n (as, am) (start_of ly as, b, c) ires" + by(rule_tac inc_inv_init, simp, simp, simp) + from h1 and h2 and h3 have h4: + "(\ stp > 0. (\ (s', l', r'). s' = + start_of ly (Suc as) \ inc_inv ly n (as, am) (s', l', r') ires) + (t_steps (a, b, c) (ci ly (start_of ly as) + (Inc n), start_of ly as - Suc 0) stp))" + apply(rule_tac inc_inv_stop, auto) + done + from h1 and h2 and h3 and h4 show + "\stp > 0. crsp_l (layout_of aprog) + (Suc as, abc_lm_s am n (Suc (abc_lm_v am n))) + (t_steps (a, b, c) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), + start_of (layout_of aprog) as - Suc 0) stp) ires" + apply(erule_tac exE) + apply(rule_tac x = stp in exI) + apply(case_tac "(t_steps (a, b, c) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), + start_of (layout_of aprog) as - Suc 0) stp)", simp) + apply(rule_tac inc_inv_stop_cond, auto) + done +qed + +text {* + The total correctness of the compilaton of @{text "Inc n"} instruction. +*} + +lemma inc_crsp_ex: + assumes layout: + -- {* For any Abacus program @{text "aprog"}, assuming its layout is @{text "ly"} *} + "ly = layout_of aprog" + and corresponds: + -- {* Abacus configuration @{text "(as, am)"} is in correspondence with + TM configuration @{text "tc"} *} + "crsp_l ly (as, am) tc ires" + and inc: + -- {* There is an instruction @{text "Inc n"} at postion @{text "as"} of @{text "aprog"} *} + "abc_fetch as aprog = Some (Inc n)" + shows + -- {* + After @{text "stp"} steps of execution, the TM compiled from this @{text "Inc n"} + stops with a configuration which corresponds to the Abacus configuration obtained + from the execution of this @{text "Inc n"} instruction. + *} + "\stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Inc n))) + (t_steps tc (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) stp) ires" +proof - + from inc_crsp_ex_pre [OF layout corresponds inc] show ?thesis . +qed + +(*******End: inc crsp********) + +(*******Begin: dec crsp******) + +subsubsection {* The compilation of @{text "Dec n e"} *} + + +text {* + The lemmas in this section lead to the correctness of the compilation + of @{text "Dec n e"} instruction using the same techniques as + @{text "Inc n"}. +*} + +type_synonym dec_inv_t = "(nat * nat list) \ t_conf \ block list \ bool" + +fun dec_first_on_right_moving :: "nat \ dec_inv_t" + where + "dec_first_on_right_moving n (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = Suc m \ length lm1 = n \ ml > 0 \ m > 0 \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ @ Bk # Bk # ires) \ + ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ @ (Bk\<^bsup>rn\<^esup>)) \ (r = (Oc\<^bsup>mr\<^esup>) \ lm2 = [])))" + +fun dec_on_right_moving :: "dec_inv_t" + where + "dec_on_right_moving (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = Suc (Suc m) \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ @ Bk # Bk # ires) \ + ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ @ (Bk\<^bsup>rn\<^esup>)) \ (r = (Oc\<^bsup>mr\<^esup>) \ lm2 = [])))" + +fun dec_after_clear :: "dec_inv_t" + where + "dec_after_clear (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = Suc m \ ml = Suc m \ r \ [] \ r \ [] \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml \<^esup>) @ [Bk] @ @ Bk # Bk # ires) \ + (tl r = Bk # @ (Bk\<^bsup>rn\<^esup>) \ tl r = [] \ lm2 = []))" + +fun dec_after_write :: "dec_inv_t" + where + "dec_after_write (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = Suc m \ ml = Suc m \ lm2 \ [] \ + (if lm1 = [] then l = Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = Bk # (Oc\<^bsup>ml \<^esup>) @ [Bk] @ @ Bk # Bk # ires) \ + tl r = @ (Bk\<^bsup>rn\<^esup>))" + +fun dec_right_move :: "dec_inv_t" + where + "dec_right_move (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 + \ ml = Suc m \ mr = (0::nat) \ + (if lm1 = [] then l = Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = Bk # Oc\<^bsup>ml\<^esup>@ [Bk] @ @ Bk # Bk # ires) + \ (r = Bk # @ Bk\<^bsup>rn\<^esup>\ r = [] \ lm2 = []))" + +fun dec_check_right_move :: "dec_inv_t" + where + "dec_check_right_move (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml = Suc m \ mr = (0::nat) \ + (if lm1 = [] then l = Bk # Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = Bk # Bk # Oc\<^bsup>ml \<^esup>@ [Bk] @ @ Bk # Bk # ires) \ + r = @ Bk\<^bsup>rn\<^esup>)" + +fun dec_left_move :: "dec_inv_t" + where + "dec_left_move (as, lm) (s, l, r) ires = + (\ lm1 m rn. (lm::nat list) = lm1 @ [m::nat] \ + rn > 0 \ + (if lm1 = [] then l = Bk # Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires + else l = Bk # Oc\<^bsup>Suc m\<^esup> @ Bk # @ Bk # Bk # ires) \ r = Bk\<^bsup>rn\<^esup>)" + +declare + dec_on_right_moving.simps[simp del] dec_after_clear.simps[simp del] + dec_after_write.simps[simp del] dec_left_move.simps[simp del] + dec_check_right_move.simps[simp del] dec_right_move.simps[simp del] + dec_first_on_right_moving.simps[simp del] + +fun inv_locate_n_b :: "inc_inv_t" + where + "inv_locate_n_b (as, lm) (s, l, r) ires= + (\ lm1 lm2 tn m ml mr rn. lm @ 0\<^bsup>tn\<^esup> = lm1 @ [m] @ lm2 \ + length lm1 = s \ m + 1 = ml + mr \ + ml = 1 \ tn = s + 1 - length lm \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = Oc\<^bsup>ml\<^esup>@Bk#@Bk#Bk#ires) \ + (r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ @ (Bk\<^bsup>rn\<^esup>) \ (lm2 = [] \ r = (Oc\<^bsup>mr\<^esup>))) + )" + +fun dec_inv_1 :: "layout \ nat \ nat \ dec_inv_t" + where + "dec_inv_1 ly n e (as, am) (s, l, r) ires = + (let ss = start_of ly as in + let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in + let am'' = abc_lm_s am n (abc_lm_v am n) in + if s = start_of ly e then inv_stop (as, am'') (s, l, r) ires + else if s = ss then False + else if ss \ s \ s < ss + 2*n then + if (s - ss) mod 2 = 0 then + inv_locate_a (as, am) ((s - ss) div 2, l, r) ires + \ inv_locate_a (as, am'') ((s - ss) div 2, l, r) ires + else + inv_locate_b (as, am) ((s - ss) div 2, l, r) ires + \ inv_locate_b (as, am'') ((s - ss) div 2, l, r) ires + else if s = ss + 2 * n then + inv_locate_a (as, am) (n, l, r) ires + \ inv_locate_a (as, am'') (n, l, r) ires + else if s = ss + 2 * n + 1 then + inv_locate_b (as, am) (n, l, r) ires + else if s = ss + 2 * n + 13 then + inv_on_left_moving (as, am'') (s, l, r) ires + else if s = ss + 2 * n + 14 then + inv_check_left_moving (as, am'') (s, l, r) ires + else if s = ss + 2 * n + 15 then + inv_after_left_moving (as, am'') (s, l, r) ires + else False)" + +fun dec_inv_2 :: "layout \ nat \ nat \ dec_inv_t" + where + "dec_inv_2 ly n e (as, am) (s, l, r) ires = + (let ss = start_of ly as in + let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in + let am'' = abc_lm_s am n (abc_lm_v am n) in + if s = 0 then False + else if s = ss then False + else if ss \ s \ s < ss + 2*n then + if (s - ss) mod 2 = 0 then + inv_locate_a (as, am) ((s - ss) div 2, l, r) ires + else inv_locate_b (as, am) ((s - ss) div 2, l, r) ires + else if s = ss + 2 * n then + inv_locate_a (as, am) (n, l, r) ires + else if s = ss + 2 * n + 1 then + inv_locate_n_b (as, am) (n, l, r) ires + else if s = ss + 2 * n + 2 then + dec_first_on_right_moving n (as, am'') (s, l, r) ires + else if s = ss + 2 * n + 3 then + dec_after_clear (as, am') (s, l, r) ires + else if s = ss + 2 * n + 4 then + dec_right_move (as, am') (s, l, r) ires + else if s = ss + 2 * n + 5 then + dec_check_right_move (as, am') (s, l, r) ires + else if s = ss + 2 * n + 6 then + dec_left_move (as, am') (s, l, r) ires + else if s = ss + 2 * n + 7 then + dec_after_write (as, am') (s, l, r) ires + else if s = ss + 2 * n + 8 then + dec_on_right_moving (as, am') (s, l, r) ires + else if s = ss + 2 * n + 9 then + dec_after_clear (as, am') (s, l, r) ires + else if s = ss + 2 * n + 10 then + inv_on_left_moving (as, am') (s, l, r) ires + else if s = ss + 2 * n + 11 then + inv_check_left_moving (as, am') (s, l, r) ires + else if s = ss + 2 * n + 12 then + inv_after_left_moving (as, am') (s, l, r) ires + else if s = ss + 2 * n + 16 then + inv_stop (as, am') (s, l, r) ires + else False)" + +(*begin: dec_fetch lemmas*) + +lemma dec_fetch_locate_a_o: + "\start_of ly as \ a; + a < start_of ly as + 2 * n; start_of ly as > 0; + a - start_of ly as = 2 * q\ + \ fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (2 * q)) Oc = (R, a + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append Suc_pre) +apply(subgoal_tac "(findnth n ! Suc (4 * q)) = + findnth (Suc q) ! (4 * q + 1)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n !(4 * q + 1) = + findnth (Suc q) ! (4 * q + 1)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma dec_fetch_locate_a_b: + "\start_of ly as \ a; + a < start_of ly as + 2 * n; + start_of ly as > 0; + a - start_of ly as = 2 * q\ + \ fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) + (Suc (2 * q)) Bk = (W1, a)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append) +apply(subgoal_tac "(findnth n ! (4 * q)) = + findnth (Suc q) ! (4 * q )") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n !(4 * q + 0) = + findnth (Suc q) ! (4 * q + 0)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma dec_fetch_locate_b_o: + "\start_of ly as \ a; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + start_of ly as> 0\ + \ fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) + (Suc (a - start_of ly as)) Oc = (R, a)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append) +apply(subgoal_tac "\ q. (a - start_of ly as) = 2 * q + 1", auto) +apply(subgoal_tac "(findnth n ! Suc (Suc (Suc (4 * q)))) = + findnth (Suc q) ! (4 * q + 3)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n ! (4 * q + 3) = + findnth (Suc q) ! (4 * q + 3)", simp add: add3_Suc) +apply(rule_tac findnth_nth, auto) +done + +lemma dec_fetch_locate_b_b: + "\\ a < start_of ly as; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + start_of ly as > 0\ + \ fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) + (Suc (a - start_of ly as)) Bk = (R, a + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append) +apply(subgoal_tac "\ q. (a - start_of ly as) = 2 * q + 1", auto) +apply(subgoal_tac "(findnth n ! Suc ((Suc (4 * q)))) = + findnth (Suc q) ! (4 * q + 2)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n ! (4 * q + 2) = + findnth (Suc q) ! (4 * q + 2)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma dec_fetch_locate_n_a_o: + "start_of ly as > 0 \ fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (2 * n)) Oc + = (R, start_of ly as + 2*n + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + +lemma dec_fetch_locate_n_a_b: + "start_of ly as > 0 \ fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (2 * n)) Bk + = (W1, start_of ly as + 2*n)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + +lemma dec_fetch_locate_n_b_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Oc + = (R, start_of ly as + 2*n + 2)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + + +lemma dec_fetch_locate_n_b_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Bk + = (L, start_of ly as + 2*n + 13)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + +lemma dec_fetch_first_on_right_move_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (Suc (Suc (2 * n)))) Oc + = (R, start_of ly as + 2*n + 2)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + +lemma dec_fetch_first_on_right_move_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) + (Suc (Suc (Suc (2 * n)))) Bk + = (L, start_of ly as + 2*n + 3)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + +lemma [simp]: "fetch x (a + 2 * n) b = fetch x (2 * n + a) b" +thm arg_cong +apply(rule_tac x = "a + 2*n" and y = "2*n + a" in arg_cong, simp) +done + +lemma dec_fetch_first_after_clear_o: + "start_of ly as > 0 \ fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 4) Oc + = (W0, start_of ly as + 2*n + 3)" +apply(auto simp: ci.simps findnth.simps tshift.simps + tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_first_after_clear_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 4) Bk + = (R, start_of ly as + 2*n + 4)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 4= Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_right_move_b: + "start_of ly as > 0 \ fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 5) Bk + = (R, start_of ly as + 2*n + 5)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 5= Suc (2*n + 4)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_right_move_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 6) Bk + = (L, start_of ly as + 2*n + 6)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_right_move_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) (start_of ly as) + (Dec n e)) (2 * n + 6) Oc + = (L, start_of ly as + 2*n + 7)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_left_move_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 7) Bk + = (L, start_of ly as + 2*n + 10)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 7 = Suc (2*n + 6)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_write_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 8) Bk + = (W1, start_of ly as + 2*n + 7)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_write_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 8) Oc + = (R, start_of ly as + 2*n + 8)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_right_move_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 9) Bk + = (L, start_of ly as + 2*n + 9)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_right_move_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 9) Oc + = (R, start_of ly as + 2*n + 8)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_clear_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 10) Bk + = (R, start_of ly as + 2*n + 4)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 10 = Suc (2*n + 9)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_clear_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 10) Oc + = (W0, start_of ly as + 2*n + 9)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 10= Suc (2*n + 9)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_left_move1_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 11) Oc + = (L, start_of ly as + 2*n + 10)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 11= Suc (2*n + 10)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_left_move1_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 11) Bk + = (L, start_of ly as + 2*n + 11)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 11 = Suc (2*n + 10)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_left_move1_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 12) Oc + = (L, start_of ly as + 2*n + 10)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 12= Suc (2*n + 11)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_left_move1_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 12) Bk + = (R, start_of ly as + 2*n + 12)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 12 = Suc (2*n + 11)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_left_move1_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 13) Bk + = (R, start_of ly as + 2*n + 16)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 13 = Suc (2*n + 12)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_left_move2_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 14) Oc + = (L, start_of ly as + 2*n + 13)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 14 = Suc (2*n + 13)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_left_move2_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 14) Bk + = (L, start_of ly as + 2*n + 14)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 14 = Suc (2*n + 13)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_left_move2_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 15) Oc + = (L, start_of ly as + 2*n + 13)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 15 = Suc (2*n + 14)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_left_move2_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 15) Bk + = (R, start_of ly as + 2*n + 15)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 15= Suc (2*n + 14)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_left_move2_b: + "\ly = layout_of aprog; + abc_fetch as aprog = Some (Dec n e); + start_of ly as > 0\ \ + fetch (ci (layout_of aprog) (start_of ly as) + (Dec n e)) (2 * n + 16) Bk + = (R, start_of ly e)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 16 = Suc (2*n + 15)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_next_state: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2* n + 17) b + = (Nop, 0)" +apply(case_tac b) +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac [!] "2*n + 17 = Suc (2*n + 16)", + simp_all only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +(*End: dec_fetch lemmas*) +lemmas dec_fetch_simps = + dec_fetch_locate_a_o dec_fetch_locate_a_b dec_fetch_locate_b_o + dec_fetch_locate_b_b dec_fetch_locate_n_a_o + dec_fetch_locate_n_a_b dec_fetch_locate_n_b_o + dec_fetch_locate_n_b_b dec_fetch_first_on_right_move_o + dec_fetch_first_on_right_move_b dec_fetch_first_after_clear_b + dec_fetch_first_after_clear_o dec_fetch_right_move_b + dec_fetch_on_right_move_b dec_fetch_on_right_move_o + dec_fetch_after_clear_b dec_fetch_after_clear_o + dec_fetch_check_right_move_b dec_fetch_check_right_move_o + dec_fetch_left_move_b dec_fetch_on_left_move1_b + dec_fetch_on_left_move1_o dec_fetch_check_left_move1_b + dec_fetch_check_left_move1_o dec_fetch_after_left_move1_b + dec_fetch_on_left_move2_b dec_fetch_on_left_move2_o + dec_fetch_check_left_move2_o dec_fetch_check_left_move2_b + dec_fetch_after_left_move2_b dec_fetch_after_write_b + dec_fetch_after_write_o dec_fetch_next_state + +lemma [simp]: + "\start_of ly as \ a; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, Bk # xs) ires\ + \ \ Suc a < start_of ly as + 2 * n \ + inv_locate_a (as, am) (n, Bk # aaa, xs) ires" +apply(insert locate_b_2_locate_a[of a ly as n am aaa xs], simp) +done + +lemma [simp]: + "\start_of ly as \ a; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, []) ires\ + \ \ Suc a < start_of ly as + 2 * n \ + inv_locate_a (as, am) (n, Bk # aaa, []) ires" +apply(insert locate_b_2_locate_a_B[of a ly as n am aaa], simp) +done + +(* +lemma [simp]: "a\<^bsup>0\<^esup>=[]" +apply(simp add: exponent_def) +done +*) + +lemma exp_ind: "a\<^bsup>Suc b\<^esup> = a\<^bsup>b\<^esup> @ [a]" +apply(simp only: exponent_def rep_ind) +done + +lemma [simp]: + "inv_locate_b (as, am) (n, l, Oc # r) ires + \ dec_first_on_right_moving n (as, abc_lm_s am n (abc_lm_v am n)) + (Suc (Suc (start_of ly as + 2 * n)), Oc # l, r) ires" +apply(simp only: inv_locate_b.simps + dec_first_on_right_moving.simps in_middle.simps + abc_lm_s.simps abc_lm_v.simps) +apply(erule_tac exE)+ +apply(erule conjE)+ +apply(case_tac "n < length am", simp) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = "Suc ml" in exI, rule_tac conjI, rule_tac [1-2] impI) +prefer 3 +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(subgoal_tac "Suc n - length am = Suc (n - length am)", + simp only:exponent_def rep_ind, simp) +apply(rule_tac x = "Suc ml" in exI, simp_all) +apply(rule_tac [!] x = "mr - 1" in exI, simp_all) +apply(case_tac [!] mr, auto) +done + +lemma [simp]: + "\inv_locate_b (as, am) (n, l, r) ires; l \ []\ \ + \ inv_on_left_moving_in_middle_B (as, abc_lm_s am n (abc_lm_v am n)) + (s, tl l, hd l # r) ires" +apply(auto simp: inv_locate_b.simps + inv_on_left_moving_in_middle_B.simps in_middle.simps) +apply(case_tac [!] ml, auto split: if_splits) +done + +lemma [simp]: "inv_locate_b (as, am) (n, l, r) ires \ l \ []" +apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits) +done + +lemma [simp]: "\inv_locate_b (as, am) (n, l, Bk # r) ires; n < length am\ + \ inv_on_left_moving_norm (as, am) (s, tl l, hd l # Bk # r) ires" +apply(simp only: inv_locate_b.simps inv_on_left_moving_norm.simps + in_middle.simps) +apply(erule_tac exE)+ +apply(erule_tac conjE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = "ml - 1" in exI, auto) +apply(rule_tac [!] x = "Suc mr" in exI) +apply(case_tac [!] mr, auto) +done + +lemma [simp]: "\inv_locate_b (as, am) (n, l, Bk # r) ires; \ n < length am\ + \ inv_on_left_moving_norm (as, am @ + replicate (n - length am) 0 @ [0]) (s, tl l, hd l # Bk # r) ires" +apply(simp only: inv_locate_b.simps inv_on_left_moving_norm.simps + in_middle.simps) +apply(erule_tac exE)+ +apply(erule_tac conjE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(subgoal_tac "Suc n - length am = Suc (n - length am)", simp only: rep_ind exponent_def, simp_all) +apply(rule_tac x = "Suc mr" in exI, auto) +done + +lemma inv_locate_b_2_on_left_moving[simp]: + "\inv_locate_b (as, am) (n, l, Bk # r) ires\ + \ (l = [] \ inv_on_left_moving (as, + abc_lm_s am n (abc_lm_v am n)) (s, [], Bk # Bk # r) ires) \ + (l \ [] \ inv_on_left_moving (as, + abc_lm_s am n (abc_lm_v am n)) (s, tl l, hd l # Bk # r) ires)" +apply(subgoal_tac "l\[]") +apply(subgoal_tac "\ inv_on_left_moving_in_middle_B + (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, hd l # Bk # r) ires") +apply(simp add:inv_on_left_moving.simps + abc_lm_s.simps abc_lm_v.simps split: if_splits, auto) +done + +lemma [simp]: + "inv_locate_b (as, am) (n, l, []) ires \ + inv_locate_b (as, am) (n, l, [Bk]) ires" +apply(auto simp: inv_locate_b.simps in_middle.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, + rule_tac x = "Suc (length lm1) - length am" in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = ml in exI, rule_tac x = mr in exI) +apply(auto) +done + +lemma nil_2_nil: " = [] \ lm = []" +apply(auto simp: tape_of_nl_abv) +apply(case_tac lm, simp) +apply(case_tac list, auto simp: tape_of_nat_list.simps) +done + +lemma inv_locate_b_2_on_left_moving_b[simp]: + "inv_locate_b (as, am) (n, l, []) ires + \ (l = [] \ inv_on_left_moving (as, + abc_lm_s am n (abc_lm_v am n)) (s, [], [Bk]) ires) \ + (l \ [] \ inv_on_left_moving (as, abc_lm_s am n + (abc_lm_v am n)) (s, tl l, [hd l]) ires)" +apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s]) +apply(simp only: inv_on_left_moving.simps, simp) +apply(subgoal_tac "\ inv_on_left_moving_in_middle_B + (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp) +apply(simp only: inv_on_left_moving_norm.simps) +apply(erule_tac exE)+ +apply(erule_tac conjE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, rule_tac x = ml in exI, + rule_tac x = mr in exI, simp) +apply(case_tac mr, simp, simp, case_tac nat, auto intro: nil_2_nil) +done + +lemma [simp]: + "\dec_first_on_right_moving n (as, am) (s, aaa, Oc # xs) ires\ + \ dec_first_on_right_moving n (as, am) (s', Oc # aaa, xs) ires" +apply(simp only: dec_first_on_right_moving.simps) +apply(erule exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = "Suc ml" in exI, + rule_tac x = "mr - 1" in exI, auto) +apply(case_tac [!] mr, auto) +done + +lemma [simp]: + "dec_first_on_right_moving n (as, am) (s, l, Bk # xs) ires \ l \ []" +apply(auto simp: dec_first_on_right_moving.simps split: if_splits) +done + +lemma [elim]: + "\\ length lm1 < length am; + am @ replicate (length lm1 - length am) 0 @ [0::nat] = + lm1 @ m # lm2; + 0 < m\ + \ RR" +apply(subgoal_tac "lm2 = []", simp) +apply(drule_tac length_equal, simp) +done + +lemma [simp]: + "\dec_first_on_right_moving n (as, + abc_lm_s am n (abc_lm_v am n)) (s, l, Bk # xs) ires\ +\ dec_after_clear (as, abc_lm_s am n + (abc_lm_v am n - Suc 0)) (s', tl l, hd l # Bk # xs) ires" +apply(simp only: dec_first_on_right_moving.simps + dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps) +apply(erule_tac exE)+ +apply(case_tac "n < length am") +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = "m - 1" in exI, auto simp: ) +apply(case_tac [!] mr, auto) +done + +lemma [simp]: + "\dec_first_on_right_moving n (as, + abc_lm_s am n (abc_lm_v am n)) (s, l, []) ires\ +\ (l = [] \ dec_after_clear (as, + abc_lm_s am n (abc_lm_v am n - Suc 0)) (s', [], [Bk]) ires) \ + (l \ [] \ dec_after_clear (as, abc_lm_s am n + (abc_lm_v am n - Suc 0)) (s', tl l, [hd l]) ires)" +apply(subgoal_tac "l \ []", + simp only: dec_first_on_right_moving.simps + dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps) +apply(erule_tac exE)+ +apply(case_tac "n < length am", simp) +apply(rule_tac x = lm1 in exI, rule_tac x = "m - 1" in exI, auto) +apply(case_tac [1-2] mr, auto) +apply(case_tac [1-2] m, auto simp: dec_first_on_right_moving.simps split: if_splits) +done + +lemma [simp]: "\dec_after_clear (as, am) (s, l, Oc # r) ires\ + \ dec_after_clear (as, am) (s', l, Bk # r) ires" +apply(auto simp: dec_after_clear.simps) +done + +lemma [simp]: "\dec_after_clear (as, am) (s, l, Bk # r) ires\ + \ dec_right_move (as, am) (s', Bk # l, r) ires" +apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits) +done + +lemma [simp]: "\dec_after_clear (as, am) (s, l, []) ires\ + \ dec_right_move (as, am) (s', Bk # l, []) ires" +apply(auto simp: dec_after_clear.simps dec_right_move.simps ) +done + +lemma [simp]: "\rn. a::block\<^bsup>rn\<^esup> = []" +apply(rule_tac x = 0 in exI, simp) +done + +lemma [simp]: "\dec_after_clear (as, am) (s, l, []) ires\ + \ dec_right_move (as, am) (s', Bk # l, [Bk]) ires" +apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits) +done + +lemma [simp]:"dec_right_move (as, am) (s, l, Oc # r) ires = False" +apply(auto simp: dec_right_move.simps) +done + +lemma dec_right_move_2_check_right_move[simp]: + "\dec_right_move (as, am) (s, l, Bk # r) ires\ + \ dec_check_right_move (as, am) (s', Bk # l, r) ires" +apply(auto simp: dec_right_move.simps dec_check_right_move.simps split: if_splits) +done + +lemma [simp]: + "dec_right_move (as, am) (s, l, []) ires= + dec_right_move (as, am) (s, l, [Bk]) ires" +apply(simp add: dec_right_move.simps) +apply(rule_tac iffI) +apply(erule_tac [!] exE)+ +apply(erule_tac [2] exE) +apply(rule_tac [!] x = lm1 in exI, rule_tac x = "[]" in exI, + rule_tac [!] x = m in exI, auto) +apply(auto intro: nil_2_nil) +done + +lemma [simp]: "\dec_right_move (as, am) (s, l, []) ires\ + \ dec_check_right_move (as, am) (s, Bk # l, []) ires" +apply(insert dec_right_move_2_check_right_move[of as am s l "[]" s'], + simp) +done + +lemma [simp]: "dec_check_right_move (as, am) (s, l, r) ires\ l \ []" +apply(auto simp: dec_check_right_move.simps split: if_splits) +done + +lemma [simp]: "\dec_check_right_move (as, am) (s, l, Oc # r) ires\ + \ dec_after_write (as, am) (s', tl l, hd l # Oc # r) ires" +apply(auto simp: dec_check_right_move.simps dec_after_write.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, auto) +done + +lemma [simp]: "\dec_check_right_move (as, am) (s, l, Bk # r) ires\ + \ dec_left_move (as, am) (s', tl l, hd l # Bk # r) ires" +apply(auto simp: dec_check_right_move.simps + dec_left_move.simps inv_after_move.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto) +apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil) +apply(rule_tac x = "Suc rn" in exI) +apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil) +done + +lemma [simp]: "\dec_check_right_move (as, am) (s, l, []) ires\ + \ dec_left_move (as, am) (s', tl l, [hd l]) ires" +apply(auto simp: dec_check_right_move.simps + dec_left_move.simps inv_after_move.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto) +apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil) +done + +lemma [simp]: "dec_left_move (as, am) (s, aaa, Oc # xs) ires = False" +apply(auto simp: dec_left_move.simps inv_after_move.simps) +apply(case_tac [!] rn, auto) +done + +lemma [simp]: "dec_left_move (as, am) (s, l, r) ires + \ l \ []" +apply(auto simp: dec_left_move.simps split: if_splits) +done + +lemma tape_of_nl_abv_cons_ex[simp]: + "\lna. Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk\<^bsup>ln\<^esup> = @ Bk\<^bsup>lna\<^esup>" +apply(case_tac "lm1=[]", auto simp: tape_of_nl_abv + tape_of_nat_list.simps) +apply(rule_tac x = "ln" in exI, simp) +apply(simp add: tape_of_nat_list_cons exponent_def) +done + +(* +lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm1 @ [m]) + (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk\<^bsup>ln\<^esup>, Bk # Bk\<^bsup>rn\<^esup>) ires" +apply(simp only: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "[]" in exI, auto) +done +*) +lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m]) + (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # ires, Bk # Bk\<^bsup>rn\<^esup>) ires" +apply(simp add: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "[m]" in exI, simp, auto simp: tape_of_nat_def) +done + +lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m]) + (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # ires, [Bk]) ires" +apply(simp add: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "[m]" in exI, simp, auto simp: tape_of_nat_def) +done + +lemma [simp]: "lm1 \ [] \ + inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', + Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk # Bk # ires, Bk # Bk\<^bsup>rn\<^esup>) ires" +apply(simp only: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp, auto) +done + +lemma [simp]: "lm1 \ [] \ + inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', + Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk # Bk # ires, [Bk]) ires" +apply(simp only: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp, auto) +done + +lemma [simp]: "dec_left_move (as, am) (s, l, Bk # r) ires + \ inv_on_left_moving (as, am) (s', tl l, hd l # Bk # r) ires" +apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits) +done + +(* +lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm1 @ [m]) + (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk\<^bsup>ln\<^esup>, [Bk]) ires" +apply(auto simp: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "[]" in exI, auto) +done +*) + +lemma [simp]: "dec_left_move (as, am) (s, l, []) ires + \ inv_on_left_moving (as, am) (s', tl l, [hd l]) ires" +apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits) +done + +lemma [simp]: "dec_after_write (as, am) (s, l, Oc # r) ires + \ dec_on_right_moving (as, am) (s', Oc # l, r) ires" +apply(auto simp: dec_after_write.simps dec_on_right_moving.simps) +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, + rule_tac x = "hd lm2" in exI, simp) +apply(rule_tac x = "Suc 0" in exI,rule_tac x = "Suc (hd lm2)" in exI) +apply(case_tac lm2, simp, simp) +apply(case_tac "list = []", + auto simp: tape_of_nl_abv tape_of_nat_list.simps split: if_splits ) +apply(case_tac rn, auto) +apply(case_tac "rev lm1", simp, simp add: tape_of_nat_list.simps) +apply(case_tac rn, auto) +apply(case_tac list, simp_all add: tape_of_nat_list.simps, auto) +apply(case_tac "rev lm1", simp, simp add: tape_of_nat_list.simps) +apply(case_tac list, simp_all add: tape_of_nat_list.simps, auto) +done + +lemma [simp]: "dec_after_write (as, am) (s, l, Bk # r) ires + \ dec_after_write (as, am) (s', l, Oc # r) ires" +apply(auto simp: dec_after_write.simps) +done + +lemma [simp]: "dec_after_write (as, am) (s, aaa, []) ires + \ dec_after_write (as, am) (s', aaa, [Oc]) ires" +apply(auto simp: dec_after_write.simps) +done + +lemma [simp]: "dec_on_right_moving (as, am) (s, l, Oc # r) ires + \ dec_on_right_moving (as, am) (s', Oc # l, r) ires" +apply(simp only: dec_on_right_moving.simps) +apply(erule_tac exE)+ +apply(erule conjE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = "m" in exI, rule_tac x = "Suc ml" in exI, + rule_tac x = "mr - 1" in exI, simp) +apply(case_tac mr, auto) +done + +lemma [simp]: "dec_on_right_moving (as, am) (s, l, r) ires\ l \ []" +apply(auto simp: dec_on_right_moving.simps split: if_splits) +done + +lemma [simp]: "dec_on_right_moving (as, am) (s, l, Bk # r) ires + \ dec_after_clear (as, am) (s', tl l, hd l # Bk # r) ires" +apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps) +apply(case_tac [!] mr, auto split: if_splits) +done + +lemma [simp]: "dec_on_right_moving (as, am) (s, l, []) ires + \ dec_after_clear (as, am) (s', tl l, [hd l]) ires" +apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps) +apply(case_tac mr, simp_all split: if_splits) +apply(rule_tac x = lm1 in exI, simp) +done + +lemma start_of_le: "a < b \ start_of ly a \ start_of ly b" +proof(induct b arbitrary: a, simp, case_tac "a = b", simp) + fix b a + show "start_of ly b \ start_of ly (Suc b)" + apply(case_tac "b::nat", + simp add: start_of.simps, simp add: start_of.simps) + done +next + fix b a + assume h1: "\a. a < b \ start_of ly a \ start_of ly b" + "a < Suc b" "a \ b" + hence "a < b" + by(simp) + from h1 and this have h2: "start_of ly a \ start_of ly b" + by(drule_tac h1, simp) + from h2 show "start_of ly a \ start_of ly (Suc b)" + proof - + have "start_of ly b \ start_of ly (Suc b)" + apply(case_tac "b::nat", + simp add: start_of.simps, simp add: start_of.simps) + done + from h2 and this show "start_of ly a \ start_of ly (Suc b)" + by simp + qed +qed + +lemma start_of_dec_length[simp]: + "\abc_fetch a aprog = Some (Dec n e)\ \ + start_of (layout_of aprog) (Suc a) + = start_of (layout_of aprog) a + 2*n + 16" +apply(case_tac a, auto simp: abc_fetch.simps start_of.simps + layout_of.simps length_of.simps + split: if_splits) +done + +lemma start_of_ge: + "\abc_fetch a aprog = Some (Dec n e); a < e\ \ + start_of (layout_of aprog) e > + start_of (layout_of aprog) a + 2*n + 15" +apply(case_tac "e = Suc a", + simp add: start_of.simps abc_fetch.simps layout_of.simps + length_of.simps split: if_splits) +apply(subgoal_tac "Suc a < e", drule_tac a = "Suc a" + and ly = "layout_of aprog" in start_of_le) +apply(subgoal_tac "start_of (layout_of aprog) (Suc a) + = start_of (layout_of aprog) a + 2*n + 16", simp) +apply(rule_tac start_of_dec_length, simp) +apply(arith) +done + +lemma starte_not_equal[simp]: + "\abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\ + \ (start_of ly e \ Suc (Suc (start_of ly as + 2 * n)) \ + start_of ly e \ start_of ly as + 2 * n + 3 \ + start_of ly e \ start_of ly as + 2 * n + 4 \ + start_of ly e \ start_of ly as + 2 * n + 5 \ + start_of ly e \ start_of ly as + 2 * n + 6 \ + start_of ly e \ start_of ly as + 2 * n + 7 \ + start_of ly e \ start_of ly as + 2 * n + 8 \ + start_of ly e \ start_of ly as + 2 * n + 9 \ + start_of ly e \ start_of ly as + 2 * n + 10 \ + start_of ly e \ start_of ly as + 2 * n + 11 \ + start_of ly e \ start_of ly as + 2 * n + 12 \ + start_of ly e \ start_of ly as + 2 * n + 13 \ + start_of ly e \ start_of ly as + 2 * n + 14 \ + start_of ly e \ start_of ly as + 2 * n + 15)" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp) +apply(drule_tac a = as and e = e in start_of_ge, simp, simp) +done + +lemma [simp]: "\abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\ + \ (Suc (Suc (start_of ly as + 2 * n)) \ start_of ly e \ + start_of ly as + 2 * n + 3 \ start_of ly e \ + start_of ly as + 2 * n + 4 \ start_of ly e \ + start_of ly as + 2 * n + 5 \start_of ly e \ + start_of ly as + 2 * n + 6 \ start_of ly e \ + start_of ly as + 2 * n + 7 \ start_of ly e \ + start_of ly as + 2 * n + 8 \ start_of ly e \ + start_of ly as + 2 * n + 9 \ start_of ly e \ + start_of ly as + 2 * n + 10 \ start_of ly e \ + start_of ly as + 2 * n + 11 \ start_of ly e \ + start_of ly as + 2 * n + 12 \ start_of ly e \ + start_of ly as + 2 * n + 13 \ start_of ly e \ + start_of ly as + 2 * n + 14 \ start_of ly e \ + start_of ly as + 2 * n + 15 \ start_of ly e)" +apply(insert starte_not_equal[of as aprog n e ly], + simp del: starte_not_equal) +apply(erule_tac conjE)+ +apply(rule_tac conjI, simp del: starte_not_equal)+ +apply(rule not_sym, simp) +done + +lemma [simp]: "start_of (layout_of aprog) as > 0 \ + fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n as)) (Suc 0) Oc = + (R, Suc (start_of (layout_of aprog) as))" + +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append + Suc_pre tdec_b_def) +apply(insert findnth_nth[of 0 n "Suc 0"], simp) +apply(simp add: findnth.simps) +done + +lemma start_of_inj[simp]: + "\abc_fetch as aprog = Some (Dec n e); e \ as; ly = layout_of aprog\ + \ start_of ly as \ start_of ly e" +apply(case_tac "e < as") +apply(case_tac "as", simp, simp) +apply(case_tac "e = nat", simp add: start_of.simps + layout_of.simps length_of.simps) +apply(subgoal_tac "e < length aprog", simp add: length_of.simps + split: abc_inst.splits) +apply(simp add: abc_fetch.simps split: if_splits) +apply(subgoal_tac "e < nat", drule_tac a = e and b = nat + and ly =ly in start_of_le, simp) +apply(subgoal_tac "start_of ly nat < start_of ly (Suc nat)", + simp, simp add: start_of.simps layout_of.simps) +apply(subgoal_tac "nat < length aprog", simp) +apply(case_tac "aprog ! nat", auto simp: length_of.simps) +apply(simp add: abc_fetch.simps split: if_splits) +apply(subgoal_tac "e > as", drule_tac start_of_ge, auto) +done + +lemma [simp]: "\abc_fetch as aprog = Some (Dec n e); e < as\ + \ Suc (start_of (layout_of aprog) e) - + start_of (layout_of aprog) as = 0" +apply(frule_tac ly = "layout_of aprog" in start_of_le, simp) +apply(subgoal_tac "start_of (layout_of aprog) as \ + start_of (layout_of aprog) e", arith) +apply(rule start_of_inj, auto) +done + +lemma [simp]: + "\abc_fetch as aprog = Some (Dec n e); + 0 < start_of (layout_of aprog) as\ + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) (Suc (start_of (layout_of aprog) e) - + start_of (layout_of aprog) as) Oc) + = (if e = as then (R, start_of (layout_of aprog) as + 1) + else (Nop, 0))" +apply(auto split: if_splits) +apply(case_tac "e < as", simp add: fetch.simps) +apply(subgoal_tac " e > as") +apply(drule start_of_ge, simp, + auto simp: fetch.simps ci_length nth_of.simps) +apply(subgoal_tac + "length (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) div 2= length_of (Dec n e)") +defer +apply(simp add: ci_length) +apply(subgoal_tac + "length (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) mod 2 = 0", auto simp: length_of.simps) +done + +lemma [simp]: + "start_of (layout_of aprog) as > 0 \ + fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n as)) (Suc 0) Bk + = (W1, start_of (layout_of aprog) as)" +apply(auto simp: ci.simps findnth.simps fetch.simps nth_of.simps + tshift.simps nth_append Suc_pre tdec_b_def) +apply(insert findnth_nth[of 0 n "0"], simp) +apply(simp add: findnth.simps) +done + +lemma [simp]: + "\abc_fetch as aprog = Some (Dec n e); + 0 < start_of (layout_of aprog) as\ +\ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) (Suc (start_of (layout_of aprog) e) - + start_of (layout_of aprog) as) Bk) + = (if e = as then (W1, start_of (layout_of aprog) as) + else (Nop, 0))" +apply(auto split: if_splits) +apply(case_tac "e < as", simp add: fetch.simps) +apply(subgoal_tac " e > as") +apply(drule start_of_ge, simp, auto simp: fetch.simps + ci_length nth_of.simps) +apply(subgoal_tac + "length (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) div 2= length_of (Dec n e)") +defer +apply(simp add: ci_length) +apply(subgoal_tac + "length (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) mod 2 = 0", auto simp: length_of.simps) +apply(simp add: ci.simps tshift.simps tdec_b_def) +done + +lemma [simp]: + "inv_stop (as, abc_lm_s am n (abc_lm_v am n)) (s, l, r) ires \ l \ []" +apply(auto simp: inv_stop.simps) +done + +lemma [simp]: + "\abc_fetch as aprog = Some (Dec n e); e \ as; ly = layout_of aprog\ + \ (\ (start_of ly as \ start_of ly e \ + start_of ly e < start_of ly as + 2 * n)) + \ start_of ly e \ start_of ly as + 2*n \ + start_of ly e \ Suc (start_of ly as + 2*n) " +apply(case_tac "e < as") +apply(drule_tac ly = ly in start_of_le, simp) +apply(case_tac n, simp, drule start_of_inj, simp, simp, simp, simp) +apply(drule_tac start_of_ge, simp, simp) +done + +lemma [simp]: + "\abc_fetch as aprog = Some (Dec n e); start_of ly as \ s; + s < start_of ly as + 2 * n; ly = layout_of aprog\ + \ Suc s \ start_of ly e " +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp) +apply(drule_tac start_of_ge, auto) +done + +lemma [simp]: "\abc_fetch as aprog = Some (Dec n e); + ly = layout_of aprog\ + \ Suc (start_of ly as + 2 * n) \ start_of ly e" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp) +apply(drule_tac start_of_ge, auto) +done + +lemma dec_false_1[simp]: + "\abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, Oc # xs) ires\ + \ False" +apply(auto simp: inv_locate_b.simps in_middle.simps exponent_def) +apply(case_tac "length lm1 \ length am", auto) +apply(subgoal_tac "lm2 = []", simp, subgoal_tac "m = 0", simp) +apply(case_tac mr, auto simp: ) +apply(subgoal_tac "Suc (length lm1) - length am = + Suc (length lm1 - length am)", + simp add: rep_ind del: replicate.simps, simp) +apply(drule_tac xs = "am @ replicate (Suc (length lm1) - length am) 0" + and ys = "lm1 @ m # lm2" in length_equal, simp) +apply(case_tac mr, auto simp: abc_lm_v.simps) +apply(case_tac "mr = 0", simp_all add: exponent_def split: if_splits) +apply(subgoal_tac "Suc (length lm1) - length am = + Suc (length lm1 - length am)", + simp add: rep_ind del: replicate.simps, simp) +done + +lemma [simp]: + "\inv_locate_b (as, am) (n, aaa, Bk # xs) ires; + abc_lm_v am n = 0\ + \ inv_on_left_moving (as, abc_lm_s am n 0) + (s, tl aaa, hd aaa # Bk # xs) ires" +apply(insert inv_locate_b_2_on_left_moving[of as am n aaa xs ires s], simp) +done + +lemma [simp]: + "\abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, []) ires\ + \ inv_on_left_moving (as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires" +apply(insert inv_locate_b_2_on_left_moving_b[of as am n aaa ires s], simp) +done + +lemma [simp]: "\am ! n = (0::nat); n < length am\ \ am[n := 0] = am" +apply(simp add: list_update_same_conv) +done + +lemma [simp]: "\abc_lm_v am n = 0; + inv_locate_b (as, abc_lm_s am n 0) (n, Oc # aaa, xs) ires\ + \ inv_locate_b (as, am) (n, Oc # aaa, xs) ires" +apply(simp only: inv_locate_b.simps in_middle.simps abc_lm_s.simps + abc_lm_v.simps) +apply(erule_tac exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, simp) +apply(case_tac "n < length am", simp_all) +apply(erule_tac conjE)+ +apply(rule_tac x = tn in exI, rule_tac x = m in exI, simp) +apply(rule_tac x = ml in exI, rule_tac x = mr in exI, simp) +defer +apply(rule_tac x = "Suc n - length am" in exI, rule_tac x = m in exI) +apply(subgoal_tac "Suc n - length am = Suc (n - length am)") +apply(simp add: exponent_def rep_ind del: replicate.simps, auto) +done + +lemma [intro]: "\abc_lm_v (a # list) 0 = 0\ \ a = 0" +apply(simp add: abc_lm_v.simps split: if_splits) +done + +lemma [simp]: + "inv_stop (as, abc_lm_s am n 0) + (start_of (layout_of aprog) e, aaa, Oc # xs) ires + \ inv_locate_a (as, abc_lm_s am n 0) (0, aaa, Oc # xs) ires" +apply(simp add: inv_locate_a.simps) +apply(rule disjI1) +apply(auto simp: inv_stop.simps at_begin_norm.simps) +done + +lemma [simp]: + "\abc_lm_v am 0 = 0; + inv_stop (as, abc_lm_s am 0 0) + (start_of (layout_of aprog) e, aaa, Oc # xs) ires\ \ + inv_locate_b (as, am) (0, Oc # aaa, xs) ires" +apply(auto simp: inv_stop.simps inv_locate_b.simps + in_middle.simps abc_lm_s.simps) +apply(case_tac "am = []", simp) +apply(rule_tac x = "[]" in exI, rule_tac x = "Suc 0" in exI, + rule_tac x = 0 in exI, simp) +apply(rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI, + simp add: tape_of_nl_abv tape_of_nat_list.simps, auto) +apply(case_tac rn, auto) +apply(rule_tac x = "tl am" in exI, rule_tac x = 0 in exI, + rule_tac x = "hd am" in exI, simp) +apply(rule_tac x = "Suc 0" in exI, rule_tac x = "hd am" in exI, simp) +apply(case_tac am, simp, simp) +apply(subgoal_tac "a = 0", case_tac list, + auto simp: tape_of_nat_list.simps tape_of_nl_abv) +apply(case_tac rn, auto) +done + +lemma [simp]: + "\inv_stop (as, abc_lm_s am n 0) + (start_of (layout_of aprog) e, aaa, Oc # xs) ires\ + \ inv_locate_b (as, am) (0, Oc # aaa, xs) ires \ + inv_locate_b (as, abc_lm_s am n 0) (0, Oc # aaa, xs) ires" +apply(simp) +done + +lemma [simp]: +"\abc_lm_v am n = 0; + inv_stop (as, abc_lm_s am n 0) + (start_of (layout_of aprog) e, aaa, Oc # xs) ires\ + \ \ Suc 0 < 2 * n \ e = as \ + inv_locate_b (as, am) (n, Oc # aaa, xs) ires" +apply(case_tac n, simp, simp) +done + +lemma dec_false2: + "inv_stop (as, abc_lm_s am n 0) + (start_of (layout_of aprog) e, aaa, Bk # xs) ires = False" +apply(auto simp: inv_stop.simps abc_lm_s.simps) +apply(case_tac "am", simp, case_tac n, simp add: tape_of_nl_abv) +apply(case_tac list, simp add: tape_of_nat_list.simps ) +apply(simp add: tape_of_nat_list.simps , simp) +apply(case_tac "list[nat := 0]", + simp add: tape_of_nat_list.simps tape_of_nl_abv) +apply(simp add: tape_of_nat_list.simps ) +apply(case_tac "am @ replicate (n - length am) 0 @ [0]", simp) +apply(case_tac list, auto simp: tape_of_nl_abv + tape_of_nat_list.simps ) +done + +lemma dec_false3: + "inv_stop (as, abc_lm_s am n 0) + (start_of (layout_of aprog) e, aaa, []) ires = False" +apply(auto simp: inv_stop.simps abc_lm_s.simps) +apply(case_tac "am", case_tac n, auto) +apply(case_tac n, auto simp: tape_of_nl_abv) +apply(case_tac "list::nat list", + simp add: tape_of_nat_list.simps tape_of_nat_list.simps) +apply(simp add: tape_of_nat_list.simps) +apply(case_tac "list[nat := 0]", + simp add: tape_of_nat_list.simps tape_of_nat_list.simps) +apply(simp add: tape_of_nat_list.simps) +apply(case_tac "(am @ replicate (n - length am) 0 @ [0])", simp) +apply(case_tac list, auto simp: tape_of_nat_list.simps) +done + +lemma [simp]: + "fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e)) 0 b = (Nop, 0)" +by(simp add: fetch.simps) + +declare dec_inv_1.simps[simp del] + +declare inv_locate_n_b.simps [simp del] + +lemma [simp]: +"\0 < abc_lm_v am n; 0 < n; + at_begin_norm (as, am) (n, aaa, Oc # xs) ires\ + \ inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires" +apply(simp only: at_begin_norm.simps inv_locate_n_b.simps) +apply(erule_tac exE)+ +apply(rule_tac x = lm1 in exI, simp) +apply(case_tac "length lm2", simp) +apply(case_tac rn, simp, simp) +apply(rule_tac x = "tl lm2" in exI, rule_tac x = "hd lm2" in exI, simp) +apply(rule conjI) +apply(case_tac "lm2", simp, simp) +apply(case_tac "lm2", auto simp: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac [!] "list", auto simp: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac rn, auto) +done +lemma [simp]: "(\rn. Oc # xs = Bk\<^bsup>rn\<^esup>) = False" +apply(auto) +apply(case_tac rn, auto simp: ) +done + +lemma [simp]: + "\0 < abc_lm_v am n; 0 < n; + at_begin_fst_bwtn (as, am) (n, aaa, Oc # xs) ires\ + \ inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires" +apply(simp add: at_begin_fst_bwtn.simps inv_locate_n_b.simps ) +done + +lemma Suc_minus:"length am + tn = n + \ Suc tn = Suc n - length am " +apply(arith) +done + +lemma [simp]: + "\0 < abc_lm_v am n; 0 < n; + at_begin_fst_awtn (as, am) (n, aaa, Oc # xs) ires\ + \ inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires" +apply(simp only: at_begin_fst_awtn.simps inv_locate_n_b.simps ) +apply(erule exE)+ +apply(erule conjE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, + rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI) +apply(simp add: exponent_def rep_ind del: replicate.simps) +apply(rule conjI)+ +apply(auto) +apply(case_tac [!] rn, auto) +done + +lemma [simp]: + "\0 < abc_lm_v am n; 0 < n; inv_locate_a (as, am) (n, aaa, Oc # xs) ires\ + \ inv_locate_n_b (as, am) (n, Oc#aaa, xs) ires" +apply(auto simp: inv_locate_a.simps) +done + +lemma [simp]: + "\inv_locate_n_b (as, am) (n, aaa, Oc # xs) ires\ + \ dec_first_on_right_moving n (as, abc_lm_s am n (abc_lm_v am n)) + (s, Oc # aaa, xs) ires" +apply(auto simp: inv_locate_n_b.simps dec_first_on_right_moving.simps + abc_lm_s.simps abc_lm_v.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = "Suc (Suc 0)" in exI, + rule_tac x = "m - 1" in exI, simp) +apply(case_tac m, auto simp: exponent_def) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, + simp add: Suc_diff_le rep_ind del: replicate.simps) +apply(rule_tac x = "Suc (Suc 0)" in exI, + rule_tac x = "m - 1" in exI, simp) +apply(case_tac m, auto simp: exponent_def) +apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = "Suc (Suc 0)" in exI, + rule_tac x = "m - 1" in exI, simp) +apply(case_tac m, auto) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, + simp add: Suc_diff_le rep_ind del: replicate.simps, simp) +done + +lemma dec_false_2: + "\0 < abc_lm_v am n; inv_locate_n_b (as, am) (n, aaa, Bk # xs) ires\ + \ False" +apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits) +apply(case_tac [!] m, auto) +done + +lemma dec_false_2_b: + "\0 < abc_lm_v am n; inv_locate_n_b (as, am) + (n, aaa, []) ires\ \ False" +apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits) +apply(case_tac [!] m, auto simp: ) +done + + +(*begin: dec halt1 lemmas*) +thm abc_inc_stage1.simps +fun abc_dec_1_stage1:: "t_conf \ nat \ nat \ nat" + where + "abc_dec_1_stage1 (s, l, r) ss n = + (if s > ss \ s \ ss + 2*n + 1 then 4 + else if s = ss + 2 * n + 13 \ s = ss + 2*n + 14 then 3 + else if s = ss + 2*n + 15 then 2 + else 0)" + +fun abc_dec_1_stage2:: "t_conf \ nat \ nat \ nat" + where + "abc_dec_1_stage2 (s, l, r) ss n = + (if s \ ss + 2 * n + 1 then (ss + 2 * n + 16 - s) + else if s = ss + 2*n + 13 then length l + else if s = ss + 2*n + 14 then length l + else 0)" + +fun abc_dec_1_stage3 :: "t_conf \ nat \ nat \ block list \ nat" + where + "abc_dec_1_stage3 (s, l, r) ss n ires = + (if s \ ss + 2*n + 1 then + if (s - ss) mod 2 = 0 then + if r \ [] \ hd r = Oc then 0 else 1 + else length r + else if s = ss + 2 * n + 13 then + if l = Bk # ires \ r \ [] \ hd r = Oc then 2 + else 1 + else if s = ss + 2 * n + 14 then + if r \ [] \ hd r = Oc then 3 else 0 + else 0)" + +fun abc_dec_1_measure :: "(t_conf \ nat \ nat \ block list) \ (nat \ nat \ nat)" + where + "abc_dec_1_measure (c, ss, n, ires) = (abc_dec_1_stage1 c ss n, + abc_dec_1_stage2 c ss n, abc_dec_1_stage3 c ss n ires)" + +definition abc_dec_1_LE :: + "(((nat \ block list \ block list) \ nat \ + nat \ block list) \ ((nat \ block list \ block list) \ nat \ nat \ block list)) set" + where "abc_dec_1_LE \ (inv_image lex_triple abc_dec_1_measure)" + +lemma wf_dec_le: "wf abc_dec_1_LE" +by(auto intro:wf_inv_image wf_lex_triple simp:abc_dec_1_LE_def) + +declare dec_inv_1.simps[simp del] dec_inv_2.simps[simp del] + +lemma [elim]: + "\abc_fetch as aprog = Some (Dec n e); + start_of (layout_of aprog) as < start_of (layout_of aprog) e; + start_of (layout_of aprog) e \ + Suc (start_of (layout_of aprog) as + 2 * n)\ \ False" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = "layout_of aprog" in + start_of_le, simp) +apply(drule_tac start_of_ge, auto) +done + +lemma [elim]: "\abc_fetch as aprog = Some (Dec n e); + start_of (layout_of aprog) e + = start_of (layout_of aprog) as + 2 * n + 13\ + \ False" +apply(insert starte_not_equal[of as aprog n e "layout_of aprog"], + simp) +done + +lemma [elim]: "\abc_fetch as aprog = Some (Dec n e); + start_of (layout_of aprog) e = + start_of (layout_of aprog) as + 2 * n + 14\ + \ False" +apply(insert starte_not_equal[of as aprog n e "layout_of aprog"], + simp) +done + +lemma [elim]: + "\abc_fetch as aprog = Some (Dec n e); + start_of (layout_of aprog) as < start_of (layout_of aprog) e; + start_of (layout_of aprog) e \ + Suc (start_of (layout_of aprog) as + 2 * n)\ + \ False" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = "layout_of aprog" in + start_of_le, simp) +apply(drule_tac start_of_ge, auto) +done + +lemma [elim]: + "\abc_fetch as aprog = Some (Dec n e); + start_of (layout_of aprog) e = + start_of (layout_of aprog) as + 2 * n + 13\ + \ False" +apply(insert starte_not_equal[of as aprog n e "layout_of aprog"], + simp) +done + +lemma [simp]: + "abc_fetch as aprog = Some (Dec n e) \ + Suc (start_of (layout_of aprog) as) \ start_of (layout_of aprog) e" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = "(layout_of aprog)" in + start_of_le, simp) +apply(drule_tac a = as and e = e in start_of_ge, simp, simp) +done + +lemma [simp]: "inv_on_left_moving (as, am) (s, [], r) ires + = False" +apply(simp add: inv_on_left_moving.simps inv_on_left_moving_norm.simps + inv_on_left_moving_in_middle_B.simps) +done + +lemma [simp]: + "inv_check_left_moving (as, abc_lm_s am n 0) + (start_of (layout_of aprog) as + 2 * n + 14, [], Oc # xs) ires + = False" +apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps) +done + +lemma dec_inv_stop1_pre: + "\abc_fetch as aprog = Some (Dec n e); abc_lm_v am n = 0; + start_of (layout_of aprog) as > 0\ + \ \na. \ (\(s, l, r) (ss, n', ires'). s = start_of (layout_of aprog) e) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) na) + (start_of (layout_of aprog) as, n, ires) \ + dec_inv_1 (layout_of aprog) n e (as, am) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) na) ires + \ dec_inv_1 (layout_of aprog) n e (as, am) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) + (Suc na)) ires \ + ((t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) (Suc na), + start_of (layout_of aprog) as, n, ires), + t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) na, + start_of (layout_of aprog) as, n, ires) + \ abc_dec_1_LE" +apply(rule allI, rule impI, simp add: t_steps_ind) +apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r) +(ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), +start_of (layout_of aprog) as - Suc 0) na)", simp) +apply(auto split:if_splits simp add:t_step.simps dec_inv_1.simps, + tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *}) +apply(simp_all add:dec_fetch_simps new_tape.simps dec_inv_1.simps) +apply(auto simp add: abc_dec_1_LE_def lex_square_def + lex_triple_def lex_pair_def + split: if_splits) +apply(rule dec_false_1, simp, simp) +done + +lemma dec_inv_stop1: + "\ly = layout_of aprog; + dec_inv_1 ly n e (as, am) (start_of ly as + 1, l, r) ires; + abc_fetch as aprog = Some (Dec n e); abc_lm_v am n = 0\ \ + (\ stp. (\ (s', l', r'). s' = start_of ly e \ + dec_inv_1 ly n e (as, am) (s', l' , r') ires) + (t_steps (start_of ly as + 1, l, r) + (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp))" +apply(insert halt_lemma2[of abc_dec_1_LE + "\ ((s, l, r), ss, n', ires'). s = start_of ly e" + "(\ stp. (t_steps (start_of ly as + 1, l, r) + (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) + stp, start_of ly as, n, ires))" + "\ ((s, l, r), ss, n, ires'). dec_inv_1 ly n e (as, am) (s, l, r) ires'"], + simp) +apply(insert wf_dec_le, simp) +apply(insert dec_inv_stop1_pre[of as aprog n e am l r], simp) +apply(subgoal_tac "start_of (layout_of aprog) as > 0", + simp add: t_steps.simps) +apply(erule_tac exE, rule_tac x = na in exI) +apply(case_tac + "(t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) na)", + case_tac b, auto) +apply(rule startof_not0) +done + +(*begin: dec halt2 lemmas*) + +lemma [simp]: + "\abc_fetch as aprog = Some (Dec n e); + ly = layout_of aprog\ \ + start_of ly (Suc as) = start_of ly as + 2*n + 16" +by simp + +fun abc_dec_2_stage1 :: "t_conf \ nat \ nat \ nat" + where + "abc_dec_2_stage1 (s, l, r) ss n = + (if s \ ss + 2*n + 1 then 7 + else if s = ss + 2*n + 2 then 6 + else if s = ss + 2*n + 3 then 5 + else if s \ ss + 2*n + 4 \ s \ ss + 2*n + 9 then 4 + else if s = ss + 2*n + 6 then 3 + else if s = ss + 2*n + 10 \ s = ss + 2*n + 11 then 2 + else if s = ss + 2*n + 12 then 1 + else 0)" + +thm new_tape.simps + +fun abc_dec_2_stage2 :: "t_conf \ nat \ nat \ nat" + where + "abc_dec_2_stage2 (s, l, r) ss n = + (if s \ ss + 2 * n + 1 then (ss + 2 * n + 16 - s) + else if s = ss + 2*n + 10 then length l + else if s = ss + 2*n + 11 then length l + else if s = ss + 2*n + 4 then length r - 1 + else if s = ss + 2*n + 5 then length r + else if s = ss + 2*n + 7 then length r - 1 + else if s = ss + 2*n + 8 then + length r + length (takeWhile (\ a. a = Oc) l) - 1 + else if s = ss + 2*n + 9 then + length r + length (takeWhile (\ a. a = Oc) l) - 1 + else 0)" + +fun abc_dec_2_stage3 :: "t_conf \ nat \ nat \ block list \ nat" + where + "abc_dec_2_stage3 (s, l, r) ss n ires = + (if s \ ss + 2*n + 1 then + if (s - ss) mod 2 = 0 then if r \ [] \ + hd r = Oc then 0 else 1 + else length r + else if s = ss + 2 * n + 10 then + if l = Bk # ires \ r \ [] \ hd r = Oc then 2 + else 1 + else if s = ss + 2 * n + 11 then + if r \ [] \ hd r = Oc then 3 + else 0 + else (ss + 2 * n + 16 - s))" + +fun abc_dec_2_stage4 :: "t_conf \ nat \ nat \ nat" + where + "abc_dec_2_stage4 (s, l, r) ss n = + (if s = ss + 2*n + 2 then length r + else if s = ss + 2*n + 8 then length r + else if s = ss + 2*n + 3 then + if r \ [] \ hd r = Oc then 1 + else 0 + else if s = ss + 2*n + 7 then + if r \ [] \ hd r = Oc then 0 + else 1 + else if s = ss + 2*n + 9 then + if r \ [] \ hd r = Oc then 1 + else 0 + else 0)" + +fun abc_dec_2_measure :: "(t_conf \ nat \ nat \ block list) \ + (nat \ nat \ nat \ nat)" + where + "abc_dec_2_measure (c, ss, n, ires) = + (abc_dec_2_stage1 c ss n, abc_dec_2_stage2 c ss n, + abc_dec_2_stage3 c ss n ires, abc_dec_2_stage4 c ss n)" + +definition abc_dec_2_LE :: + "(((nat \ block list \ block list) \ nat \ nat \ block list) \ + ((nat \ block list \ block list) \ nat \ nat \ block list)) set" + where "abc_dec_2_LE \ (inv_image lex_square abc_dec_2_measure)" + +lemma wf_dec_2_le: "wf abc_dec_2_LE" +by(auto intro:wf_inv_image wf_lex_triple wf_lex_square + simp:abc_dec_2_LE_def) + +lemma [simp]: "dec_after_write (as, am) (s, aa, r) ires + \ takeWhile (\a. a = Oc) aa = []" +apply(simp only : dec_after_write.simps) +apply(erule exE)+ +apply(erule_tac conjE)+ +apply(case_tac aa, simp) +apply(case_tac a, simp only: takeWhile.simps , simp, simp split: if_splits) +done + +lemma [simp]: + "\dec_on_right_moving (as, lm) (s, aa, []) ires; + length (takeWhile (\a. a = Oc) (tl aa)) + \ length (takeWhile (\a. a = Oc) aa) - Suc 0\ + \ length (takeWhile (\a. a = Oc) (tl aa)) < + length (takeWhile (\a. a = Oc) aa) - Suc 0" +apply(simp only: dec_on_right_moving.simps) +apply(erule_tac exE)+ +apply(erule_tac conjE)+ +apply(case_tac mr, auto split: if_splits) +done + +lemma [simp]: + "dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) + (start_of (layout_of aprog) as + 2 * n + 9, aa, Bk # xs) ires + \ length xs - Suc 0 < length xs + + length (takeWhile (\a. a = Oc) aa)" +apply(simp only: dec_after_clear.simps) +apply(erule_tac exE)+ +apply(erule conjE)+ +apply(simp split: if_splits ) +done + +lemma [simp]: + "\dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) + (start_of (layout_of aprog) as + 2 * n + 9, aa, []) ires\ + \ Suc 0 < length (takeWhile (\a. a = Oc) aa)" +apply(simp add: dec_after_clear.simps split: if_splits) +done + +lemma [simp]: + "\dec_on_right_moving (as, am) (s, aa, Bk # xs) ires; + Suc (length (takeWhile (\a. a = Oc) (tl aa))) + \ length (takeWhile (\a. a = Oc) aa)\ + \ Suc (length (takeWhile (\a. a = Oc) (tl aa))) + < length (takeWhile (\a. a = Oc) aa)" +apply(simp only: dec_on_right_moving.simps) +apply(erule exE)+ +apply(erule conjE)+ +apply(case_tac ml, auto split: if_splits ) +done + +(* +lemma abc_dec_2_wf: + "\ly = layout_of aprog; dec_inv_2 ly n e (as, am) (start_of ly as + 1, l, r); abc_fetch as aprog = Dec n e; abc_lm_v am n > 0\ + \ \na. \ (\(s, l, r) (ss, n'). s = start_of (layout_of aprog) as + 2*n + 16) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na) + (start_of (layout_of aprog) as, n) \ + ((t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) (Suc na), + start_of (layout_of aprog) as, n), + t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na, + start_of (layout_of aprog) as, n) + \ abc_dec_2_LE" +proof(rule allI, rule impI, simp add: t_steps_ind) + fix na + assume h1 :"ly = layout_of aprog" "dec_inv_2 (layout_of aprog) n e (as, am) (Suc (start_of (layout_of aprog) as), l, r)" + "abc_fetch as aprog = Dec n e" "abc_lm_v am n > 0" + "\ (\(s, l, r) (ss, n'). s = start_of (layout_of aprog) as + 2*n + 16) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na) + (start_of (layout_of aprog) as, n)" + thus "((t_step (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0), + start_of (layout_of aprog) as, n), + t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na, + start_of (layout_of aprog) as, n) + \ abc_dec_2_LE" + proof(insert dec_inv_2_steps[of "layout_of aprog" n e as am "(start_of (layout_of aprog) as + 1, l, r)" aprog na], + case_tac "(t_steps (start_of (layout_of aprog) as + 1, l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na)", case_tac b, simp) + fix a b aa ba + assume "dec_inv_2 (layout_of aprog) n e (as, am) (a, aa, ba)" " a \ start_of (layout_of aprog) as + 2*n + 16" "abc_lm_v am n > 0" "abc_fetch as aprog = Dec n e " + thus "((t_step (a, aa, ba) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0), start_of (layout_of aprog) as, n), (a, aa, ba), + start_of (layout_of aprog) as, n) + \ abc_dec_2_LE" + apply(case_tac "a = 0", auto split:if_splits simp add:t_step.simps dec_inv_2.simps, + tactic {* ALLGOALS (resolve_tac (thms "fetch_intro")) *}) + apply(simp_all add:dec_fetch_simps new_tape.simps) + apply(auto simp add: abc_dec_2_LE_def lex_square_def lex_triple_def lex_pair_def + split: if_splits) + + done + qed +qed +*) + +lemma [simp]: "inv_check_left_moving (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) + (start_of (layout_of aprog) as + 2 * n + 11, [], Oc # xs) ires = False" +apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps) +done + +lemma dec_inv_stop2_pre: + "\abc_fetch as aprog = Some (Dec n e); abc_lm_v am n > 0\ \ + \na. \ (\(s, l, r) (ss, n', ires'). + s = start_of (layout_of aprog) as + 2 * n + 16) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) na) + (start_of (layout_of aprog) as, n, ires) \ + dec_inv_2 (layout_of aprog) n e (as, am) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) na) ires + \ + dec_inv_2 (layout_of aprog) n e (as, am) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) (Suc na)) ires \ + ((t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) (Suc na), + start_of (layout_of aprog) as, n, ires), + t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) na, + start_of (layout_of aprog) as, n, ires) + \ abc_dec_2_LE" +apply(subgoal_tac "start_of (layout_of aprog) as > 0") +apply(rule allI, rule impI, simp add: t_steps_ind) +apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) na)", simp) +apply(auto split:if_splits simp add:t_step.simps dec_inv_2.simps, + tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *}) +apply(simp_all add:dec_fetch_simps new_tape.simps dec_inv_2.simps) +apply(auto simp add: abc_dec_2_LE_def lex_square_def lex_triple_def + lex_pair_def split: if_splits) +apply(auto intro: dec_false_2_b dec_false_2) +apply(rule startof_not0) +done + +lemma dec_stop2: + "\ly = layout_of aprog; + dec_inv_2 ly n e (as, am) (start_of ly as + 1, l, r) ires; + abc_fetch as aprog = Some (Dec n e); + abc_lm_v am n > 0\ \ + (\ stp. (\ (s', l', r'). s' = start_of ly (Suc as) \ + dec_inv_2 ly n e (as, am) (s', l', r') ires) + (t_steps (start_of ly as+1, l, r) (ci ly (start_of ly as) + (Dec n e), start_of ly as - Suc 0) stp))" +apply(insert halt_lemma2[of abc_dec_2_LE + "\ ((s, l, r), ss, n', ires'). s = start_of ly (Suc as)" + "(\ stp. (t_steps (start_of ly as + 1, l, r) + (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp, + start_of ly as, n, ires))" + "(\ ((s, l, r), ss, n, ires'). dec_inv_2 ly n e (as, am) (s, l, r) ires')"]) +apply(insert wf_dec_2_le, simp) +apply(insert dec_inv_stop2_pre[of as aprog n e am l r], + simp add: t_steps.simps) +apply(erule_tac exE) +apply(rule_tac x = na in exI) +apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r) +(ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) na)", + case_tac b, auto) +done + +lemma dec_inv_stop_cond1: + "\ly = layout_of aprog; + dec_inv_1 ly n e (as, lm) (s, (l, r)) ires; s = start_of ly e; + abc_fetch as aprog = Some (Dec n e); abc_lm_v lm n = 0\ + \ crsp_l ly (e, abc_lm_s lm n 0) (s, l, r) ires" +apply(simp add: dec_inv_1.simps split: if_splits) +apply(auto simp: crsp_l.simps inv_stop.simps ) +done + +lemma dec_inv_stop_cond2: + "\ly = layout_of aprog; s = start_of ly (Suc as); + dec_inv_2 ly n e (as, lm) (s, (l, r)) ires; + abc_fetch as aprog = Some (Dec n e); + abc_lm_v lm n > 0\ + \ crsp_l ly (Suc as, + abc_lm_s lm n (abc_lm_v lm n - Suc 0)) (s, l, r) ires" +apply(simp add: dec_inv_2.simps split: if_splits) +apply(auto simp: crsp_l.simps inv_stop.simps ) +done + +lemma [simp]: "(case Bk\<^bsup>rn\<^esup> of [] \ Bk | + Bk # xs \ Bk | Oc # xs \ Oc) = Bk" +apply(case_tac rn, auto) +done + +lemma [simp]: "t_steps tc (p,off) (m + n) = + t_steps (t_steps tc (p, off) m) (p, off) n" +apply(induct m arbitrary: n) +apply(simp add: t_steps.simps) +proof - + fix m n + assume h1: "\n. t_steps tc (p, off) (m + n) = + t_steps (t_steps tc (p, off) m) (p, off) n" + hence h2: "t_steps tc (p, off) (Suc m + n) = + t_steps tc (p, off) (m + Suc n)" + by simp + from h1 and this show + "t_steps tc (p, off) (Suc m + n) = + t_steps (t_steps tc (p, off) (Suc m)) (p, off) n" + proof(simp only: h2, simp add: t_steps.simps) + have h3: "(t_step (t_steps tc (p, off) m) (p, off)) = + (t_steps (t_step tc (p, off)) (p, off) m)" + apply(simp add: t_steps.simps[THEN sym] t_steps_ind[THEN sym]) + done + from h3 show + "t_steps (t_step (t_steps tc (p, off) m) (p, off)) (p, off) n = t_steps (t_steps (t_step tc (p, off)) (p, off) m) (p, off) n" + by simp + qed +qed + +lemma [simp]: " abc_fetch as aprog = Some (Dec n e) \ + Suc (start_of (layout_of aprog) as) \ + start_of (layout_of aprog) e" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = "layout_of aprog" + in start_of_le, simp) +apply(drule_tac start_of_ge, auto) +done + +lemma [simp]: "inv_locate_b (as, []) (0, Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires" +apply(auto simp: inv_locate_b.simps in_middle.simps) +apply(rule_tac x = "[]" in exI, rule_tac x = "Suc 0" in exI, + rule_tac x = 0 in exI, simp) +apply(rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI, auto) +apply(case_tac rn, simp, case_tac nat, auto) +done + +lemma [simp]: + "inv_locate_n_b (as, []) (0, Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires" +apply(auto simp: inv_locate_n_b.simps in_middle.simps) +apply(case_tac rn, simp, case_tac nat, auto) +done + +lemma [simp]: +"abc_fetch as aprog = Some (Dec n e) \ + dec_inv_1 (layout_of aprog) n e (as, []) + (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires +\ + dec_inv_2 (layout_of aprog) n e (as, []) + (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires" +apply(simp add: dec_inv_1.simps dec_inv_2.simps) +apply(case_tac n, auto) +done + +lemma [simp]: + "\am \ []; = Oc # r'; + abc_fetch as aprog = Some (Dec n e)\ + \ inv_locate_b (as, am) (0, Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires" +apply(auto simp: inv_locate_b.simps in_middle.simps) +apply(rule_tac x = "tl am" in exI, rule_tac x = 0 in exI, + rule_tac x = "hd am" in exI, simp) +apply(rule_tac x = "Suc 0" in exI) +apply(rule_tac x = "hd am" in exI, simp) +apply(case_tac am, simp, case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac rn, auto) +done + +lemma [simp]: + "\ = Oc # r'; abc_fetch as aprog = Some (Dec n e)\ \ + inv_locate_n_b (as, am) (0, Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires" +apply(auto simp: inv_locate_n_b.simps) +apply(rule_tac x = "tl am" in exI, rule_tac x = "hd am" in exI, auto) +apply(case_tac [!] am, auto simp: tape_of_nl_abv tape_of_nat_list.simps ) +apply(case_tac [!]list, auto simp: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac rn, simp, simp) +apply(erule_tac x = nat in allE, simp) +done + +lemma [simp]: + "\am \ []; + = Oc # r'; + abc_fetch as aprog = Some (Dec n e)\ \ + dec_inv_1 (layout_of aprog) n e (as, am) + (Suc (start_of (layout_of aprog) as), + Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires \ + dec_inv_2 (layout_of aprog) n e (as, am) + (Suc (start_of (layout_of aprog) as), + Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires" +apply(simp add: dec_inv_1.simps dec_inv_2.simps) +apply(case_tac n, auto) +done + +lemma [simp]: "am \ [] \ \r'. = Oc # r'" +apply(case_tac am, simp, case_tac list) +apply(auto simp: tape_of_nl_abv tape_of_nat_list.simps ) +done + +lemma [simp]: "start_of (layout_of aprog) as > 0 \ + (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e)) (Suc 0) Bk) + = (W1, start_of (layout_of aprog) as)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append Suc_pre tdec_b_def) +thm findnth_nth +apply(insert findnth_nth[of 0 n 0], simp) +apply(simp add: findnth.simps) +done + +lemma [simp]: + "start_of (layout_of aprog) as > 0 + \ (t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Bk\<^bsup>rn\<^esup>) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0)) + = (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # Bk\<^bsup>rn- Suc 0\<^esup>)" +apply(simp add: t_step.simps) +apply(case_tac "start_of (layout_of aprog) as", + auto simp: new_tape.simps) +apply(case_tac rn, auto) +done + +lemma [simp]: "start_of (layout_of aprog) as > 0 \ + (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) (Suc 0) Oc) + = (R, Suc (start_of (layout_of aprog) as))" + +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append + Suc_pre tdec_b_def) +apply(insert findnth_nth[of 0 n "Suc 0"], simp) +apply(simp add: findnth.simps) +done + +lemma [simp]: "start_of (layout_of aprog) as > 0 \ + (t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # Bk\<^bsup>rn - Suc 0\<^esup>) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0)) = + (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn-Suc 0\<^esup>)" +apply(simp add: t_step.simps) +apply(case_tac "start_of (layout_of aprog) as", + auto simp: new_tape.simps) +done + +lemma [simp]: "start_of (layout_of aprog) as > 0 \ + t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # r' @ Bk\<^bsup>rn\<^esup>) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) = + (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>)" +apply(simp add: t_step.simps) +apply(case_tac "start_of (layout_of aprog) as", + auto simp: new_tape.simps) +done + +lemma crsp_next_state: + "\crsp_l (layout_of aprog) (as, am) tc ires; + abc_fetch as aprog = Some (Dec n e)\ + \ \ stp' > 0. (\ (s, l, r). + (s = Suc (start_of (layout_of aprog) as) + \ (dec_inv_1 (layout_of aprog) n e (as, am) (s, l, r) ires) + \ (dec_inv_2 (layout_of aprog) n e (as, am) (s, l, r)) ires)) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')" +apply(subgoal_tac "start_of (layout_of aprog) as > 0") +apply(case_tac tc, case_tac b, auto simp: crsp_l.simps) +apply(case_tac "am = []", simp) +apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: t_steps.simps) +proof- + fix rn + assume h1: "am \ []" "abc_fetch as aprog = Some (Dec n e)" + "start_of (layout_of aprog) as > 0" + hence h2: "\ r'. = Oc # r'" + by simp + from h1 and h2 show + "\stp'>0. case t_steps (start_of (layout_of aprog) as, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp' of + (s, ab) \ s = Suc (start_of (layout_of aprog) as) \ + dec_inv_1 (layout_of aprog) n e (as, am) (s, ab) ires \ + dec_inv_2 (layout_of aprog) n e (as, am) (s, ab) ires" + proof(erule_tac exE, simp, rule_tac x = "Suc 0" in exI, + simp add: t_steps.simps) + qed +next + assume "abc_fetch as aprog = Some (Dec n e)" + thus "0 < start_of (layout_of aprog) as" + apply(insert startof_not0[of "layout_of aprog" as], simp) + done +qed + +lemma dec_crsp_ex1: + "\crsp_l (layout_of aprog) (as, am) tc ires; + abc_fetch as aprog = Some (Dec n e); + abc_lm_v am n = 0\ + \ \stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires" +proof - + assume h1: "crsp_l (layout_of aprog) (as, am) tc ires" + "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n = 0" + hence h2: "\ stp' > 0. (\ (s, l, r). + (s = Suc (start_of (layout_of aprog) as) \ + (dec_inv_1 (layout_of aprog) n e (as, am) (s, l, r)) ires)) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')" + apply(insert crsp_next_state[of aprog as am tc ires n e], auto) + done + from h1 and h2 show + "\stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) + (t_steps tc (ci (layout_of aprog) (start_of + (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp) ires" + proof(erule_tac exE, case_tac "(t_steps tc (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e), start_of + (layout_of aprog) as - Suc 0) stp')", simp) + fix stp' a b c + assume h3: "stp' > 0 \ a = Suc (start_of (layout_of aprog) as) \ + dec_inv_1 (layout_of aprog) n e (as, am) (a, b, c) ires" + "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n = 0" + "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp' + = (Suc (start_of (layout_of aprog) as), b, c)" + thus "\stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires" + proof(erule_tac conjE, simp) + assume "dec_inv_1 (layout_of aprog) n e (as, am) + (Suc (start_of (layout_of aprog) as), b, c) ires" + "abc_fetch as aprog = Some (Dec n e)" + "abc_lm_v am n = 0" + " t_steps tc (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp' + = (Suc (start_of (layout_of aprog) as), b, c)" + hence h4: "\stp. (\(s', l', r'). s' = + start_of (layout_of aprog) e \ + dec_inv_1 (layout_of aprog) n e (as, am) (s', l', r') ires) + (t_steps (start_of (layout_of aprog) as + 1, b, c) + (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp)" + apply(rule_tac dec_inv_stop1, auto) + done + from h3 and h4 show ?thesis + apply(erule_tac exE) + apply(rule_tac x = "stp' + stp" in exI, simp) + apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), + b, c) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp)", + simp) + apply(rule_tac dec_inv_stop_cond1, auto) + done + qed + qed +qed + +lemma dec_crsp_ex2: + "\crsp_l (layout_of aprog) (as, am) tc ires; + abc_fetch as aprog = Some (Dec n e); + 0 < abc_lm_v am n\ + \ \stp > 0. crsp_l (layout_of aprog) + (Suc as, abc_lm_s am n (abc_lm_v am n - Suc 0)) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires" +proof - + assume h1: + "crsp_l (layout_of aprog) (as, am) tc ires" + "abc_fetch as aprog = Some (Dec n e)" + "abc_lm_v am n > 0" + hence h2: + "\ stp' > 0. (\ (s, l, r). (s = Suc (start_of (layout_of aprog) as) + \ (dec_inv_2 (layout_of aprog) n e (as, am) (s, l, r)) ires)) +(t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')" + apply(insert crsp_next_state[of aprog as am tc ires n e], auto) + done + from h1 and h2 show + "\stp >0. crsp_l (layout_of aprog) + (Suc as, abc_lm_s am n (abc_lm_v am n - Suc 0)) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires" + proof(erule_tac exE, + case_tac + "(t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')", simp) + fix stp' a b c + assume h3: "0 < stp' \ a = Suc (start_of (layout_of aprog) as) \ + dec_inv_2 (layout_of aprog) n e (as, am) (a, b, c) ires" + "abc_fetch as aprog = Some (Dec n e)" + "abc_lm_v am n > 0" + "t_steps tc (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp' + = (Suc (start_of (layout_of aprog) as), b, c)" + thus "?thesis" + proof(erule_tac conjE, simp) + assume + "dec_inv_2 (layout_of aprog) n e (as, am) + (Suc (start_of (layout_of aprog) as), b, c) ires" + "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n > 0" + "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp' + = (Suc (start_of (layout_of aprog) as), b, c)" + hence h4: + "\stp. (\(s', l', r'). s' = start_of (layout_of aprog) (Suc as) \ + dec_inv_2 (layout_of aprog) n e (as, am) (s', l', r') ires) + (t_steps (start_of (layout_of aprog) as + 1, b, c) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp)" + apply(rule_tac dec_stop2, auto) + done + from h3 and h4 show ?thesis + apply(erule_tac exE) + apply(rule_tac x = "stp' + stp" in exI, simp) + apply(case_tac + "(t_steps (Suc (start_of (layout_of aprog) as), b, c) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp)" + ,simp) + apply(rule_tac dec_inv_stop_cond2, auto) + done + qed + qed +qed + +lemma dec_crsp_ex_pre: + "\ly = layout_of aprog; crsp_l ly (as, am) tc ires; + abc_fetch as aprog = Some (Dec n e)\ + \ \stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Dec n e))) + (t_steps tc (ci (layout_of aprog) (start_of ly as) (Dec n e), + start_of ly as - Suc 0) stp) ires" +apply(auto simp: abc_step_l.simps intro: dec_crsp_ex2 dec_crsp_ex1) +done + +lemma dec_crsp_ex: + assumes layout: -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"} *} + "ly = layout_of aprog" + and dec: -- {* There is an @{text "Dec n e"} instruction at postion @{text "as"} of @{text "aprog"} *} + "abc_fetch as aprog = Some (Dec n e)" + and correspond: + -- {* Abacus configuration @{text "(as, am)"} is in correspondence with TM + configuration @{text "tc"} + *} + "crsp_l ly (as, am) tc ires" +shows + "\stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Dec n e))) + (t_steps tc (ci (layout_of aprog) (start_of ly as) (Dec n e), + start_of ly as - Suc 0) stp) ires" +proof - + from dec_crsp_ex_pre layout dec correspond show ?thesis by blast +qed + + +(*******End: dec crsp********) + + +subsubsection {* Compilation of @{text "Goto n"}*} + + +(*******Begin: goto crsp********) +lemma goto_fetch: + "fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Goto n)) (Suc 0) b + = (Nop, start_of (layout_of aprog) n)" +apply(auto simp: ci.simps fetch.simps nth_of.simps + split: block.splits) +done + +text {* + Correctness of complied @{text "Goto n"} + *} + +lemma goto_crsp_ex_pre: + "\ly = layout_of aprog; + crsp_l ly (as, am) tc ires; + abc_fetch as aprog = Some (Goto n)\ + \ \stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Goto n))) + (t_steps tc (ci (layout_of aprog) (start_of ly as) (Goto n), + start_of ly as - Suc 0) stp) ires" +apply(rule_tac x = 1 in exI) +apply(simp add: abc_step_l.simps t_steps.simps t_step.simps) +apply(case_tac tc, simp) +apply(subgoal_tac "a = start_of (layout_of aprog) as", auto) +apply(subgoal_tac "start_of (layout_of aprog) as > 0", simp) +apply(auto simp: goto_fetch new_tape.simps crsp_l.simps) +apply(rule startof_not0) +done + +lemma goto_crsp_ex: + assumes layout: "ly = layout_of aprog" + and goto: "abc_fetch as aprog = Some (Goto n)" + and correspondence: "crsp_l ly (as, am) tc ires" + shows "\stp>0. crsp_l ly (abc_step_l (as, am) (Some (Goto n))) + (t_steps tc (ci (layout_of aprog) (start_of ly as) (Goto n), + start_of ly as - Suc 0) stp) ires" +proof - + from goto_crsp_ex_pre and layout goto correspondence show "?thesis" by blast +qed +(*******End : goto crsp*********) + +subsubsection {* + The correctness of the compiler + *} + +declare abc_step_l.simps[simp del] + +lemma tm_crsp_ex: + "\ly = layout_of aprog; + crsp_l ly (as, am) tc ires; + as < length aprog; + abc_fetch as aprog = Some ins\ + \ \ n > 0. crsp_l ly (abc_step_l (as,am) (Some ins)) + (t_steps tc (ci (layout_of aprog) (start_of ly as) + (ins), (start_of ly as) - (Suc 0)) n) ires" +apply(case_tac "ins", simp) +apply(auto intro: inc_crsp_ex_pre dec_crsp_ex goto_crsp_ex) +done + +lemma start_of_pre: + "n < length aprog \ start_of (layout_of aprog) n + = start_of (layout_of (butlast aprog)) n" +apply(induct n, simp add: start_of.simps, simp) +apply(simp add: layout_of.simps start_of.simps) +apply(subgoal_tac "n < length aprog - Suc 0", simp) +apply(subgoal_tac "(aprog ! n) = (butlast aprog ! n)", simp) +proof - + fix n + assume h1: "Suc n < length aprog" + thus "aprog ! n = butlast aprog ! n" + apply(case_tac "length aprog", simp, simp) + apply(insert nth_append[of "butlast aprog" "[last aprog]" n]) + apply(subgoal_tac "(butlast aprog @ [last aprog]) = aprog") + apply(simp split: if_splits) + apply(rule append_butlast_last_id, case_tac aprog, simp, simp) + done +next + fix n + assume "Suc n < length aprog" + thus "n < length aprog - Suc 0" + apply(case_tac aprog, simp, simp) + done +qed + +lemma zip_eq: "xs = ys \ zip xs zs = zip ys zs" +by simp + +lemma tpairs_of_append_iff: "length aprog = Suc n \ + tpairs_of aprog = tpairs_of (butlast aprog) @ + [(start_of (layout_of aprog) n, aprog ! n)]" +apply(simp add: tpairs_of.simps) +apply(insert zip_append[of "map (start_of (layout_of aprog)) [0..(n, tm). abacus.t_ncorrect (ci layout n tm)) + (zip (map (start_of layout) [0..aprog. x = length aprog \ + list_all (\(n, tm). abacus.t_ncorrect (ci layout n tm)) + (zip (map (start_of layout) [0..(n, tm). abacus.t_ncorrect (ci layout n tm)) + (zip (map (start_of layout) [0.. xs a. aprog = xs @ [a]" + using h + apply(rule_tac x = "butlast aprog" in exI, + rule_tac x = "last aprog" in exI) + apply(case_tac "aprog = []", simp, simp) + done + from this obtain xs where "\ a. aprog = xs @ [a]" .. + from this obtain a where g3: "aprog = xs @ [a]" .. + from g1 and g2 and g3 show "list_all (\(n, tm). + abacus.t_ncorrect (ci layout n tm)) + (zip (map (start_of layout) [0..ly = layout_of aprog; tprog = tm_of aprog; + crsp_l ly (as, am) tc ires; length aprog \ as\ + \ abc_step_l (as, am) (abc_fetch as aprog) = (as, am)" +apply(simp add: abc_fetch.simps abc_step_l.simps) +done + +lemma tm_merge_ex: + "\crsp_l (layout_of aprog) (as, am) tc ires; + as < length aprog; + abc_fetch as aprog = Some a; + abc2t_correct aprog; + crsp_l (layout_of aprog) (abc_step_l (as, am) (Some a)) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + a, start_of (layout_of aprog) as - Suc 0) n) ires; + n > 0\ + \ \stp > 0. crsp_l (layout_of aprog) (abc_step_l (as, am) + (Some a)) (t_steps tc (tm_of aprog, 0) stp) ires" +apply(case_tac "(t_steps tc (ci (layout_of aprog) + (start_of (layout_of aprog) as) a, + start_of (layout_of aprog) as - Suc 0) n)", simp) +apply(case_tac "(abc_step_l (as, am) (Some a))", simp) +proof - + fix aa b c aaa ba + assume h: + "crsp_l (layout_of aprog) (as, am) tc ires" + "as < length aprog" + "abc_fetch as aprog = Some a" + "crsp_l (layout_of aprog) (aaa, ba) (aa, b, c) ires" + "abc2t_correct aprog" + "n>0" + "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) a, + start_of (layout_of aprog) as - Suc 0) n = (aa, b, c)" + "abc_step_l (as, am) (Some a) = (aaa, ba)" + hence "aa = start_of (layout_of aprog) aaa" + apply(simp add: crsp_l.simps) + done + from this and h show + "\stp > 0. crsp_l (layout_of aprog) (aaa, ba) + (t_steps tc (tm_of aprog, 0) stp) ires" + apply(insert tms_out_ex[of "layout_of aprog" aprog + "tm_of aprog" as am tc ires a n aa b c aaa ba], auto) + done +qed + +lemma crsp_inside: + "\ly = layout_of aprog; + tprog = tm_of aprog; + crsp_l ly (as, am) tc ires; + as < length aprog\ \ + (\ stp > 0. crsp_l ly (abc_step_l (as,am) (abc_fetch as aprog)) + (t_steps tc (tprog, 0) stp) ires)" +apply(case_tac "abc_fetch as aprog", simp add: abc_fetch.simps) +proof - + fix a + assume "ly = layout_of aprog" + "tprog = tm_of aprog" + "crsp_l ly (as, am) tc ires" + "as < length aprog" + "abc_fetch as aprog = Some a" + thus "\stp > 0. crsp_l ly (abc_step_l (as, am) + (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires" + proof(insert tm_crsp_ex[of ly aprog as am tc ires a], + auto intro: tm_merge_ex) + qed +qed + +lemma crsp_outside: + "\ly = layout_of aprog; tprog = tm_of aprog; + crsp_l ly (as, am) tc ires; as \ length aprog\ + \ (\ stp. crsp_l ly (abc_step_l (as,am) (abc_fetch as aprog)) + (t_steps tc (tprog, 0) stp) ires)" +apply(subgoal_tac "abc_step_l (as, am) (abc_fetch as aprog) + = (as, am)", simp) +apply(rule_tac x = 0 in exI, simp add: t_steps.simps) +apply(rule as_out, simp+) +done + +text {* + Single-step correntess of the compiler. +*} +lemma astep_crsp_pre: + "\ly = layout_of aprog; + tprog = tm_of aprog; + crsp_l ly (as, am) tc ires\ + \ (\ stp. crsp_l ly (abc_step_l (as,am) + (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)" +apply(case_tac "as < length aprog") +apply(drule_tac crsp_inside, auto) +apply(rule_tac crsp_outside, simp+) +done + +text {* + Single-step correntess of the compiler. +*} +lemma astep_crsp_pre1: + "\ly = layout_of aprog; + tprog = tm_of aprog; + crsp_l ly (as, am) tc ires\ + \ (\ stp. crsp_l ly (abc_step_l (as,am) + (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)" +apply(case_tac "as < length aprog") +apply(drule_tac crsp_inside, auto) +apply(rule_tac crsp_outside, simp+) +done + +lemma astep_crsp: + assumes layout: + -- {* There is a Abacus program @{text "aprog"} with layout @{text "ly"} *} + "ly = layout_of aprog" + and compiled: + -- {* @{text "tprog"} is the TM compiled from @{text "aprog"} *} + "tprog = tm_of aprog" + and corresponds: + -- {* Abacus configuration @{text "(as, am)"} is in correspondence with TM configuration + @{text "tc"} *} + "crsp_l ly (as, am) tc ires" + -- {* One step execution of @{text "aprog"} can be simulated by multi-step execution + of @{text "tprog"} *} + shows "(\ stp. crsp_l ly (abc_step_l (as,am) + (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)" +proof - + from astep_crsp_pre1 [OF layout compiled corresponds] show ?thesis . +qed + +lemma steps_crsp_pre: + "\ly = layout_of aprog; tprog = tm_of aprog; + crsp_l ly ac tc ires; ac' = abc_steps_l ac aprog n\ \ + (\ n'. crsp_l ly ac' (t_steps tc (tprog, 0) n') ires)" +apply(induct n arbitrary: ac' ac tc, simp add: abc_steps_l.simps) +apply(rule_tac x = 0 in exI) +apply(case_tac ac, simp add: abc_steps_l.simps t_steps.simps) +apply(case_tac ac, simp add: abc_steps_l.simps) +apply(subgoal_tac + "(\ stp. crsp_l ly (abc_step_l (a, b) + (abc_fetch a aprog)) (t_steps tc (tprog, 0) stp) ires)") +apply(erule exE) +apply(subgoal_tac + "\n'. crsp_l (layout_of aprog) + (abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) aprog n) + (t_steps ((t_steps tc (tprog, 0) stp)) (tm_of aprog, 0) n') ires") +apply(erule exE) +apply(subgoal_tac + "t_steps (t_steps tc (tprog, 0) stp) (tm_of aprog, 0) n' = + t_steps tc (tprog, 0) (stp + n')") +apply(rule_tac x = "stp + n'" in exI, simp) +apply(auto intro: astep_crsp simp: t_step_add) +done + +text {* + Multi-step correctess of the compiler. +*} + +lemma steps_crsp: + assumes layout: + -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"} *} + "ly = layout_of aprog" + and compiled: + -- {* @{text "tprog"} is the TM compiled from @{text "aprog"} *} + "tprog = tm_of aprog" + and correspond: + -- {* Abacus configuration @{text "ac"} is in correspondence with TM configuration @{text "tc"} *} + "crsp_l ly ac tc ires" + and execution: + -- {* @{text "ac'"} is the configuration obtained from @{text "n"}-step execution + of @{text "aprog"} starting from configuration @{text "ac"} *} + "ac' = abc_steps_l ac aprog n" + -- {* There exists steps @{text "n'"} steps, after these steps of execution, + the Turing configuration such obtained is in correspondence with @{text "ac'"} *} + shows "(\ n'. crsp_l ly ac' (t_steps tc (tprog, 0) n') ires)" +proof - + from steps_crsp_pre [OF layout compiled correspond execution] show ?thesis . +qed + + +subsubsection {* The Mop-up machine *} + +fun mop_bef :: "nat \ tprog" + where + "mop_bef 0 = []" | + "mop_bef (Suc n) = mop_bef n @ + [(R, 2*n + 3), (W0, 2*n + 2), (R, 2*n + 1), (W1, 2*n + 2)]" + +definition mp_up :: "tprog" + where + "mp_up \ [(R, 2), (R, 1), (L, 5), (W0, 3), (R, 4), (W0, 3), + (R, 2), (W0, 3), (L, 5), (L, 6), (R, 0), (L, 6)]" + +fun tMp :: "nat \ nat \ tprog" + where + "tMp n off = tshift (mop_bef n @ tshift mp_up (2*n)) off" + +declare mp_up_def[simp del] tMp.simps[simp del] mop_bef.simps[simp del] +(**********Begin: equiv among aba and turing***********) + +lemma tm_append_step: + "\t_ncorrect tp1; t_step tc (tp1, 0) = (s, l, r); s \ 0\ + \ t_step tc (tp1 @ tp2, 0) = (s, l, r)" +apply(simp add: t_step.simps) +apply(case_tac tc, simp) +apply(case_tac + "(fetch tp1 a (case c of [] \ Bk | + Bk # xs \ Bk | Oc # xs \ Oc))", simp) +apply(case_tac a, simp add: fetch.simps) +apply(simp add: fetch.simps) +apply(case_tac c, simp) +apply(case_tac [!] "ab::block") +apply(auto simp: nth_of.simps nth_append t_ncorrect.simps + split: if_splits) +done + +lemma state0_ind: "t_steps (0, l, r) (tp, 0) stp = (0, l, r)" +apply(induct stp, simp add: t_steps.simps) +apply(simp add: t_steps.simps t_step.simps fetch.simps new_tape.simps) +done + +lemma tm_append_steps: + "\t_ncorrect tp1; t_steps tc (tp1, 0) stp = (s, l ,r); s \ 0\ + \ t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)" +apply(induct stp arbitrary: tc s l r) +apply(case_tac tc, simp) +apply(simp add: t_steps.simps) +proof - + fix stp tc s l r + assume h1: "\tc s l r. \t_ncorrect tp1; t_steps tc (tp1, 0) stp = + (s, l, r); s \ 0\ \ t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)" + and h2: "t_steps tc (tp1, 0) (Suc stp) = (s, l, r)" "s \ 0" + "t_ncorrect tp1" + thus "t_steps tc (tp1 @ tp2, 0) (Suc stp) = (s, l, r)" + apply(simp add: t_steps.simps) + apply(case_tac "(t_step tc (tp1, 0))", simp) + proof- + fix a b c + assume g1: "\tc s l r. \t_steps tc (tp1, 0) stp = (s, l, r); + 0 < s\ \ t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)" + and g2: "t_step tc (tp1, 0) = (a, b, c)" + "t_steps (a, b, c) (tp1, 0) stp = (s, l, r)" + "0 < s" + "t_ncorrect tp1" + hence g3: "a > 0" + apply(case_tac "a::nat", auto simp: t_steps.simps) + apply(simp add: state0_ind) + done + from g1 and g2 and this have g4: + "(t_step tc (tp1 @ tp2, 0)) = (a, b, c)" + apply(rule_tac tm_append_step, simp, simp, simp) + done + from g1 and g2 and g3 and g4 show + "t_steps (t_step tc (tp1 @ tp2, 0)) (tp1 @ tp2, 0) stp + = (s, l, r)" + apply(simp) + done + qed +qed + +lemma shift_fetch: + "\n < length tp; + (tp:: (taction \ nat) list) ! n = (aa, ba); + ba \ 0\ + \ (tshift tp (length tp div 2)) ! n = + (aa , ba + length tp div 2)" +apply(simp add: tshift.simps) +done + +lemma tshift_length_equal: "length (tshift tp q) = length tp" +apply(auto simp: tshift.simps) +done + +thm nth_of.simps + +lemma [simp]: "t_ncorrect tp \ 2 * (length tp div 2) = length tp" +apply(auto simp: t_ncorrect.simps) +done + +lemma tm_append_step_equal': + "\t_ncorrect tp; t_ncorrect tp'; off = length tp div 2\ \ + (\ (s, l, r). ((\ (s', l', r'). + (s'\ 0 \ (s = s' + off \ l = l' \ r = r'))) + (t_step (a, b, c) (tp', 0)))) + (t_step (a + off, b, c) (tp @ tshift tp' off, 0))" +apply(simp add: t_step.simps) +apply(case_tac a, simp add: fetch.simps) +apply(case_tac +"(fetch tp' a (case c of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc))", + simp) +apply(case_tac +"(fetch (tp @ tshift tp' (length tp div 2)) + (Suc (nat + length tp div 2)) + (case c of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc))", + simp) +apply(case_tac "(new_tape aa (b, c))", + case_tac "(new_tape aaa (b, c))", simp, + rule impI, simp add: fetch.simps split: block.splits option.splits) +apply (auto simp: nth_of.simps t_ncorrect.simps + nth_append tshift_length_equal tshift.simps split: if_splits) +done + + +lemma tm_append_step_equal: + "\t_ncorrect tp; t_ncorrect tp'; off = length tp div 2; + t_step (a, b, c) (tp', 0) = (aa, ab, bb); aa \ 0\ + \ t_step (a + length tp div 2, b, c) + (tp @ tshift tp' (length tp div 2), 0) + = (aa + length tp div 2, ab, bb)" +apply(insert tm_append_step_equal'[of tp tp' off a b c], simp) +apply(case_tac "(t_step (a + length tp div 2, b, c) + (tp @ tshift tp' (length tp div 2), 0))", simp) +done + +lemma tm_append_steps_equal: + "\t_ncorrect tp; t_ncorrect tp'; off = length tp div 2\ \ + (\ (s, l, r). ((\ (s', l', r'). ((s'\ 0 \ s = s' + off \ l = l' + \ r = r'))) (t_steps (a, b, c) (tp', 0) stp))) + (t_steps (a + off, b, c) (tp @ tshift tp' off, 0) stp)" +apply(induct stp arbitrary : a b c, simp add: t_steps.simps) +apply(simp add: t_steps.simps) +apply(case_tac "(t_step (a, b, c) (tp', 0))", simp) +apply(case_tac "aa = 0", simp add: state0_ind) +apply(subgoal_tac "(t_step (a + length tp div 2, b, c) + (tp @ tshift tp' (length tp div 2), 0)) + = (aa + length tp div 2, ba, ca)", simp) +apply(rule tm_append_step_equal, auto) +done + +(*********Begin: mop_up***************) +type_synonym mopup_type = "t_conf \ nat list \ nat \ block list \ bool" + +fun mopup_stop :: "mopup_type" + where + "mopup_stop (s, l, r) lm n ires= + (\ ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \ r = @ Bk\<^bsup>rn\<^esup>)" + +fun mopup_bef_erase_a :: "mopup_type" + where + "mopup_bef_erase_a (s, l, r) lm n ires= + (\ ln m rn. l = Bk \<^bsup>ln\<^esup> @ Bk # Bk # ires \ + r = Oc\<^bsup>m \<^esup>@ Bk # <(drop ((s + 1) div 2) lm)> @ Bk\<^bsup>rn\<^esup>)" + +fun mopup_bef_erase_b :: "mopup_type" + where + "mopup_bef_erase_b (s, l, r) lm n ires = + (\ ln m rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \ r = Bk # Oc\<^bsup>m\<^esup> @ Bk # + <(drop (s div 2) lm)> @ Bk\<^bsup>rn\<^esup>)" + + +fun mopup_jump_over1 :: "mopup_type" + where + "mopup_jump_over1 (s, l, r) lm n ires = + (\ ln m1 m2 rn. m1 + m2 = Suc (abc_lm_v lm n) \ + l = Oc\<^bsup>m1\<^esup> @ Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \ + (r = Oc\<^bsup>m2\<^esup> @ Bk # <(drop (Suc n) lm)> @ Bk\<^bsup>rn \<^esup>\ + (r = Oc\<^bsup>m2\<^esup> \ (drop (Suc n) lm) = [])))" + +fun mopup_aft_erase_a :: "mopup_type" + where + "mopup_aft_erase_a (s, l, r) lm n ires = + (\ lnl lnr rn (ml::nat list) m. + m = Suc (abc_lm_v lm n) \ l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \ + (r = @ Bk\<^bsup>rn\<^esup>))" + +fun mopup_aft_erase_b :: "mopup_type" + where + "mopup_aft_erase_b (s, l, r) lm n ires= + (\ lnl lnr rn (ml::nat list) m. + m = Suc (abc_lm_v lm n) \ + l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \ + (r = Bk # @ Bk\<^bsup>rn \<^esup>\ + r = Bk # Bk # @ Bk\<^bsup>rn\<^esup>))" + +fun mopup_aft_erase_c :: "mopup_type" + where + "mopup_aft_erase_c (s, l, r) lm n ires = + (\ lnl lnr rn (ml::nat list) m. + m = Suc (abc_lm_v lm n) \ + l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \ + (r = @ Bk\<^bsup>rn \<^esup>\ r = Bk # @ Bk\<^bsup>rn\<^esup>))" + +fun mopup_left_moving :: "mopup_type" + where + "mopup_left_moving (s, l, r) lm n ires = + (\ lnl lnr rn m. + m = Suc (abc_lm_v lm n) \ + ((l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \ r = Bk\<^bsup>rn\<^esup>) \ + (l = Oc\<^bsup>m - 1\<^esup> @ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \ r = Oc # Bk\<^bsup>rn\<^esup>)))" + +fun mopup_jump_over2 :: "mopup_type" + where + "mopup_jump_over2 (s, l, r) lm n ires = + (\ ln rn m1 m2. + m1 + m2 = Suc (abc_lm_v lm n) + \ r \ [] + \ (hd r = Oc \ (l = Oc\<^bsup>m1\<^esup> @ Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \ r = Oc\<^bsup>m2\<^esup> @ Bk\<^bsup>rn\<^esup>)) + \ (hd r = Bk \ (l = Bk\<^bsup>ln\<^esup> @ Bk # ires \ r = Bk # Oc\<^bsup>m1 + m2\<^esup> @ Bk\<^bsup>rn\<^esup>)))" + + +fun mopup_inv :: "mopup_type" + where + "mopup_inv (s, l, r) lm n ires = + (if s = 0 then mopup_stop (s, l, r) lm n ires + else if s \ 2*n then + if s mod 2 = 1 then mopup_bef_erase_a (s, l, r) lm n ires + else mopup_bef_erase_b (s, l, r) lm n ires + else if s = 2*n + 1 then + mopup_jump_over1 (s, l, r) lm n ires + else if s = 2*n + 2 then mopup_aft_erase_a (s, l, r) lm n ires + else if s = 2*n + 3 then mopup_aft_erase_b (s, l, r) lm n ires + else if s = 2*n + 4 then mopup_aft_erase_c (s, l, r) lm n ires + else if s = 2*n + 5 then mopup_left_moving (s, l, r) lm n ires + else if s = 2*n + 6 then mopup_jump_over2 (s, l, r) lm n ires + else False)" + +declare + mopup_jump_over2.simps[simp del] mopup_left_moving.simps[simp del] + mopup_aft_erase_c.simps[simp del] mopup_aft_erase_b.simps[simp del] + mopup_aft_erase_a.simps[simp del] mopup_jump_over1.simps[simp del] + mopup_bef_erase_a.simps[simp del] mopup_bef_erase_b.simps[simp del] + mopup_stop.simps[simp del] + +lemma mopup_fetch_0[simp]: + "(fetch (mop_bef n @ tshift mp_up (2 * n)) 0 b) = (Nop, 0)" +by(simp add: fetch.simps) + +lemma mop_bef_length[simp]: "length (mop_bef n) = 4 * n" +apply(induct n, simp add: mop_bef.simps, simp add: mop_bef.simps) +done + +thm findnth_nth +lemma mop_bef_nth: + "\q < n; x < 4\ \ mop_bef n ! (4 * q + x) = + mop_bef (Suc q) ! ((4 * q) + x)" +apply(induct n, simp) +apply(case_tac "q < n", simp add: mop_bef.simps, auto) +apply(simp add: nth_append) +apply(subgoal_tac "q = n", simp) +apply(arith) +done + +lemma fetch_bef_erase_a_o[simp]: + "\0 < s; s \ 2 * n; s mod 2 = Suc 0\ + \ (fetch (mop_bef n @ tshift mp_up (2 * n)) s Oc) = (W0, s + 1)" +apply(subgoal_tac "\ q. s = 2*q + 1", auto) +apply(subgoal_tac "length (mop_bef n) = 4*n") +apply(auto simp: fetch.simps nth_of.simps nth_append) +apply(subgoal_tac "mop_bef n ! (4 * q + 1) = + mop_bef (Suc q) ! ((4 * q) + 1)", + simp add: mop_bef.simps nth_append) +apply(rule mop_bef_nth, auto) +done + +lemma fetch_bef_erase_a_b[simp]: + "\n < length lm; 0 < s; s \ 2 * n; s mod 2 = Suc 0\ + \ (fetch (mop_bef n @ tshift mp_up (2 * n)) s Bk) = (R, s + 2)" +apply(subgoal_tac "\ q. s = 2*q + 1", auto) +apply(subgoal_tac "length (mop_bef n) = 4*n") +apply(auto simp: fetch.simps nth_of.simps nth_append) +apply(subgoal_tac "mop_bef n ! (4 * q + 0) = + mop_bef (Suc q) ! ((4 * q + 0))", + simp add: mop_bef.simps nth_append) +apply(rule mop_bef_nth, auto) +done + +lemma fetch_bef_erase_b_b: + "\n < length lm; 0 < s; s \ 2 * n; s mod 2 = 0\ \ + (fetch (mop_bef n @ tshift mp_up (2 * n)) s Bk) = (R, s - 1)" +apply(subgoal_tac "\ q. s = 2 * q", auto) +apply(case_tac qa, simp, simp) +apply(auto simp: fetch.simps nth_of.simps nth_append) +apply(subgoal_tac "mop_bef n ! (4 * nat + 2) = + mop_bef (Suc nat) ! ((4 * nat) + 2)", + simp add: mop_bef.simps nth_append) +apply(rule mop_bef_nth, auto) +done + +lemma fetch_jump_over1_o: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (2 * n)) Oc + = (R, Suc (2 * n))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(auto simp: fetch.simps nth_of.simps mp_up_def nth_append + tshift.simps) +done + +lemma fetch_jump_over1_b: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (2 * n)) Bk + = (R, Suc (Suc (2 * n)))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(auto simp: fetch.simps nth_of.simps mp_up_def + nth_append tshift.simps) +done + +lemma fetch_aft_erase_a_o: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (Suc (2 * n))) Oc + = (W0, Suc (2 * n + 2))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(auto simp: fetch.simps nth_of.simps mp_up_def + nth_append tshift.simps) +done + +lemma fetch_aft_erase_a_b: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (Suc (2 * n))) Bk + = (L, Suc (2 * n + 4))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(auto simp: fetch.simps nth_of.simps mp_up_def + nth_append tshift.simps) +done + +lemma fetch_aft_erase_b_b: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (2*n + 3) Bk + = (R, Suc (2 * n + 3))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 3 = Suc (2*n + 2)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_aft_erase_c_o: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 4) Oc + = (W0, Suc (2 * n + 2))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_aft_erase_c_b: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 4) Bk + = (R, Suc (2 * n + 1))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_left_moving_o: + "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 5) Oc) + = (L, 2*n + 6)" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_left_moving_b: + "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 5) Bk) + = (L, 2*n + 5)" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_jump_over2_b: + "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 6) Bk) + = (R, 0)" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_jump_over2_o: +"(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 6) Oc) + = (L, 2*n + 6)" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemmas mopupfetchs = +fetch_bef_erase_a_o fetch_bef_erase_a_b fetch_bef_erase_b_b +fetch_jump_over1_o fetch_jump_over1_b fetch_aft_erase_a_o +fetch_aft_erase_a_b fetch_aft_erase_b_b fetch_aft_erase_c_o +fetch_aft_erase_c_b fetch_left_moving_o fetch_left_moving_b +fetch_jump_over2_b fetch_jump_over2_o + +lemma [simp]: +"\n < length lm; 0 < s; s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, Oc # xs) lm n ires; + Suc s \ 2 * n\ \ + mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires" +apply(auto simp: mopup_bef_erase_a.simps mopup_bef_erase_b.simps ) +apply(rule_tac x = "m - 1" in exI, rule_tac x = rn in exI) +apply(case_tac m, simp, simp) +done + +lemma mopup_false1: + "\0 < s; s \ 2 * n; s mod 2 = Suc 0; \ Suc s \ 2 * n\ + \ RR" +apply(arith) +done + +lemma [simp]: + "\n < length lm; 0 < s; s \ 2 * n; s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, Oc # xs) lm n ires; r = Oc # xs\ + \ (Suc s \ 2 * n \ mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires) \ + (\ Suc s \ 2 * n \ mopup_jump_over1 (Suc s, l, Bk # xs) lm n ires) " +apply(auto elim: mopup_false1) +done + +lemma drop_abc_lm_v_simp[simp]: + "n < length lm \ drop n lm = abc_lm_v lm n # drop (Suc n) lm" +apply(auto simp: abc_lm_v.simps) +apply(drule drop_Suc_conv_tl, simp) +done +lemma [simp]: "(\rna. Bk\<^bsup>rn\<^esup> = Bk # Bk\<^bsup>rna\<^esup>) \ Bk\<^bsup>rn\<^esup> = []" +apply(case_tac rn, simp, auto) +done + +lemma [simp]: "\lna. Bk # Bk\<^bsup>ln\<^esup> = Bk\<^bsup>lna\<^esup>" +apply(rule_tac x = "Suc ln" in exI, auto) +done + +lemma mopup_bef_erase_a_2_jump_over[simp]: + "\n < length lm; 0 < s; s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, Bk # xs) lm n ires; Suc s = 2 * n\ +\ mopup_jump_over1 (Suc (2 * n), Bk # l, xs) lm n ires" +apply(auto simp: mopup_bef_erase_a.simps mopup_jump_over1.simps) +apply(case_tac m, simp) +apply(rule_tac x = "Suc ln" in exI, rule_tac x = 0 in exI, + simp add: tape_of_nl_abv) +apply(case_tac "drop (Suc n) lm", auto simp: tape_of_nat_list.simps ) +done + +lemma Suc_Suc_div: "\0 < s; s mod 2 = Suc 0; Suc (Suc s) \ 2 * n\ + \ (Suc (Suc (s div 2))) \ n" +apply(arith) +done + +lemma mopup_bef_erase_a_2_a[simp]: + "\n < length lm; 0 < s; s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, Bk # xs) lm n ires; + Suc (Suc s) \ 2 * n\ \ + mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires" +apply(auto simp: mopup_bef_erase_a.simps ) +apply(subgoal_tac "drop (Suc (Suc (s div 2))) lm \ []") +apply(case_tac m, simp) +apply(rule_tac x = "Suc (abc_lm_v lm (Suc (s div 2)))" in exI, + rule_tac x = rn in exI, simp, simp) +apply(subgoal_tac "(Suc (Suc (s div 2))) \ n", simp, + rule_tac Suc_Suc_div, auto) +done + +lemma mopup_false2: + "\n < length lm; 0 < s; s \ 2 * n; + s mod 2 = Suc 0; Suc s \ 2 * n; + \ Suc (Suc s) \ 2 * n\ \ RR" +apply(arith) +done + +lemma [simp]: + "\n < length lm; 0 < s; s \ 2 * n; + s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, Bk # xs) lm n ires; + r = Bk # xs\ + \ (Suc s = 2 * n \ + mopup_jump_over1 (Suc (2 * n), Bk # l, xs) lm n ires) \ + (Suc s \ 2 * n \ + (Suc (Suc s) \ 2 * n \ + mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires) \ + (\ Suc (Suc s) \ 2 * n \ + mopup_aft_erase_a (Suc (Suc s), Bk # l, xs) lm n ires))" +apply(auto elim: mopup_false2) +done + +lemma [simp]: "mopup_bef_erase_a (s, l, []) lm n ires \ + mopup_bef_erase_a (s, l, [Bk]) lm n ires" +apply(auto simp: mopup_bef_erase_a.simps) +done + +lemma [simp]: + "\n < length lm; 0 < s; s \ 2 * n; s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, []) lm n ires; r = []\ + \ (Suc s = 2 * n \ + mopup_jump_over1 (Suc (2 * n), Bk # l, []) lm n ires) \ + (Suc s \ 2 * n \ + (Suc (Suc s) \ 2 * n \ + mopup_bef_erase_a (Suc (Suc s), Bk # l, []) lm n ires) \ + (\ Suc (Suc s) \ 2 * n \ + mopup_aft_erase_a (Suc (Suc s), Bk # l, []) lm n ires))" +apply(auto) +done + +lemma "mopup_bef_erase_b (s, l, Oc # xs) lm n ires \ l \ []" +apply(auto simp: mopup_bef_erase_b.simps) +done + +lemma [simp]: "mopup_bef_erase_b (s, l, Oc # xs) lm n ires = False" +apply(auto simp: mopup_bef_erase_b.simps ) +done + +lemma [simp]: "\0 < s; s \ 2 *n; s mod 2 \ Suc 0\ \ + (s - Suc 0) mod 2 = Suc 0" +apply(arith) +done + +lemma [simp]: "\0 < s; s \ 2 *n; s mod 2 \ Suc 0\ \ + s - Suc 0 \ 2 * n" +apply(simp) +done + +lemma [simp]: "\0 < s; s \ 2 *n; s mod 2 \ Suc 0\ \ \ s \ Suc 0" +apply(arith) +done + +lemma [simp]: "\n < length lm; 0 < s; s \ 2 * n; + s mod 2 \ Suc 0; + mopup_bef_erase_b (s, l, Bk # xs) lm n ires; r = Bk # xs\ + \ mopup_bef_erase_a (s - Suc 0, Bk # l, xs) lm n ires" +apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps) +done + +lemma [simp]: "\mopup_bef_erase_b (s, l, []) lm n ires\ \ + mopup_bef_erase_a (s - Suc 0, Bk # l, []) lm n ires" +apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps) +done + +lemma [simp]: + "\n < length lm; + mopup_jump_over1 (Suc (2 * n), l, Oc # xs) lm n ires; + r = Oc # xs\ + \ mopup_jump_over1 (Suc (2 * n), Oc # l, xs) lm n ires" +apply(auto simp: mopup_jump_over1.simps) +apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI, + rule_tac x = "m2 - 1" in exI) +apply(case_tac "m2", simp, simp, rule_tac x = rn in exI, simp) +apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI, + rule_tac x = "m2 - 1" in exI) +apply(case_tac m2, simp, simp) +done + +lemma mopup_jump_over1_2_aft_erase_a[simp]: + "\n < length lm; mopup_jump_over1 (Suc (2 * n), l, Bk # xs) lm n ires\ + \ mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires" +apply(simp only: mopup_jump_over1.simps mopup_aft_erase_a.simps) +apply(erule_tac exE)+ +apply(rule_tac x = ln in exI, rule_tac x = "Suc 0" in exI) +apply(case_tac m2, simp) +apply(rule_tac x = rn in exI, rule_tac x = "drop (Suc n) lm" in exI, + simp) +apply(simp) +done + +lemma [simp]: + "\n < length lm; mopup_jump_over1 (Suc (2 * n), l, []) lm n ires\ \ + mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires" +apply(rule mopup_jump_over1_2_aft_erase_a, simp) +apply(auto simp: mopup_jump_over1.simps) +apply(rule_tac x = ln in exI, rule_tac x = m1 in exI, + rule_tac x = m2 in exI, simp add: ) +apply(rule_tac x = 0 in exI, auto) +done + +lemma [simp]: + "\n < length lm; + mopup_aft_erase_a (Suc (Suc (2 * n)), l, Oc # xs) lm n ires\ + \ mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires" +apply(auto simp: mopup_aft_erase_a.simps mopup_aft_erase_b.simps ) +apply(case_tac ml, simp, case_tac rn, simp, simp) +apply(case_tac list, auto simp: tape_of_nl_abv + tape_of_nat_list.simps ) +apply(case_tac a, simp, rule_tac x = rn in exI, + rule_tac x = "[]" in exI, + simp add: tape_of_nat_list.simps, simp) +apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, + simp add: tape_of_nat_list.simps ) +apply(case_tac a, simp, rule_tac x = rn in exI, + rule_tac x = "aa # lista" in exI, simp, simp) +apply(rule_tac x = rn in exI, rule_tac x = "nat # aa # lista" in exI, + simp add: tape_of_nat_list.simps ) +done + +lemma [simp]: + "mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires \ l \ []" +apply(auto simp: mopup_aft_erase_a.simps) +done + +lemma [simp]: + "\n < length lm; + mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires\ + \ mopup_left_moving (5 + 2 * n, tl l, hd l # Bk # xs) lm n ires" +apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps) +apply(erule exE)+ +apply(case_tac lnr, simp) +apply(rule_tac x = lnl in exI, simp, rule_tac x = rn in exI, simp) +apply(subgoal_tac "ml = []", simp) +apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp, auto) +apply(subgoal_tac "ml = []", auto) +apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp) +done + +lemma [simp]: + "mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires \ l \ []" +apply(simp only: mopup_aft_erase_a.simps) +apply(erule exE)+ +apply(auto) +done + +lemma [simp]: + "\n < length lm; mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires\ + \ mopup_left_moving (5 + 2 * n, tl l, [hd l]) lm n ires" +apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps) +apply(erule exE)+ +apply(subgoal_tac "ml = [] \ rn = 0", erule conjE, erule conjE, simp) +apply(case_tac lnr, simp, rule_tac x = lnl in exI, simp, + rule_tac x = 0 in exI, simp) +apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, + rule_tac x = "Suc 0" in exI, simp) +apply(case_tac ml, simp, case_tac rn, simp, simp) +apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma [simp]: "mopup_aft_erase_b (2 * n + 3, l, Oc # xs) lm n ires = False" +apply(auto simp: mopup_aft_erase_b.simps ) +done + +lemma [simp]: + "\n < length lm; + mopup_aft_erase_c (2 * n + 4, l, Oc # xs) lm n ires\ + \ mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires" +apply(auto simp: mopup_aft_erase_c.simps mopup_aft_erase_b.simps ) +apply(case_tac ml, simp, case_tac rn, simp, simp, simp) +apply(case_tac list, auto simp: tape_of_nl_abv + tape_of_nat_list.simps tape_of_nat_abv ) +apply(case_tac a, rule_tac x = rn in exI, + rule_tac x = "[]" in exI, simp add: tape_of_nat_list.simps) +apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, + simp add: tape_of_nat_list.simps ) +apply(case_tac a, simp, rule_tac x = rn in exI, + rule_tac x = "aa # lista" in exI, simp) +apply(rule_tac x = rn in exI, rule_tac x = "nat # aa # lista" in exI, + simp add: tape_of_nat_list.simps ) +done + +lemma mopup_aft_erase_c_aft_erase_a[simp]: + "\n < length lm; mopup_aft_erase_c (2 * n + 4, l, Bk # xs) lm n ires\ + \ mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires" +apply(simp only: mopup_aft_erase_c.simps mopup_aft_erase_a.simps ) +apply(erule_tac exE)+ +apply(erule conjE, erule conjE, erule disjE) +apply(subgoal_tac "ml = []", simp, case_tac rn, + simp, simp, rule conjI) +apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp) +apply(rule_tac x = nat in exI, rule_tac x = "[]" in exI, simp) +apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp, simp, + rule conjI) +apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp) +apply(rule_tac x = rn in exI, rule_tac x = "ml" in exI, simp) +done + +lemma [simp]: + "\n < length lm; mopup_aft_erase_c (2 * n + 4, l, []) lm n ires\ + \ mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires" +apply(rule mopup_aft_erase_c_aft_erase_a, simp) +apply(simp only: mopup_aft_erase_c.simps) +apply(erule exE)+ +apply(rule_tac x = lnl in exI, rule_tac x = lnr in exI, simp add: ) +apply(rule_tac x = 0 in exI, rule_tac x = "[]" in exI, simp) +done + +lemma mopup_aft_erase_b_2_aft_erase_c[simp]: + "\n < length lm; mopup_aft_erase_b (2 * n + 3, l, Bk # xs) lm n ires\ + \ mopup_aft_erase_c (4 + 2 * n, Bk # l, xs) lm n ires" +apply(auto simp: mopup_aft_erase_b.simps mopup_aft_erase_c.simps) +apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp) +apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp) +done + +lemma [simp]: + "\n < length lm; mopup_aft_erase_b (2 * n + 3, l, []) lm n ires\ + \ mopup_aft_erase_c (4 + 2 * n, Bk # l, []) lm n ires" +apply(rule_tac mopup_aft_erase_b_2_aft_erase_c, simp) +apply(simp add: mopup_aft_erase_b.simps) +done + +lemma [simp]: + "mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires \ l \ []" +apply(auto simp: mopup_left_moving.simps) +done + +lemma [simp]: + "\n < length lm; mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires\ + \ mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires" +apply(simp only: mopup_left_moving.simps mopup_jump_over2.simps) +apply(erule_tac exE)+ +apply(erule conjE, erule disjE, erule conjE) +apply(case_tac rn, simp, simp add: ) +apply(case_tac "hd l", simp add: ) +apply(case_tac "abc_lm_v lm n", simp) +apply(rule_tac x = "lnl" in exI, rule_tac x = rn in exI, + rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI) +apply(case_tac lnl, simp, simp, simp add: exp_ind[THEN sym], simp) +apply(case_tac "abc_lm_v lm n", simp) +apply(case_tac lnl, simp, simp) +apply(rule_tac x = lnl in exI, rule_tac x = rn in exI) +apply(rule_tac x = nat in exI, rule_tac x = "Suc (Suc 0)" in exI, simp) +done + +lemma [simp]: "mopup_left_moving (2 * n + 5, l, xs) lm n ires \ l \ []" +apply(auto simp: mopup_left_moving.simps) +done + +lemma [simp]: + "\n < length lm; mopup_left_moving (2 * n + 5, l, Bk # xs) lm n ires\ + \ mopup_left_moving (2 * n + 5, tl l, hd l # Bk # xs) lm n ires" +apply(simp only: mopup_left_moving.simps) +apply(erule exE)+ +apply(case_tac lnr, simp) +apply(rule_tac x = lnl in exI, rule_tac x = 0 in exI, + rule_tac x = rn in exI, simp, simp) +apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, simp) +done + +lemma [simp]: +"\n < length lm; mopup_left_moving (2 * n + 5, l, []) lm n ires\ + \ mopup_left_moving (2 * n + 5, tl l, [hd l]) lm n ires" +apply(simp only: mopup_left_moving.simps) +apply(erule exE)+ +apply(case_tac lnr, simp) +apply(rule_tac x = lnl in exI, rule_tac x = 0 in exI, + rule_tac x = 0 in exI, simp, auto) +done + +lemma [simp]: + "mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires \ l \ []" +apply(auto simp: mopup_jump_over2.simps ) +done + +lemma [intro]: "\lna. Bk # Bk\<^bsup>ln\<^esup> = Bk\<^bsup>lna\<^esup> @ [Bk]" +apply(simp only: exp_ind[THEN sym], auto) +done + +lemma [simp]: +"\n < length lm; mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires\ + \ mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires" +apply(simp only: mopup_jump_over2.simps) +apply(erule_tac exE)+ +apply(simp add: , erule conjE, erule_tac conjE) +apply(case_tac m1, simp) +apply(rule_tac x = ln in exI, rule_tac x = rn in exI, + rule_tac x = 0 in exI, simp) +apply(case_tac ln, simp, simp, simp only: exp_ind[THEN sym], simp) +apply(rule_tac x = ln in exI, rule_tac x = rn in exI, + rule_tac x = nat in exI, rule_tac x = "Suc m2" in exI, simp) +done + +lemma [simp]: "\rna. Oc # Oc\<^bsup>a\<^esup> @ Bk\<^bsup>rn\<^esup> = @ Bk\<^bsup>rna\<^esup>" +apply(case_tac a, auto simp: tape_of_nat_abv ) +done + +lemma [simp]: + "\n < length lm; mopup_jump_over2 (2 * n + 6, l, Bk # xs) lm n ires\ + \ mopup_stop (0, Bk # l, xs) lm n ires" +apply(auto simp: mopup_jump_over2.simps mopup_stop.simps) +done + +lemma [simp]: "mopup_jump_over2 (2 * n + 6, l, []) lm n ires = False" +apply(simp only: mopup_jump_over2.simps, simp) +done + +lemma mopup_inv_step: + "\n < length lm; mopup_inv (s, l, r) lm n ires\ + \ mopup_inv (t_step (s, l, r) + ((mop_bef n @ tshift mp_up (2 * n)), 0)) lm n ires" +apply(auto split:if_splits simp add:t_step.simps, + tactic {* ALLGOALS (resolve_tac [@{thm "fetch_intro"}]) *}) +apply(simp_all add: mopupfetchs new_tape.simps) +done + +declare mopup_inv.simps[simp del] + +lemma mopup_inv_steps: +"\n < length lm; mopup_inv (s, l, r) lm n ires\ \ + mopup_inv (t_steps (s, l, r) + ((mop_bef n @ tshift mp_up (2 * n)), 0) stp) lm n ires" +apply(induct stp, simp add: t_steps.simps) +apply(simp add: t_steps_ind) +apply(case_tac "(t_steps (s, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)", simp) +apply(rule_tac mopup_inv_step, simp, simp) +done + +lemma [simp]: + "\n < length lm; Suc 0 \ n\ \ + mopup_bef_erase_a (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) lm n ires" +apply(auto simp: mopup_bef_erase_a.simps abc_lm_v.simps) +apply(case_tac lm, simp, case_tac list, simp, simp) +apply(rule_tac x = "Suc a" in exI, rule_tac x = rn in exI, simp) +done + +lemma [simp]: + "lm \ [] \ mopup_jump_over1 (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) lm 0 ires" +apply(auto simp: mopup_jump_over1.simps) +apply(rule_tac x = ln in exI, rule_tac x = 0 in exI, simp add: ) +apply(case_tac lm, simp, simp add: abc_lm_v.simps) +apply(case_tac rn, simp) +apply(case_tac list, rule_tac disjI2, + simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(rule_tac disjI1, + simp add: tape_of_nl_abv tape_of_nat_list.simps ) +apply(rule_tac disjI1, case_tac list, + simp add: tape_of_nl_abv tape_of_nat_list.simps, + rule_tac x = nat in exI, simp) +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps ) +done + +lemma mopup_init: + "\n < length lm; crsp_l ly (as, lm) (ac, l, r) ires\ \ + mopup_inv (Suc 0, l, r) lm n ires" +apply(auto simp: crsp_l.simps mopup_inv.simps) +apply(case_tac n, simp, auto simp: mopup_bef_erase_a.simps ) +apply(rule_tac x = "Suc (hd lm)" in exI, rule_tac x = rn in exI, simp) +apply(case_tac lm, simp, case_tac list, simp, case_tac lista, simp add: abc_lm_v.simps) +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) +apply(simp add: mopup_jump_over1.simps) +apply(rule_tac x = 0 in exI, rule_tac x = 0 in exI, auto) +apply(case_tac [!] n, simp_all) +apply(case_tac [!] lm, simp, case_tac list, simp) +apply(case_tac rn, simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) +apply(erule_tac x = nat in allE, simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) +apply(simp add: abc_lm_v.simps, auto) +apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) +apply(erule_tac x = rn in allE, simp_all) +done + +(***Begin: mopup stop***) +fun abc_mopup_stage1 :: "t_conf \ nat \ nat" + where + "abc_mopup_stage1 (s, l, r) n = + (if s > 0 \ s \ 2*n then 6 + else if s = 2*n + 1 then 4 + else if s \ 2*n + 2 \ s \ 2*n + 4 then 3 + else if s = 2*n + 5 then 2 + else if s = 2*n + 6 then 1 + else 0)" + +fun abc_mopup_stage2 :: "t_conf \ nat \ nat" + where + "abc_mopup_stage2 (s, l, r) n = + (if s > 0 \ s \ 2*n then length r + else if s = 2*n + 1 then length r + else if s = 2*n + 5 then length l + else if s = 2*n + 6 then length l + else if s \ 2*n + 2 \ s \ 2*n + 4 then length r + else 0)" + +fun abc_mopup_stage3 :: "t_conf \ nat \ nat" + where + "abc_mopup_stage3 (s, l, r) n = + (if s > 0 \ s \ 2*n then + if hd r = Bk then 0 + else 1 + else if s = 2*n + 2 then 1 + else if s = 2*n + 3 then 0 + else if s = 2*n + 4 then 2 + else 0)" + +fun abc_mopup_measure :: "(t_conf \ nat) \ (nat \ nat \ nat)" + where + "abc_mopup_measure (c, n) = + (abc_mopup_stage1 c n, abc_mopup_stage2 c n, + abc_mopup_stage3 c n)" + +definition abc_mopup_LE :: + "(((nat \ block list \ block list) \ nat) \ + ((nat \ block list \ block list) \ nat)) set" + where + "abc_mopup_LE \ (inv_image lex_triple abc_mopup_measure)" + +lemma wf_abc_mopup_le[intro]: "wf abc_mopup_LE" +by(auto intro:wf_inv_image wf_lex_triple simp:abc_mopup_LE_def) + +lemma [simp]: "mopup_bef_erase_a (a, aa, []) lm n ires = False" +apply(auto simp: mopup_bef_erase_a.simps) +done + +lemma [simp]: "mopup_bef_erase_b (a, aa, []) lm n ires = False" +apply(auto simp: mopup_bef_erase_b.simps) +done + +lemma [simp]: "mopup_aft_erase_b (2 * n + 3, aa, []) lm n ires = False" +apply(auto simp: mopup_aft_erase_b.simps) +done + +lemma mopup_halt_pre: + "\n < length lm; mopup_inv (Suc 0, l, r) lm n ires; wf abc_mopup_LE\ + \ \na. \ (\(s, l, r) n. s = 0) (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) na) n \ + ((t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) + (Suc na), n), + t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) + na, n) \ abc_mopup_LE" +apply(rule allI, rule impI, simp add: t_steps_ind) +apply(subgoal_tac "mopup_inv (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) na) lm n ires") +apply(case_tac "(t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) na)", simp) +proof - + fix na a b c + assume "n < length lm" "mopup_inv (a, b, c) lm n ires" "0 < a" + thus "((t_step (a, b, c) (mop_bef n @ tshift mp_up (2 * n), 0), n), + (a, b, c), n) \ abc_mopup_LE" + apply(auto split:if_splits simp add:t_step.simps mopup_inv.simps, + tactic {* ALLGOALS (resolve_tac [@{thm "fetch_intro"}]) *}) + apply(simp_all add: mopupfetchs new_tape.simps abc_mopup_LE_def + lex_triple_def lex_pair_def ) + done +next + fix na + assume "n < length lm" "mopup_inv (Suc 0, l, r) lm n ires" + thus "mopup_inv (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) na) lm n ires" + apply(rule mopup_inv_steps) + done +qed + +lemma mopup_halt: "\n < length lm; crsp_l ly (as, lm) (s, l, r) ires\ \ + \ stp. (\ (s, l, r). s = 0) (t_steps (Suc 0, l, r) + ((mop_bef n @ tshift mp_up (2 * n)), 0) stp)" +apply(subgoal_tac "mopup_inv (Suc 0, l, r) lm n ires") +apply(insert wf_abc_mopup_le) +apply(insert halt_lemma[of abc_mopup_LE + "\ ((s, l, r), n). s = 0" + "\ stp. (t_steps (Suc 0, l, r) ((mop_bef n @ tshift mp_up (2 * n)) + , 0) stp, n)"], auto) +apply(insert mopup_halt_pre[of n lm l r], simp, erule exE) +apply(rule_tac x = na in exI, case_tac "(t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) na)", simp) +apply(rule_tac mopup_init, auto) +done +(***End: mopup stop****) +(* +lemma mopup_stop_cond: "mopup_inv (0, l, r) lm n ires \ + (\ln rn. ?l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ?ires \ ?r = @ Bk\<^bsup>rn\<^esup>) " + t_halt_conf (0, l, r) \ t_result r = Suc (abc_lm_v lm n)" +apply(simp add: mopup_inv.simps mopup_stop.simps t_halt_conf.simps + t_result.simps, auto simp: tape_of_nat_abv) +apply(rule_tac x = rn in exI, + rule_tac x = "Suc (abc_lm_v lm n)" in exI, + simp add: tape_of_nat_abv) +apply(simp add: tape_of_nat_abv exponent_def) +apply(subgoal_tac "takeWhile (\a. a = Oc) + (replicate (abc_lm_v lm n) Oc @ replicate rn Bk) + = replicate (abc_lm_v lm n) Oc @ takeWhile (\a. a = Oc) + (replicate rn Bk)", simp) +apply(case_tac rn, simp, simp) +apply(rule takeWhile_append2) +apply(case_tac x, auto) +done +*) + + +lemma mopup_halt_conf_pre: + "\n < length lm; crsp_l ly (as, lm) (s, l, r) ires\ + \ \ na. (\ (s', l', r'). s' = 0 \ mopup_stop (s', l', r') lm n ires) + (t_steps (Suc 0, l, r) + ((mop_bef n @ tshift mp_up (2 * n)), 0) na)" +apply(subgoal_tac "\ stp. (\ (s, l, r). s = 0) + (t_steps (Suc 0, l, r) ((mop_bef n @ tshift mp_up (2 * n)), 0) stp)", + erule exE) +apply(rule_tac x = stp in exI, + case_tac "(t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)", simp) +apply(subgoal_tac " mopup_inv (Suc 0, l, r) lm n ires") +apply(subgoal_tac "mopup_inv (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp) lm n ires", simp) +apply(simp only: mopup_inv.simps) +apply(rule_tac mopup_inv_steps, simp, simp) +apply(rule_tac mopup_init, simp, simp) +apply(rule_tac mopup_halt, simp, simp) +done + +thm mopup_stop.simps + +lemma mopup_halt_conf: + assumes len: "n < length lm" + and correspond: "crsp_l ly (as, lm) (s, l, r) ires" + shows + "\ na. (\ (s', l', r'). ((\ln rn. s' = 0 \ l' = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \ r' = Oc\<^bsup>Suc (abc_lm_v lm n)\<^esup> @ Bk\<^bsup>rn\<^esup>))) + (t_steps (Suc 0, l, r) + ((mop_bef n @ tshift mp_up (2 * n)), 0) na)" +using len correspond mopup_halt_conf_pre[of n lm ly as s l r ires] +apply(simp add: mopup_stop.simps tape_of_nat_abv tape_of_nat_list.simps) +done +(*********End: mop_up****************************) + + +subsubsection {* Final results about Abacus machine *} + +thm mopup_halt +lemma mopup_halt_bef: "\n < length lm; crsp_l ly (as, lm) (s, l, r) ires\ + \ \stp. (\(s, l, r). s \ 0 \ ((\ (s', l', r'). s' = 0) + (t_step (s, l, r) (mop_bef n @ tshift mp_up (2 * n), 0)))) + (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp)" +apply(insert mopup_halt[of n lm ly as s l r ires], simp, erule_tac exE) +proof - + fix stp + assume "n < length lm" + "crsp_l ly (as, lm) (s, l, r) ires" + "(\(s, l, r). s = 0) + (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)" + thus "\stp. (\(s, ab). 0 < s \ (\(s', l', r'). s' = 0) + (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) + (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp)" + proof(induct stp, simp add: t_steps.simps, simp) + fix stpa + assume h1: + "(\(s, l, r). s = 0) (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stpa) \ + \stp. (\(s, ab). 0 < s \ (\(s', l', r'). s' = 0) + (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) + (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)" + and h2: + "(\(s, l, r). s = 0) (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) (Suc stpa))" + "n < length lm" + "crsp_l ly (as, lm) (s, l, r) ires" + thus "\stp. (\(s, ab). 0 < s \ (\(s', l', r'). s' = 0) + (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) ( + t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)" + apply(case_tac "(\(s, l, r). s = 0) (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", + simp) + apply(rule_tac x = "stpa" in exI) + apply(case_tac "(t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", + simp add: t_steps_ind) + done + qed +qed + +lemma tshift_nth_state0: "\n < length tp; tp ! n = (a, 0)\ + \ tshift tp off ! n = (a, 0)" +apply(induct n, case_tac tp, simp) +apply(auto simp: tshift.simps) +done + +lemma shift_length: "length (tshift tp n) = length tp" +apply(auto simp: tshift.simps) +done + +lemma even_Suc_le: "\y mod 2 = 0; 2 * x < y\ \ Suc (2 * x) < y" +by arith + +lemma [simp]: "(4::nat) * n mod 2 = 0" +by arith + +lemma tm_append_fetch_equal: + "\t_ncorrect (tm_of aprog); s'> 0; + fetch (mop_bef n @ tshift mp_up (2 * n)) s' b = (a, 0)\ +\ fetch (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (length (tm_of aprog) div 2)) (s' + length (tm_of aprog) div 2) b + = (a, 0)" +apply(case_tac s', simp) +apply(auto simp: fetch.simps nth_of.simps t_ncorrect.simps shift_length nth_append + tshift.simps split: list.splits block.splits split: if_splits) +done + +lemma [simp]: + "\t_ncorrect (tm_of aprog); + t_step (s', l', r') (mop_bef n @ tshift mp_up (2 * n), 0) = + (0, l'', r''); s' > 0\ + \ t_step (s' + length (tm_of aprog) div 2, l', r') + (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (length (tm_of aprog) div 2), 0) = (0, l'', r'')" +apply(simp add: t_step.simps) +apply(subgoal_tac + "(fetch (mop_bef n @ tshift mp_up (2 * n)) s' + (case r' of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc)) + = (fetch (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (length (tm_of aprog) div 2)) (s' + length (tm_of aprog) div 2) + (case r' of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc))", simp) +apply(case_tac "(fetch (mop_bef n @ tshift mp_up (2 * n)) s' + (case r' of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc))", simp) +apply(drule_tac tm_append_fetch_equal, auto) +done + +lemma [intro]: + "start_of (layout_of aprog) (length aprog) - Suc 0 = + length (tm_of aprog) div 2" +apply(subgoal_tac "abc2t_correct aprog") +apply(insert pre_lheq[of "concat (take (length aprog) + (tms_of aprog))" "length aprog" aprog], simp add: tm_of.simps) +by auto + +lemma tm_append_stop_step: + "\t_ncorrect (tm_of aprog); + t_ncorrect (mop_bef n @ tshift mp_up (2 * n)); n < length lm; + (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp) = + (s', l', r'); + s' \ 0; + t_step (s', l', r') (mop_bef n @ tshift mp_up (2 * n), 0) + = (0, l'', r'')\ + \ +(t_steps ((start_of (layout_of aprog) (length aprog), l, r)) + (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (start_of (layout_of aprog) (length aprog) - Suc 0), 0) (Suc stp)) + = (0, l'', r'')" +apply(insert tm_append_steps_equal[of "tm_of aprog" + "(mop_bef n @ tshift mp_up (2 * n))" + "(start_of (layout_of aprog) (length aprog) - Suc 0)" + "Suc 0" l r stp], simp) +apply(subgoal_tac "(start_of (layout_of aprog) (length aprog) - Suc 0) + = (length (tm_of aprog) div 2)", simp add: t_steps_ind) +apply(case_tac + "(t_steps (start_of (layout_of aprog) (length aprog), l, r) + (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (length (tm_of aprog) div 2), 0) stp)", simp) +apply(subgoal_tac "start_of (layout_of aprog) (length aprog) > 0", + case_tac "start_of (layout_of aprog) (length aprog)", + simp, simp) +apply(rule startof_not0, auto) +done + +(* +lemma stop_conf: "mopup_inv (0, aca, bc) am n + \ t_halt_conf (0, aca, bc) \ t_result bc = Suc (abc_lm_v am n)" +apply(case_tac n, + auto simp: mopup_inv.simps mopup_stop.simps t_halt_conf.simps + t_result.simps tape_of_nl_abv tape_of_nat_abv ) +apply(rule_tac x = "rn" in exI, + rule_tac x = "Suc (abc_lm_v am 0)" in exI, simp) +apply(subgoal_tac "takeWhile (\a. a = Oc) (Oc\<^bsup>abc_lm_v am 0\<^esup> @ Bk\<^bsup>rn\<^esup>) + = Oc\<^bsup>abc_lm_v am 0\<^esup> @ takeWhile (\a. a = Oc) (Bk\<^bsup>rn\<^esup>)", simp) +apply(simp add: exponent_def, case_tac rn, simp, simp) +apply(rule_tac takeWhile_append2, simp add: exponent_def) +apply(rule_tac x = rn in exI, + rule_tac x = "Suc (abc_lm_v am (Suc nat))" in exI, simp) +apply(subgoal_tac + "takeWhile (\a. a = Oc) (Oc\<^bsup>abc_lm_v am (Suc nat)\<^esup> @ Bk\<^bsup>rn\<^esup>) = + Oc\<^bsup>abc_lm_v am (Suc nat)\<^esup> @ takeWhile (\a. a = Oc) (Bk\<^bsup>rn\<^esup>)", simp) +apply(simp add: exponent_def, case_tac rn, simp, simp) +apply(rule_tac takeWhile_append2, simp add: exponent_def) +done +*) + + +lemma start_of_out_range: +"as \ length aprog \ + start_of (layout_of aprog) as = + start_of (layout_of aprog) (length aprog)" +apply(induct as, simp) +apply(case_tac "length aprog = Suc as", simp) +apply(simp add: start_of.simps) +done + +lemma [intro]: "t_ncorrect (tm_of aprog)" +apply(simp add: tm_of.simps) +apply(insert tms_mod2[of "length aprog" aprog], + simp add: t_ncorrect.simps) +done + +lemma abacus_turing_eq_halt_case_pre: + "\ly = layout_of aprog; + tprog = tm_of aprog; + crsp_l ly ac tc ires; + n < length am; + abc_steps_l ac aprog stp = (as, am); + mop_ss = start_of ly (length aprog); + as \ length aprog\ + \ \ stp. (\ (s, l, r). s = 0 \ mopup_inv (0, l, r) am n ires) + (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)" +apply(insert steps_crsp[of ly aprog tprog ac tc ires "(as, am)" stp], auto) +apply(case_tac "(t_steps tc (tm_of aprog, 0) n')", + simp add: tMp.simps) +apply(subgoal_tac "t_ncorrect (mop_bef n @ tshift mp_up (2 * n))") +proof - + fix n' a b c + assume h1: + "crsp_l (layout_of aprog) ac tc ires" + "abc_steps_l ac aprog stp = (as, am)" + "length aprog \ as" + "crsp_l (layout_of aprog) (as, am) (a, b, c) ires" + "t_steps tc (tm_of aprog, 0) n' = (a, b, c)" + "n < length am" + "t_ncorrect (mop_bef n @ tshift mp_up (2 * n))" + hence h2: + "t_steps tc (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (start_of (layout_of aprog) (length aprog) - Suc 0), 0) n' + = (a, b, c)" + apply(rule_tac tm_append_steps, simp) + apply(simp add: crsp_l.simps, auto) + apply(simp add: crsp_l.simps) + apply(rule startof_not0) + done + from h1 have h3: + "\stp. (\(s, l, r). s \ 0 \ ((\ (s', l', r'). s' = 0) + (t_step (s, l, r) (mop_bef n @ tshift mp_up (2 * n), 0)))) + (t_steps (Suc 0, b, c) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)" + apply(rule_tac mopup_halt_bef, auto) + done + from h1 and h2 and h3 show + "\stp. case t_steps tc (tm_of aprog @ abacus.tshift (mop_bef n @ abacus.tshift mp_up (2 * n)) + (start_of (layout_of aprog) (length aprog) - Suc 0), 0) stp of (s, ab) + \ s = 0 \ mopup_inv (0, ab) am n ires" + proof(erule_tac exE, + case_tac "(t_steps (Suc 0, b, c) + (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", simp, + case_tac "(t_step (aa, ba, ca) + (mop_bef n @ tshift mp_up (2 * n), 0))", simp) + fix stpa aa ba ca aaa baa caa + assume g1: "0 < aa \ aaa = 0" + "t_steps (Suc 0, b, c) + (mop_bef n @ tshift mp_up (2 * n), 0) stpa = (aa, ba,ca)" + "t_step (aa, ba, ca) (mop_bef n @ tshift mp_up (2 * n), 0) + = (0, baa, caa)" + from h1 and this have g2: + "t_steps (start_of (layout_of aprog) (length aprog), b, c) + (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (start_of (layout_of aprog) (length aprog) - Suc 0), 0) + (Suc stpa) = (0, baa, caa)" + apply(rule_tac tm_append_stop_step, auto) + done + from h1 and h2 and g1 and this show "?thesis" + apply(rule_tac x = "n' + Suc stpa" in exI) + apply(simp add: t_steps_ind del: t_steps.simps) + apply(subgoal_tac "a = start_of (layout_of aprog) + (length aprog)", simp) + apply(insert mopup_inv_steps[of n am "Suc 0" b c ires "Suc stpa"], + simp add: t_steps_ind) + apply(subgoal_tac "mopup_inv (Suc 0, b, c) am n ires", simp) + apply(rule_tac mopup_init, simp, simp) + apply(simp add: crsp_l.simps) + apply(erule_tac start_of_out_range) + done + qed +next + show " t_ncorrect (mop_bef n @ tshift mp_up (2 * n))" + apply(auto simp: t_ncorrect.simps tshift.simps mp_up_def) + done +qed + +text {* + One of the main theorems about Abacus compilation. +*} + +lemma abacus_turing_eq_halt_case: + assumes layout: + -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} + "ly = layout_of aprog" + and complied: + -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} + "tprog = tm_of aprog" + and correspond: + -- {* TM configuration @{text "tc"} and Abacus configuration @{text "ac"} + are in correspondence: *} + "crsp_l ly ac tc ires" + and halt_state: + -- {* @{text "as"} is a program label outside the range of @{text "aprog"}. So + if Abacus is in such a state, it is in halt state: *} + "as \ length aprog" + and abc_exec: + -- {* Supposing after @{text "stp"} step of execution, Abacus program @{text "aprog"} + reaches such a halt state: *} + "abc_steps_l ac aprog stp = (as, am)" + and rs_len: + -- {* @{text "n"} is a memory address in the range of Abacus memory @{text "am"}: *} + "n < length am" + and mopup_start: + -- {* The startling label for mopup mahines, according to the layout and Abacus program + should be @{text "mop_ss"}: *} + "mop_ss = start_of ly (length aprog)" + shows + -- {* + After @{text "stp"} steps of execution of the TM composed of @{text "tprog"} and the mopup + TM @{text "(tMp n (mop_ss - 1))"} will halt and gives rise to a configuration which + only hold the content of memory cell @{text "n"}: + *} + "\ stp. (\ (s, l, r). \ ln rn. s = 0 \ l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \ r = Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>rn\<^esup>) + (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)" +proof - + from layout complied correspond halt_state abc_exec rs_len mopup_start + and abacus_turing_eq_halt_case_pre [of ly aprog tprog ac tc ires n am stp as mop_ss] + show "?thesis" + apply(simp add: mopup_inv.simps mopup_stop.simps tape_of_nat_abv) + done +qed + +lemma abc_unhalt_case_zero: +"\crsp_l (layout_of aprog) ac tc ires; + n < length am; + \stp. (\(as, am). as < length aprog) (abc_steps_l ac aprog stp)\ + \ \astp bstp. 0 \ bstp \ + crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) + (t_steps tc (tm_of aprog, 0) bstp) ires" +apply(rule_tac x = "Suc 0" in exI) +apply(case_tac " abc_steps_l ac aprog (Suc 0)", simp) +proof - + fix a b + assume "crsp_l (layout_of aprog) ac tc ires" + "abc_steps_l ac aprog (Suc 0) = (a, b)" + thus "\bstp. crsp_l (layout_of aprog) (a, b) + (t_steps tc (tm_of aprog, 0) bstp) ires" + apply(insert steps_crsp[of "layout_of aprog" aprog + "tm_of aprog" ac tc ires "(a, b)" "Suc 0"], auto) + done +qed + +declare abc_steps_l.simps[simp del] + +lemma abc_steps_ind: + "let (as, am) = abc_steps_l ac aprog stp in + abc_steps_l ac aprog (Suc stp) = + abc_step_l (as, am) (abc_fetch as aprog) " +proof(simp) + show "(\(as, am). abc_steps_l ac aprog (Suc stp) = + abc_step_l (as, am) (abc_fetch as aprog)) + (abc_steps_l ac aprog stp)" + proof(induct stp arbitrary: ac) + fix ac + show "(\(as, am). abc_steps_l ac aprog (Suc 0) = + abc_step_l (as, am) (abc_fetch as aprog)) + (abc_steps_l ac aprog 0)" + apply(case_tac "ac:: nat \ nat list", + simp add: abc_steps_l.simps) + apply(case_tac "(abc_step_l (a, b) (abc_fetch a aprog))", + simp add: abc_steps_l.simps) + done + next + fix stp ac + assume h1: + "(\ac. (\(as, am). abc_steps_l ac aprog (Suc stp) = + abc_step_l (as, am) (abc_fetch as aprog)) + (abc_steps_l ac aprog stp))" + thus + "(\(as, am). abc_steps_l ac aprog (Suc (Suc stp)) = + abc_step_l (as, am) (abc_fetch as aprog)) + (abc_steps_l ac aprog (Suc stp))" + apply(case_tac "ac::nat \ nat list", simp) + apply(subgoal_tac + "abc_steps_l (a, b) aprog (Suc (Suc stp)) = + abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) + aprog (Suc stp)", simp) + apply(case_tac "(abc_step_l (a, b) (abc_fetch a aprog))", simp) + proof - + fix a b aa ba + assume h2: "abc_step_l (a, b) (abc_fetch a aprog) = (aa, ba)" + from h1 and h2 show + "(\(as, am). abc_steps_l (aa, ba) aprog (Suc stp) = + abc_step_l (as, am) (abc_fetch as aprog)) + (abc_steps_l (a, b) aprog (Suc stp))" + apply(insert h1[of "(aa, ba)"]) + apply(simp add: abc_steps_l.simps) + apply(insert h2, simp) + done + next + fix a b + show + "abc_steps_l (a, b) aprog (Suc (Suc stp)) = + abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) + aprog (Suc stp)" + apply(simp only: abc_steps_l.simps) + done + qed + qed +qed + +lemma abc_unhalt_case_induct: + "\crsp_l (layout_of aprog) ac tc ires; + n < length am; + \stp. (\(as, am). as < length aprog) (abc_steps_l ac aprog stp); + stp \ bstp; + crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) + (t_steps tc (tm_of aprog, 0) bstp) ires\ + \ \astp bstp. Suc stp \ bstp \ crsp_l (layout_of aprog) + (abc_steps_l ac aprog astp) (t_steps tc (tm_of aprog, 0) bstp) ires" +apply(rule_tac x = "Suc astp" in exI) +apply(case_tac "abc_steps_l ac aprog astp") +proof - + fix a b + assume + "\stp. (\(as, am). as < length aprog) + (abc_steps_l ac aprog stp)" + "stp \ bstp" + "crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) + (t_steps tc (tm_of aprog, 0) bstp) ires" + "abc_steps_l ac aprog astp = (a, b)" + thus + "\bstp\Suc stp. crsp_l (layout_of aprog) + (abc_steps_l ac aprog (Suc astp)) + (t_steps tc (tm_of aprog, 0) bstp) ires" + apply(insert crsp_inside[of "layout_of aprog" aprog + "tm_of aprog" a b "(t_steps tc (tm_of aprog, 0) bstp)" "ires"], auto) + apply(erule_tac x = astp in allE, auto) + apply(rule_tac x = "bstp + stpa" in exI, simp) + apply(insert abc_steps_ind[of ac aprog "astp"], simp) + done +qed + +lemma abc_unhalt_case: + "\crsp_l (layout_of aprog) ac tc ires; + \stp. (\(as, am). as < length aprog) (abc_steps_l ac aprog stp)\ + \ (\ astp bstp. bstp \ stp \ + crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) + (t_steps tc (tm_of aprog, 0) bstp) ires)" +apply(induct stp) +apply(rule_tac abc_unhalt_case_zero, auto) +apply(rule_tac abc_unhalt_case_induct, auto) +done + +lemma abacus_turing_eq_unhalt_case_pre: + "\ly = layout_of aprog; + tprog = tm_of aprog; + crsp_l ly ac tc ires; + \ stp. ((\ (as, am). as < length aprog) + (abc_steps_l ac aprog stp)); + mop_ss = start_of ly (length aprog)\ + \ (\ (\ stp. (\ (s, l, r). s = 0) + (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)))" + apply(auto) +proof - + fix stp a b + assume h1: + "crsp_l (layout_of aprog) ac tc ires" + "\stp. (\(as, am). as < length aprog) (abc_steps_l ac aprog stp)" + "t_steps tc (tm_of aprog @ tMp n (start_of (layout_of aprog) + (length aprog) - Suc 0), 0) stp = (0, a, b)" + thus "False" + proof(insert abc_unhalt_case[of aprog ac tc ires stp], auto, + case_tac "(abc_steps_l ac aprog astp)", + case_tac "(t_steps tc (tm_of aprog, 0) bstp)", simp) + fix astp bstp aa ba aaa baa c + assume h2: + "abc_steps_l ac aprog astp = (aa, ba)" "stp \ bstp" + "t_steps tc (tm_of aprog, 0) bstp = (aaa, baa, c)" + "crsp_l (layout_of aprog) (aa, ba) (aaa, baa, c) ires" + hence h3: + "t_steps tc (tm_of aprog @ tMp n + (start_of (layout_of aprog) (length aprog) - Suc 0), 0) bstp + = (aaa, baa, c)" + apply(intro tm_append_steps, auto) + apply(simp add: crsp_l.simps, rule startof_not0) + done + from h2 have h4: "\ diff. bstp = stp + diff" + apply(rule_tac x = "bstp - stp" in exI, simp) + done + from h4 and h3 and h2 and h1 show "?thesis" + apply(auto) + apply(simp add: state0_ind crsp_l.simps) + apply(subgoal_tac "start_of (layout_of aprog) aa > 0", simp) + apply(rule startof_not0) + done + qed +qed + +lemma abacus_turing_eq_unhalt_case: + assumes layout: + -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} + "ly = layout_of aprog" + and compiled: + -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} + "tprog = tm_of aprog" + and correspond: + -- {* + TM configuration @{text "tc"} and Abacus configuration @{text "ac"} + are in correspondence: + *} + "crsp_l ly ac tc ires" + and abc_unhalt: + -- {* + If, no matter how many steps the Abacus program @{text "aprog"} executes, it + may never reach a halt state. + *} + "\ stp. ((\ (as, am). as < length aprog) + (abc_steps_l ac aprog stp))" + and mopup_start: "mop_ss = start_of ly (length aprog)" + shows + -- {* + The the TM composed of TM @{text "tprog"} and the moupup TM may never reach a halt state as well. + *} + "\ (\ stp. (\ (s, l, r). s = 0) + (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp))" + using layout compiled correspond abc_unhalt mopup_start + apply(rule_tac abacus_turing_eq_unhalt_case_pre, auto) + done + +definition abc_list_crsp:: "nat list \ nat list \ bool" + where + "abc_list_crsp xs ys = (\ n. xs = ys @ 0\<^bsup>n\<^esup> \ ys = xs @ 0\<^bsup>n\<^esup>)" +lemma [intro]: "abc_list_crsp (lm @ 0\<^bsup>m\<^esup>) lm" +apply(auto simp: abc_list_crsp_def) +done + +thm abc_lm_v.simps +lemma abc_list_crsp_lm_v: + "abc_list_crsp lma lmb \ abc_lm_v lma n = abc_lm_v lmb n" +apply(auto simp: abc_list_crsp_def abc_lm_v.simps + nth_append exponent_def) +done + +lemma rep_app_cons_iff: + "k < n \ replicate n a[k:=b] = + replicate k a @ b # replicate (n - k - 1) a" +apply(induct n arbitrary: k, simp) +apply(simp split:nat.splits) +done + +lemma abc_list_crsp_lm_s: + "abc_list_crsp lma lmb \ + abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)" +apply(auto simp: abc_list_crsp_def abc_lm_v.simps abc_lm_s.simps) +apply(simp_all add: list_update_append, auto simp: exponent_def) +proof - + fix na + assume h: "m < length lmb + na" " \ m < length lmb" + hence "m - length lmb < na" by simp + hence "replicate na 0[(m- length lmb):= n] = + replicate (m - length lmb) 0 @ n # + replicate (na - (m - length lmb) - 1) 0" + apply(erule_tac rep_app_cons_iff) + done + thus "\nb. replicate na 0[m - length lmb := n] = + replicate (m - length lmb) 0 @ n # replicate nb 0 \ + replicate (m - length lmb) 0 @ [n] = + replicate na 0[m - length lmb := n] @ replicate nb 0" + apply(auto) + done +next + fix na + assume h: "\ m < length lmb + na" + show + "\nb. replicate na 0 @ replicate (m - (length lmb + na)) 0 @ [n] = + replicate (m - length lmb) 0 @ n # replicate nb 0 \ + replicate (m - length lmb) 0 @ [n] = + replicate na 0 @ + replicate (m - (length lmb + na)) 0 @ n # replicate nb 0" + apply(rule_tac x = 0 in exI, simp, auto) + using h + apply(simp add: replicate_add[THEN sym]) + done +next + fix na + assume h: "\ m < length lma" "m < length lma + na" + hence "m - length lma < na" by simp + hence + "replicate na 0[(m- length lma):= n] = replicate (m - length lma) + 0 @ n # replicate (na - (m - length lma) - 1) 0" + apply(erule_tac rep_app_cons_iff) + done + thus "\nb. replicate (m - length lma) 0 @ [n] = + replicate na 0[m - length lma := n] @ replicate nb 0 + \ replicate na 0[m - length lma := n] = + replicate (m - length lma) 0 @ n # replicate nb 0" + apply(auto) + done +next + fix na + assume "\ m < length lma + na" + thus " \nb. replicate (m - length lma) 0 @ [n] = + replicate na 0 @ + replicate (m - (length lma + na)) 0 @ n # replicate nb 0 + \ replicate na 0 @ + replicate (m - (length lma + na)) 0 @ [n] = + replicate (m - length lma) 0 @ n # replicate nb 0" + apply(rule_tac x = 0 in exI, simp, auto) + apply(simp add: replicate_add[THEN sym]) + done +qed + +lemma abc_list_crsp_step: + "\abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); + abc_step_l (aa, lmb) i = (a', lmb')\ + \ a' = a \ abc_list_crsp lma' lmb'" +apply(case_tac i, auto simp: abc_step_l.simps + abc_list_crsp_lm_s abc_list_crsp_lm_v Let_def + split: abc_inst.splits if_splits) +done + +thm abc_step_l.simps + +lemma abc_steps_red: + "abc_steps_l ac aprog stp = (as, am) \ + abc_steps_l ac aprog (Suc stp) = + abc_step_l (as, am) (abc_fetch as aprog)" +using abc_steps_ind[of ac aprog stp] +apply(simp) +done + +lemma abc_list_crsp_steps: + "\abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp = (a, lm'); aprog \ []\ + \ \ lma. abc_steps_l (0, lm) aprog stp = (a, lma) \ + abc_list_crsp lm' lma" +apply(induct stp arbitrary: a lm', simp add: abc_steps_l.simps, auto) +apply(case_tac "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp", + simp add: abc_steps_ind) +proof - + fix stp a lm' aa b + assume ind: + "\a lm'. aa = a \ b = lm' \ + \lma. abc_steps_l (0, lm) aprog stp = (a, lma) \ + abc_list_crsp lm' lma" + and h: "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog (Suc stp) = (a, lm')" + "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp = (aa, b)" + "aprog \ []" + hence g1: "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog (Suc stp) + = abc_step_l (aa, b) (abc_fetch aa aprog)" + apply(rule_tac abc_steps_red, simp) + done + have "\lma. abc_steps_l (0, lm) aprog stp = (aa, lma) \ + abc_list_crsp b lma" + apply(rule_tac ind, simp) + done + from this obtain lma where g2: + "abc_steps_l (0, lm) aprog stp = (aa, lma) \ + abc_list_crsp b lma" .. + hence g3: "abc_steps_l (0, lm) aprog (Suc stp) + = abc_step_l (aa, lma) (abc_fetch aa aprog)" + apply(rule_tac abc_steps_red, simp) + done + show "\lma. abc_steps_l (0, lm) aprog (Suc stp) = (a, lma) \ + abc_list_crsp lm' lma" + using g1 g2 g3 h + apply(auto) + apply(case_tac "abc_step_l (aa, b) (abc_fetch aa aprog)", + case_tac "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp) + apply(rule_tac abc_list_crsp_step, auto) + done +qed + +text {* Begin: equvilence between steps and t_steps*} +lemma [simp]: "(case ca of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc) = + (case ca of [] \ Bk | x # xs \ x)" +by(case_tac ca, simp_all, case_tac a, simp, simp) + +text {* needed to interpret*} + +lemma steps_eq: "length t mod 2 = 0 \ + t_steps c (t, 0) stp = steps c t stp" +apply(induct stp) +apply(simp add: steps.simps t_steps.simps) +apply(simp add:tstep_red t_steps_ind) +apply(case_tac "steps c t stp", simp) +apply(auto simp: t_step.simps tstep.simps) +done + +text{* end: equvilence between steps and t_steps*} + +lemma crsp_l_start: "crsp_l ly (0, lm) (Suc 0, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) ires" +apply(simp add: crsp_l.simps, auto simp: start_of.simps) +done + +lemma t_ncorrect_app: "\t_ncorrect t1; t_ncorrect t2\ \ + t_ncorrect (t1 @ t2)" +apply(simp add: t_ncorrect.simps, auto) +done + +lemma [simp]: + "(length (tm_of aprog) + + length (tMp n (start_of ly (length aprog) - Suc 0))) mod 2 = 0" +apply(subgoal_tac + "t_ncorrect (tm_of aprog @ tMp n + (start_of ly (length aprog) - Suc 0))") +apply(simp add: t_ncorrect.simps) +apply(rule_tac t_ncorrect_app, + auto simp: tMp.simps t_ncorrect.simps tshift.simps mp_up_def) +apply(subgoal_tac + "t_ncorrect (tm_of aprog)", simp add: t_ncorrect.simps) +apply(auto) +done + +lemma [simp]: "takeWhile (\a. a = Oc) + (replicate rs Oc @ replicate rn Bk) = replicate rs Oc" +apply(induct rs, auto) +apply(induct rn, auto) +done + +lemma abacus_turing_eq_halt': + "\ly = layout_of aprog; + tprog = tm_of aprog; + n < length am; + abc_steps_l (0, lm) aprog stp = (as, am); + mop_ss = start_of ly (length aprog); + as \ length aprog\ + \ \ stp m l. steps (Suc 0, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) + (tprog @ (tMp n (mop_ss - 1))) stp + = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)" +apply(drule_tac tc = "(Suc 0, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>)" in + abacus_turing_eq_halt_case, auto intro: crsp_l_start) +apply(subgoal_tac + "length (tm_of aprog @ tMp n + (start_of ly (length aprog) - Suc 0)) mod 2 = 0") +apply(simp add: steps_eq) +apply(rule_tac x = stpa in exI, + simp add: exponent_def, auto) +done + + +thm tinres_steps +lemma list_length: "xs = ys \ length xs = length ys" +by simp +lemma [elim]: "tinres (Bk\<^bsup>m\<^esup>) b \ \m. b = Bk\<^bsup>m\<^esup>" +apply(auto simp: tinres_def) +apply(rule_tac x = "m-n" in exI, + auto simp: exponent_def replicate_add[THEN sym]) +apply(case_tac "m < n", auto) +apply(drule_tac list_length, auto) +apply(subgoal_tac "\ d. m = d + n", auto simp: replicate_add) +apply(rule_tac x = "m - n" in exI, simp) +done +lemma [intro]: "tinres [Bk] (Bk\<^bsup>k\<^esup>) " +apply(auto simp: tinres_def exponent_def) +apply(case_tac k, auto) +apply(rule_tac x = "Suc 0" in exI, simp) +done + +lemma abacus_turing_eq_halt_pre: + "\ly = layout_of aprog; + tprog = tm_of aprog; + n < length am; + abc_steps_l (0, lm) aprog stp = (as, am); + mop_ss = start_of ly (length aprog); + as \ length aprog\ + \ \ stp m l. steps (Suc 0, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) + (tprog @ (tMp n (mop_ss - 1))) stp + = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)" +using abacus_turing_eq_halt' +apply(simp) +done + + +text {* + Main theorem for the case when the original Abacus program does halt. +*} +lemma abacus_turing_eq_halt: + assumes layout: + "ly = layout_of aprog" + -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} + and compiled: "tprog = tm_of aprog" + -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} + and halt_state: + -- {* @{text "as"} is a program label outside the range of @{text "aprog"}. So + if Abacus is in such a state, it is in halt state: *} + "as \ length aprog" + and abc_exec: + -- {* Supposing after @{text "stp"} step of execution, Abacus program @{text "aprog"} + reaches such a halt state: *} + "abc_steps_l (0, lm) aprog stp = (as, am)" + and rs_locate: + -- {* @{text "n"} is a memory address in the range of Abacus memory @{text "am"}: *} + "n < length am" + and mopup_start: + -- {* The startling label for mopup mahines, according to the layout and Abacus program + should be @{text "mop_ss"}: *} + "mop_ss = start_of ly (length aprog)" + shows + -- {* + After @{text "stp"} steps of execution of the TM composed of @{text "tprog"} and the mopup + TM @{text "(tMp n (mop_ss - 1))"} will halt and gives rise to a configuration which + only hold the content of memory cell @{text "n"}: + *} + "\ stp m l. steps (Suc 0, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) (tprog @ (tMp n (mop_ss - 1))) stp + = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)" + using layout compiled halt_state abc_exec rs_locate mopup_start + by(rule_tac abacus_turing_eq_halt_pre, auto) + +lemma abacus_turing_eq_uhalt': + "\ly = layout_of aprog; + tprog = tm_of aprog; + \ stp. ((\ (as, am). as < length aprog) + (abc_steps_l (0, lm) aprog stp)); + mop_ss = start_of ly (length aprog)\ + \ (\ (\ stp. isS0 (steps (Suc 0, [Bk, Bk], ) + (tprog @ (tMp n (mop_ss - 1))) stp)))" +apply(drule_tac tc = "(Suc 0, [Bk, Bk], )" and n = n and ires = "[]" in + abacus_turing_eq_unhalt_case, auto intro: crsp_l_start) +apply(simp add: crsp_l.simps start_of.simps) +apply(erule_tac x = stp in allE, erule_tac x = stp in allE) +apply(subgoal_tac + "length (tm_of aprog @ tMp n + (start_of ly (length aprog) - Suc 0)) mod 2 = 0") +apply(simp add: steps_eq, auto simp: isS0_def) +done +(* +lemma abacus_turing_eq_uhalt_pre: + "\ly = layout_of aprog; + tprog = tm_of aprog; + \ stp. ((\ (as, am). as < length aprog) + (abc_steps_l (0, lm) aprog stp)); + mop_ss = start_of ly (length aprog)\ + \ (\ (\ stp. isS0 (steps (Suc 0, [Bk, Bk], ) + (tprog @ (tMp n (mop_ss - 1))) stp)))" +apply(drule_tac k = 0 and n = n in abacus_turing_eq_uhalt', auto) +apply(erule_tac x = stp in allE, erule_tac x = stp in allE) +apply(subgoal_tac "tinres ([Bk]) (Bk\<^bsup>k\<^esup>)") +apply(case_tac "steps (Suc 0, Bk\<^bsup>k\<^esup>, ) + (tm_of aprog @ tMp n (start_of ly (length aprog) - Suc 0)) stp") +apply(case_tac + "steps (Suc 0, [Bk], ) + (tm_of aprog @ tMp n (start_of ly (length aprog) - Suc 0)) stp") +apply(drule_tac tinres_steps, auto simp: isS0_def) +done +*) +text {* + Main theorem for the case when the original Abacus program does not halt. + *} +lemma abacus_turing_eq_uhalt: + assumes layout: + -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} + "ly = layout_of aprog" + and compiled: + -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} + "tprog = tm_of aprog" + and abc_unhalt: + -- {* + If, no matter how many steps the Abacus program @{text "aprog"} executes, it + may never reach a halt state. + *} + "\ stp. ((\ (as, am). as < length aprog) + (abc_steps_l (0, lm) aprog stp))" + and mop_start: "mop_ss = start_of ly (length aprog)" + shows + -- {* + The the TM composed of TM @{text "tprog"} and the moupup TM may never reach a halt state as well. + *} + "\ (\ stp. isS0 (steps (Suc 0, [Bk, Bk], ) + (tprog @ (tMp n (mop_ss - 1))) stp))" + using abacus_turing_eq_uhalt' + layout compiled abc_unhalt mop_start + by(auto) + + +end + diff -r cbb4ac6c8081 -r 1ce04eb1c8ad utm/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/utm/document/root.tex Sat Sep 29 12:38:12 2012 +0000 @@ -0,0 +1,60 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +%\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage[greek,english]{babel} + %option greek for \ + %option english (default language) for \, \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + + +\begin{document} + +\title{utm} +\author{By xujian} +\maketitle + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +% optional bibliography +%\bibliographystyle{abbrv} +%\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r cbb4ac6c8081 -r 1ce04eb1c8ad utm/rec_def.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/utm/rec_def.thy Sat Sep 29 12:38:12 2012 +0000 @@ -0,0 +1,87 @@ +theory rec_def +imports Main +begin + +section {* + Recursive functions +*} + +text {* + Datatype of recursive operators. +*} + +datatype recf = + -- {* The zero function, which always resturns @{text "0"} as result. *} + z | + -- {* The successor function, which increments its arguments. *} + s | + -- {* + The projection function, where @{text "id i j"} returns the @{text "j"}-th + argment out of the @{text "i"} arguments. + *} + id nat nat | + -- {* + The compostion operator, where "@{text "Cn n f [g1; g2; \ ;gm]"} + computes @{text "f (g1(x1, x2, \, xn), g2(x1, x2, \, xn), \ , + gm(x1, x2, \ , xn))"} for input argments @{text "x1, \, xn"}. + *} + Cn nat recf "recf list" | +-- {* + The primitive resursive operator, where @{text "Pr n f g"} computes: + @{text "Pr n f g (x1, x2, \, xn-1, 0) = f(x1, \, xn-1)"} + and @{text "Pr n f g (x1, x2, \, xn-1, k') = g(x1, x2, \, xn-1, k, + Pr n f g (x1, \, xn-1, k))"}. + *} + Pr nat recf recf | +-- {* + The minimization operator, where @{text "Mn n f (x1, x2, \ , xn)"} + computes the first i such that @{text "f (x1, \, xn, i) = 0"} and for all + @{text "j"}, @{text "f (x1, x2, \, xn, j) > 0"}. + *} + Mn nat recf + +text {* + The semantis of recursive operators is given by an inductively defined + relation as follows, where + @{text "rec_calc_rel R [x1, x2, \, xn] r"} means the computation of + @{text "R"} over input arguments @{text "[x1, x2, \, xn"} terminates + and gives rise to a result @{text "r"} +*} + +inductive rec_calc_rel :: "recf \ nat list \ nat \ bool" +where + calc_z: "rec_calc_rel z [n] 0" | + calc_s: "rec_calc_rel s [n] (Suc n)" | + calc_id: "\length args = i; j < i; args!j = r\ \ rec_calc_rel (id i j) args r" | + calc_cn: "\length args = n; + \ k < length gs. rec_calc_rel (gs ! k) args (rs ! k); + length rs = length gs; + rec_calc_rel f rs r\ + \ rec_calc_rel (Cn n f gs) args r" | + calc_pr_zero: + "\length args = n; + rec_calc_rel f args r0 \ + \ rec_calc_rel (Pr n f g) (args @ [0]) r0" | + calc_pr_ind: " + \ length args = n; + rec_calc_rel (Pr n f g) (args @ [k]) rk; + rec_calc_rel g (args @ [k] @ [rk]) rk'\ + \ rec_calc_rel (Pr n f g) (args @ [Suc k]) rk'" | + calc_mn: "\length args = n; + rec_calc_rel f (args@[r]) 0; + \ i < r. (\ ri. rec_calc_rel f (args@[i]) ri \ ri \ 0)\ + \ rec_calc_rel (Mn n f) args r" + +inductive_cases calc_pr_reverse: + "rec_calc_rel (Pr n f g) (lm) rSucy" + +inductive_cases calc_z_reverse: "rec_calc_rel z lm x" + +inductive_cases calc_s_reverse: "rec_calc_rel s lm x" + +inductive_cases calc_id_reverse: "rec_calc_rel (id m n) lm x" + +inductive_cases calc_cn_reverse: "rec_calc_rel (Cn n f gs) lm x" + +inductive_cases calc_mn_reverse:"rec_calc_rel (Mn n f) lm x" +end \ No newline at end of file diff -r cbb4ac6c8081 -r 1ce04eb1c8ad utm/recursive.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/utm/recursive.thy Sat Sep 29 12:38:12 2012 +0000 @@ -0,0 +1,5024 @@ +theory recursive +imports Main rec_def abacus +begin + +section {* + Compiling from recursive functions to Abacus machines + *} + +text {* + Some auxilliary Abacus machines used to construct the result Abacus machines. +*} + +text {* + @{text "get_paras_num recf"} returns the arity of recursive function @{text "recf"}. +*} +fun get_paras_num :: "recf \ nat" + where + "get_paras_num z = 1" | + "get_paras_num s = 1" | + "get_paras_num (id m n) = m" | + "get_paras_num (Cn n f gs) = n" | + "get_paras_num (Pr n f g) = Suc n" | + "get_paras_num (Mn n f) = n" + +fun addition :: "nat \ nat \ nat \ abc_prog" + where + "addition m n p = [Dec m 4, Inc n, Inc p, Goto 0, Dec p 7, + Inc m, Goto 4]" + +fun empty :: "nat \ nat \ abc_prog" + where + "empty m n = [Dec m 3, Inc n, Goto 0]" + +fun abc_inst_shift :: "abc_inst \ nat \ abc_inst" + where + "abc_inst_shift (Inc m) n = Inc m" | + "abc_inst_shift (Dec m e) n = Dec m (e + n)" | + "abc_inst_shift (Goto m) n = Goto (m + n)" + +fun abc_shift :: "abc_inst list \ nat \ abc_inst list" + where + "abc_shift xs n = map (\ x. abc_inst_shift x n) xs" + +fun abc_append :: "abc_inst list \ abc_inst list \ + abc_inst list" (infixl "[+]" 60) + where + "abc_append al bl = (let al_len = length al in + al @ abc_shift bl al_len)" + +text {* + The compilation of @{text "z"}-operator. +*} +definition rec_ci_z :: "abc_inst list" + where + "rec_ci_z \ [Goto 1]" + +text {* + The compilation of @{text "s"}-operator. +*} +definition rec_ci_s :: "abc_inst list" + where + "rec_ci_s \ (addition 0 1 2 [+] [Inc 1])" + + +text {* + The compilation of @{text "id i j"}-operator +*} + +fun rec_ci_id :: "nat \ nat \ abc_inst list" + where + "rec_ci_id i j = addition j i (i + 1)" + + +fun mv_boxes :: "nat \ nat \ nat \ abc_inst list" + where + "mv_boxes ab bb 0 = []" | + "mv_boxes ab bb (Suc n) = mv_boxes ab bb n [+] empty (ab + n) + (bb + n)" + +fun empty_boxes :: "nat \ abc_inst list" + where + "empty_boxes 0 = []" | + "empty_boxes (Suc n) = empty_boxes n [+] [Dec n 2, Goto 0]" + +fun cn_merge_gs :: + "(abc_inst list \ nat \ nat) list \ nat \ abc_inst list" + where + "cn_merge_gs [] p = []" | + "cn_merge_gs (g # gs) p = + (let (gprog, gpara, gn) = g in + gprog [+] empty gpara p [+] cn_merge_gs gs (Suc p))" + + +text {* + The compiler of recursive functions, where @{text "rec_ci recf"} return + @{text "(ap, arity, fp)"}, where @{text "ap"} is the Abacus program, @{text "arity"} is the + arity of the recursive function @{text "recf"}, +@{text "fp"} is the amount of memory which is going to be + used by @{text "ap"} for its execution. +*} + +function rec_ci :: "recf \ abc_inst list \ nat \ nat" + where + "rec_ci z = (rec_ci_z, 1, 2)" | + "rec_ci s = (rec_ci_s, 1, 3)" | + "rec_ci (id m n) = (rec_ci_id m n, m, m + 2)" | + "rec_ci (Cn n f gs) = + (let cied_gs = map (\ g. rec_ci g) (f # gs) in + let (fprog, fpara, fn) = hd cied_gs in + let pstr = + Max (set (Suc n # fn # (map (\ (aprog, p, n). n) cied_gs))) in + let qstr = pstr + Suc (length gs) in + (cn_merge_gs (tl cied_gs) pstr [+] mv_boxes 0 qstr n [+] + mv_boxes pstr 0 (length gs) [+] fprog [+] + empty fpara pstr [+] empty_boxes (length gs) [+] + empty pstr n [+] mv_boxes qstr 0 n, n, qstr + n))" | + "rec_ci (Pr n f g) = + (let (fprog, fpara, fn) = rec_ci f in + let (gprog, gpara, gn) = rec_ci g in + let p = Max (set ([n + 3, fn, gn])) in + let e = length gprog + 7 in + (empty n p [+] fprog [+] empty n (Suc n) [+] + (([Dec p e] [+] gprog [+] + [Inc n, Dec (Suc n) 3, Goto 1]) @ + [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gprog + 4)]), + Suc n, p + 1))" | + "rec_ci (Mn n f) = + (let (fprog, fpara, fn) = rec_ci f in + let len = length (fprog) in + (fprog @ [Dec (Suc n) (len + 5), Dec (Suc n) (len + 3), + Goto (len + 1), Inc n, Goto 0], n, max (Suc n) fn) )" + by pat_completeness auto +termination +proof +term size + show "wf (measure size)" by auto +next + fix n f gs x + assume "(x::recf) \ set (f # gs)" + thus "(x, Cn n f gs) \ measure size" + by(induct gs, auto) +next + fix n f g + show "(f, Pr n f g) \ measure size" by auto +next + fix n f g x xa y xb ya + show "(g, Pr n f g) \ measure size" by auto +next + fix n f + show "(f, Mn n f) \ measure size" by auto +qed + +declare rec_ci.simps [simp del] rec_ci_s_def[simp del] + rec_ci_z_def[simp del] rec_ci_id.simps[simp del] + mv_boxes.simps[simp del] abc_append.simps[simp del] + empty.simps[simp del] addition.simps[simp del] + +thm rec_calc_rel.induct + +declare abc_steps_l.simps[simp del] abc_fetch.simps[simp del] + abc_step_l.simps[simp del] + +lemma abc_steps_add: + "abc_steps_l (as, lm) ap (m + n) = + abc_steps_l (abc_steps_l (as, lm) ap m) ap n" +apply(induct m arbitrary: n as lm, simp add: abc_steps_l.simps) +proof - + fix m n as lm + assume ind: + "\n as lm. abc_steps_l (as, lm) ap (m + n) = + abc_steps_l (abc_steps_l (as, lm) ap m) ap n" + show "abc_steps_l (as, lm) ap (Suc m + n) = + abc_steps_l (abc_steps_l (as, lm) ap (Suc m)) ap n" + apply(insert ind[of as lm "Suc n"], simp) + apply(insert ind[of as lm "Suc 0"], simp add: abc_steps_l.simps) + apply(case_tac "(abc_steps_l (as, lm) ap m)", simp) + apply(simp add: abc_steps_l.simps) + apply(case_tac "abc_step_l (a, b) (abc_fetch a ap)", + simp add: abc_steps_l.simps) + done +qed + +(*lemmas: rec_ci and rec_calc_rel*) + +lemma rec_calc_inj_case_z: + "\rec_calc_rel z l x; rec_calc_rel z l y\ \ x = y" +apply(auto elim: calc_z_reverse) +done + +lemma rec_calc_inj_case_s: + "\rec_calc_rel s l x; rec_calc_rel s l y\ \ x = y" +apply(auto elim: calc_s_reverse) +done + +lemma rec_calc_inj_case_id: + "\rec_calc_rel (recf.id nat1 nat2) l x; + rec_calc_rel (recf.id nat1 nat2) l y\ \ x = y" +apply(auto elim: calc_id_reverse) +done + +lemma rec_calc_inj_case_mn: + assumes ind: "\ l x y. \rec_calc_rel f l x; rec_calc_rel f l y\ + \ x = y" + and h: "rec_calc_rel (Mn n f) l x" "rec_calc_rel (Mn n f) l y" + shows "x = y" + apply(insert h) + apply(elim calc_mn_reverse) + apply(case_tac "x > y", simp) + apply(erule_tac x = "y" in allE, auto) +proof - + fix v va + assume "rec_calc_rel f (l @ [y]) 0" + "rec_calc_rel f (l @ [y]) v" + "0 < v" + thus "False" + apply(insert ind[of "l @ [y]" 0 v], simp) + done +next + fix v va + assume + "rec_calc_rel f (l @ [x]) 0" + "\xv. rec_calc_rel f (l @ [x]) v \ 0 < v" "\ y < x" + thus "x = y" + apply(erule_tac x = "x" in allE) + apply(case_tac "x = y", auto) + apply(drule_tac y = v in ind, simp, simp) + done +qed + +lemma rec_calc_inj_case_pr: + assumes f_ind: + "\l x y. \rec_calc_rel f l x; rec_calc_rel f l y\ \ x = y" + and g_ind: + "\x xa y xb ya l xc yb. + \x = rec_ci f; (xa, y) = x; (xb, ya) = y; + rec_calc_rel g l xc; rec_calc_rel g l yb\ \ xc = yb" + and h: "rec_calc_rel (Pr n f g) l x" "rec_calc_rel (Pr n f g) l y" + shows "x = y" + apply(case_tac "rec_ci f") +proof - + fix a b c + assume "rec_ci f = (a, b, c)" + hence ng_ind: + "\ l xc yb. \rec_calc_rel g l xc; rec_calc_rel g l yb\ + \ xc = yb" + apply(insert g_ind[of "(a, b, c)" "a" "(b, c)" b c], simp) + done + from h show "x = y" + apply(erule_tac calc_pr_reverse, erule_tac calc_pr_reverse) + apply(erule f_ind, simp, simp) + apply(erule_tac calc_pr_reverse, simp, simp) + proof - + fix la ya ry laa yaa rya + assume k1: "rec_calc_rel g (la @ [ya, ry]) x" + "rec_calc_rel g (la @ [ya, rya]) y" + and k2: "rec_calc_rel (Pr (length la) f g) (la @ [ya]) ry" + "rec_calc_rel (Pr (length la) f g) (la @ [ya]) rya" + from k2 have "ry = rya" + apply(induct ya arbitrary: ry rya) + apply(erule_tac calc_pr_reverse, + erule_tac calc_pr_reverse, simp) + apply(erule f_ind, simp, simp, simp) + apply(erule_tac calc_pr_reverse, simp) + apply(erule_tac rSucy = rya in calc_pr_reverse, simp, simp) + proof - + fix ya ry rya l y ryb laa yb ryc + assume ind: + "\ry rya. \rec_calc_rel (Pr (length l) f g) (l @ [y]) ry; + rec_calc_rel (Pr (length l) f g) (l @ [y]) rya\ \ ry = rya" + and j: "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryb" + "rec_calc_rel g (l @ [y, ryb]) ry" + "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryc" + "rec_calc_rel g (l @ [y, ryc]) rya" + from j show "ry = rya" + apply(insert ind[of ryb ryc], simp) + apply(insert ng_ind[of "l @ [y, ryc]" ry rya], simp) + done + qed + from k1 and this show "x = y" + apply(simp) + apply(insert ng_ind[of "la @ [ya, rya]" x y], simp) + done + qed +qed + +lemma Suc_nth_part_eq: + "\k \k<(length list). (xs) ! k = (list) ! k" +apply(rule allI, rule impI) +apply(erule_tac x = "Suc k" in allE, simp) +done + + +lemma list_eq_intro: + "\length xs = length ys; \ k < length xs. xs ! k = ys ! k\ + \ xs = ys" +apply(induct xs arbitrary: ys, simp) +apply(case_tac ys, simp, simp) +proof - + fix a xs ys aa list + assume ind: + "\ys. \length list = length ys; \k + \ xs = ys" + and h: "length xs = length list" + "\k xs = list" + apply(insert ind[of list], simp) + apply(frule Suc_nth_part_eq, simp) + apply(erule_tac x = "0" in allE, simp) + done +qed + +lemma rec_calc_inj_case_cn: + assumes ind: + "\x l xa y. + \x = f \ x \ set gs; rec_calc_rel x l xa; rec_calc_rel x l y\ + \ xa = y" + and h: "rec_calc_rel (Cn n f gs) l x" + "rec_calc_rel (Cn n f gs) l y" + shows "x = y" + apply(insert h, elim calc_cn_reverse) + apply(subgoal_tac "rs = rsa") + apply(rule_tac x = f and l = rsa and xa = x and y = y in ind, + simp, simp, simp) + apply(intro list_eq_intro, simp, rule allI, rule impI) + apply(erule_tac x = k in allE, rule_tac x = k in allE, simp, simp) + apply(rule_tac x = "gs ! k" in ind, simp, simp, simp) + done + +lemma rec_calc_inj: + "\rec_calc_rel f l x; + rec_calc_rel f l y\ \ x = y" +apply(induct f arbitrary: l x y rule: rec_ci.induct) +apply(simp add: rec_calc_inj_case_z) +apply(simp add: rec_calc_inj_case_s) +apply(simp add: rec_calc_inj_case_id, simp) +apply(erule rec_calc_inj_case_cn,simp, simp) +apply(erule rec_calc_inj_case_pr, auto) +apply(erule rec_calc_inj_case_mn, auto) +done + + +lemma calc_rel_reverse_ind_step_ex: + "\rec_calc_rel (Pr n f g) (lm @ [Suc x]) rs\ + \ \ rs. rec_calc_rel (Pr n f g) (lm @ [x]) rs" +apply(erule calc_pr_reverse, simp, simp) +apply(rule_tac x = rk in exI, simp) +done + +lemma [simp]: "Suc x \ y \ Suc (y - Suc x) = y - x" +by arith + +lemma calc_pr_para_not_null: + "rec_calc_rel (Pr n f g) lm rs \ lm \ []" +apply(erule calc_pr_reverse, simp, simp) +done + +lemma calc_pr_less_ex: + "\rec_calc_rel (Pr n f g) lm rs; x \ last lm\ \ + \rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rs" +apply(subgoal_tac "lm \ []") +apply(induct x, rule_tac x = rs in exI, simp, simp, erule exE) +apply(rule_tac rs = xa in calc_rel_reverse_ind_step_ex, simp) +apply(simp add: calc_pr_para_not_null) +done + +lemma calc_pr_zero_ex: + "rec_calc_rel (Pr n f g) lm rs \ + \rs. rec_calc_rel f (butlast lm) rs" +apply(drule_tac x = "last lm" in calc_pr_less_ex, simp, + erule_tac exE, simp) +apply(erule_tac calc_pr_reverse, simp) +apply(rule_tac x = rs in exI, simp, simp) +done + + +lemma abc_steps_ind: + "abc_steps_l (as, am) ap (Suc stp) = + abc_steps_l (abc_steps_l (as, am) ap stp) ap (Suc 0)" +apply(insert abc_steps_add[of as am ap stp "Suc 0"], simp) +done + +lemma abc_steps_zero: "abc_steps_l asm ap 0 = asm" +apply(case_tac asm, simp add: abc_steps_l.simps) +done + +lemma abc_append_nth: + "n < length ap + length bp \ + (ap [+] bp) ! n = + (if n < length ap then ap ! n + else abc_inst_shift (bp ! (n - length ap)) (length ap))" +apply(simp add: abc_append.simps nth_append map_nth split: if_splits) +done + +lemma abc_state_keep: + "as \ length bp \ abc_steps_l (as, lm) bp stp = (as, lm)" +apply(induct stp, simp add: abc_steps_zero) +apply(simp add: abc_steps_ind) +apply(simp add: abc_steps_zero) +apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps) +done + +lemma abc_halt_equal: + "\abc_steps_l (0, lm) bp stpa = (length bp, lm1); + abc_steps_l (0, lm) bp stpb = (length bp, lm2)\ \ lm1 = lm2" +apply(case_tac "stpa - stpb > 0") +apply(insert abc_steps_add[of 0 lm bp stpb "stpa - stpb"], simp) +apply(insert abc_state_keep[of bp "length bp" lm2 "stpa - stpb"], + simp, simp add: abc_steps_zero) +apply(insert abc_steps_add[of 0 lm bp stpa "stpb - stpa"], simp) +apply(insert abc_state_keep[of bp "length bp" lm1 "stpb - stpa"], + simp) +done + +lemma abc_halt_point_ex: + "\\stp. abc_steps_l (0, lm) bp stp = (bs, lm'); + bs = length bp; bp \ []\ + \ \ stp. (\ (s, l). s < bs \ + (abc_steps_l (s, l) bp (Suc 0)) = (bs, lm')) + (abc_steps_l (0, lm) bp stp) " +apply(erule_tac exE) +proof - + fix stp + assume "bs = length bp" + "abc_steps_l (0, lm) bp stp = (bs, lm')" + "bp \ []" + thus + "\stp. (\(s, l). s < bs \ + abc_steps_l (s, l) bp (Suc 0) = (bs, lm')) + (abc_steps_l (0, lm) bp stp)" + apply(induct stp, simp add: abc_steps_zero, simp) + proof - + fix stpa + assume ind: + "abc_steps_l (0, lm) bp stpa = (length bp, lm') + \ \stp. (\(s, l). s < length bp \ abc_steps_l (s, l) bp + (Suc 0) = (length bp, lm')) (abc_steps_l (0, lm) bp stp)" + and h: "abc_steps_l (0, lm) bp (Suc stpa) = (length bp, lm')" + "abc_steps_l (0, lm) bp stp = (length bp, lm')" + "bp \ []" + from h show + "\stp. (\(s, l). s < length bp \ abc_steps_l (s, l) bp (Suc 0) + = (length bp, lm')) (abc_steps_l (0, lm) bp stp)" + apply(case_tac "abc_steps_l (0, lm) bp stpa", + case_tac "a = length bp") + apply(insert ind, simp) + apply(subgoal_tac "b = lm'", simp) + apply(rule_tac abc_halt_equal, simp, simp) + apply(rule_tac x = stpa in exI, simp add: abc_steps_ind) + apply(simp add: abc_steps_zero) + apply(rule classical, simp add: abc_steps_l.simps + abc_fetch.simps abc_step_l.simps) + done + qed +qed + + +lemma abc_append_empty_r[simp]: "[] [+] ab = ab" +apply(simp add: abc_append.simps abc_inst_shift.simps) +apply(induct ab, simp, simp) +apply(case_tac a, simp_all add: abc_inst_shift.simps) +done + +lemma abc_append_empty_l[simp]: "ab [+] [] = ab" +apply(simp add: abc_append.simps abc_inst_shift.simps) +done + + +lemma abc_append_length[simp]: + "length (ap [+] bp) = length ap + length bp" +apply(simp add: abc_append.simps) +done + +lemma abc_append_commute: "as [+] bs [+] cs = as [+] (bs [+] cs)" +apply(simp add: abc_append.simps abc_shift.simps abc_inst_shift.simps) +apply(induct cs, simp, simp) +apply(case_tac a, auto simp: abc_inst_shift.simps) +done + +lemma abc_halt_point_step[simp]: + "\a < length bp; abc_steps_l (a, b) bp (Suc 0) = (length bp, lm')\ + \ abc_steps_l (length ap + a, b) (ap [+] bp [+] cp) (Suc 0) = + (length ap + length bp, lm')" +apply(simp add: abc_steps_l.simps abc_fetch.simps abc_append_nth) +apply(case_tac "bp ! a", + auto simp: abc_steps_l.simps abc_step_l.simps) +done + +lemma abc_step_state_in: + "\bs < length bp; abc_steps_l (a, b) bp (Suc 0) = (bs, l)\ + \ a < length bp" +apply(simp add: abc_steps_l.simps abc_fetch.simps) +apply(rule_tac classical, + simp add: abc_step_l.simps abc_steps_l.simps) +done + + +lemma abc_append_state_in_exc: + "\bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\ + \ abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = + (length ap + bs, l)" +apply(induct stpa arbitrary: bs l, simp add: abc_steps_zero) +proof - + fix stpa bs l + assume ind: + "\bs l. \bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\ + \ abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = + (length ap + bs, l)" + and h: "bs < length bp" + "abc_steps_l (0, lm) bp (Suc stpa) = (bs, l)" + from h show + "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) (Suc stpa) = + (length ap + bs, l)" + apply(simp add: abc_steps_ind) + apply(case_tac "(abc_steps_l (0, lm) bp stpa)", simp) + proof - + fix a b + assume g: "abc_steps_l (0, lm) bp stpa = (a, b)" + "abc_steps_l (a, b) bp (Suc 0) = (bs, l)" + from h and g have k1: "a < length bp" + apply(simp add: abc_step_state_in) + done + from h and g and k1 show + "abc_steps_l (abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa) + (ap [+] bp [+] cp) (Suc 0) = (length ap + bs, l)" + apply(insert ind[of a b], simp) + apply(simp add: abc_steps_l.simps abc_fetch.simps + abc_append_nth) + apply(case_tac "bp ! a", auto simp: + abc_steps_l.simps abc_step_l.simps) + done + qed +qed + +lemma [simp]: "abc_steps_l (0, am) [] stp = (0, am)" +apply(induct stp, simp add: abc_steps_zero) +apply(simp add: abc_steps_ind) +apply(simp add: abc_steps_zero abc_steps_l.simps + abc_fetch.simps abc_step_l.simps) +done + +lemma abc_append_exc1: + "\\ stp. abc_steps_l (0, lm) bp stp = (bs, lm'); + bs = length bp; + as = length ap\ + \ \ stp. abc_steps_l (as, lm) (ap [+] bp [+] cp) stp + = (as + bs, lm')" +apply(case_tac "bp = []", erule_tac exE, simp, + rule_tac x = 0 in exI, simp add: abc_steps_zero) +apply(frule_tac abc_halt_point_ex, simp, simp, + erule_tac exE, erule_tac exE) +apply(rule_tac x = "stpa + Suc 0" in exI) +apply(case_tac "(abc_steps_l (0, lm) bp stpa)", + simp add: abc_steps_ind) +apply(subgoal_tac + "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa + = (length ap + a, b)", simp) +apply(simp add: abc_steps_zero) +apply(rule_tac abc_append_state_in_exc, simp, simp) +done + +lemma abc_append_exc3: + "\\ stp. abc_steps_l (0, am) bp stp = (bs, bm); ss = length ap\ + \ \ stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)" +apply(erule_tac exE) +proof - + fix stp + assume h: "abc_steps_l (0, am) bp stp = (bs, bm)" "ss = length ap" + thus " \stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)" + proof(induct stp arbitrary: bs bm) + fix bs bm + assume "abc_steps_l (0, am) bp 0 = (bs, bm)" + thus "\stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)" + apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps) + done + next + fix stp bs bm + assume ind: + "\bs bm. \abc_steps_l (0, am) bp stp = (bs, bm); + ss = length ap\ \ + \stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)" + and g: "abc_steps_l (0, am) bp (Suc stp) = (bs, bm)" + from g show + "\stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)" + apply(insert abc_steps_add[of 0 am bp stp "Suc 0"], simp) + apply(case_tac "(abc_steps_l (0, am) bp stp)", simp) + proof - + fix a b + assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)" + "abc_steps_l (0, am) bp (Suc stp) = + abc_steps_l (a, b) bp (Suc 0)" + "abc_steps_l (0, am) bp stp = (a, b)" + thus "?thesis" + apply(insert ind[of a b], simp add: h, erule_tac exE) + apply(rule_tac x = "Suc stp" in exI) + apply(simp only: abc_steps_ind, simp add: abc_steps_zero) + proof - + fix stp + assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)" + thus "abc_steps_l (a + length ap, b) (ap [+] bp) (Suc 0) + = (bs + length ap, bm)" + apply(simp add: abc_steps_l.simps abc_steps_zero + abc_fetch.simps split: if_splits) + apply(case_tac "bp ! a", + simp_all add: abc_inst_shift.simps abc_append_nth + abc_steps_l.simps abc_steps_zero abc_step_l.simps) + apply(auto) + done + qed + qed + qed +qed + +lemma abc_add_equal: + "\ap \ []; + abc_steps_l (0, am) ap astp = (a, b); + a < length ap\ + \ (abc_steps_l (0, am) (ap @ bp) astp) = (a, b)" +apply(induct astp arbitrary: a b, simp add: abc_steps_l.simps, simp) +apply(simp add: abc_steps_ind) +apply(case_tac "(abc_steps_l (0, am) ap astp)") +proof - + fix astp a b aa ba + assume ind: + "\a b. \abc_steps_l (0, am) ap astp = (a, b); + a < length ap\ \ + abc_steps_l (0, am) (ap @ bp) astp = (a, b)" + and h: "abc_steps_l (abc_steps_l (0, am) ap astp) ap (Suc 0) + = (a, b)" + "a < length ap" + "abc_steps_l (0, am) ap astp = (aa, ba)" + from h show "abc_steps_l (abc_steps_l (0, am) (ap @ bp) astp) + (ap @ bp) (Suc 0) = (a, b)" + apply(insert ind[of aa ba], simp) + apply(subgoal_tac "aa < length ap", simp) + apply(simp add: abc_steps_l.simps abc_fetch.simps + nth_append abc_steps_zero) + apply(rule abc_step_state_in, auto) + done +qed + + +lemma abc_add_exc1: + "\\ astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap\ + \ \ stp. abc_steps_l (0, am) (ap @ bp) stp = (as, bm)" +apply(case_tac "ap = []", simp, + rule_tac x = 0 in exI, simp add: abc_steps_zero) +apply(drule_tac abc_halt_point_ex, simp, simp) +apply(erule_tac exE, case_tac "(abc_steps_l (0, am) ap astp)", simp) +apply(rule_tac x = "Suc astp" in exI, simp add: abc_steps_ind, auto) +apply(frule_tac bp = bp in abc_add_equal, simp, simp, simp) +apply(simp add: abc_steps_l.simps abc_steps_zero + abc_fetch.simps nth_append) +done + +declare abc_shift.simps[simp del] + +lemma abc_append_exc2: + "\\ astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap; + \ bstp. abc_steps_l (0, bm) bp bstp = (bs, bm'); bs = length bp; + cs = as + bs; bp \ []\ + \ \ stp. abc_steps_l (0, am) (ap [+] bp) stp = (cs, bm')" +apply(insert abc_append_exc1[of bm bp bs bm' as ap "[]"], simp) +apply(drule_tac bp = "abc_shift bp (length ap)" in abc_add_exc1, simp) +apply(subgoal_tac "ap @ abc_shift bp (length ap) = ap [+] bp", + simp, auto) +apply(rule_tac x = "stpa + stp" in exI, simp add: abc_steps_add) +apply(simp add: abc_append.simps) +done +lemma exp_length[simp]: "length (a\<^bsup>b\<^esup>) = b" +by(simp add: exponent_def) +lemma exponent_add_iff: "a\<^bsup>b\<^esup> @ a\<^bsup>c \<^esup>@ xs = a\<^bsup>b + c \<^esup>@ xs" +apply(auto simp: exponent_def replicate_add) +done +lemma exponent_cons_iff: "a # a\<^bsup>c \<^esup>@ xs = a\<^bsup>Suc c \<^esup>@ xs" +apply(auto simp: exponent_def replicate_add) +done + + +lemma [simp]: "length lm = n \ + abc_steps_l (Suc 0, lm @ Suc x # 0 # suf_lm) + [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0)) + = (3, lm @ Suc x # 0 # suf_lm)" +apply(simp add: abc_steps_l.simps abc_fetch.simps + abc_step_l.simps abc_lm_v.simps abc_lm_s.simps + nth_append list_update_append) +done + +lemma [simp]: + "length lm = n \ + abc_steps_l (Suc 0, lm @ Suc x # Suc y # suf_lm) + [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0)) + = (Suc 0, lm @ Suc x # y # suf_lm)" +apply(simp add: abc_steps_l.simps abc_fetch.simps + abc_step_l.simps abc_lm_v.simps abc_lm_s.simps + nth_append list_update_append) +done + +lemma pr_cycle_part_middle_inv: + "\length lm = n\ \ + \ stp. abc_steps_l (0, lm @ x # y # suf_lm) + [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp + = (3, lm @ Suc x # 0 # suf_lm)" +proof - + assume h: "length lm = n" + hence k1: "\ stp. abc_steps_l (0, lm @ x # y # suf_lm) + [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp + = (Suc 0, lm @ Suc x # y # suf_lm)" + apply(rule_tac x = "Suc 0" in exI) + apply(simp add: abc_steps_l.simps abc_step_l.simps + abc_lm_v.simps abc_lm_s.simps nth_append + list_update_append abc_fetch.simps) + done + from h have k2: + "\ stp. abc_steps_l (Suc 0, lm @ Suc x # y # suf_lm) + [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp + = (3, lm @ Suc x # 0 # suf_lm)" + apply(induct y) + apply(rule_tac x = "Suc (Suc 0)" in exI, simp, simp, + erule_tac exE) + apply(rule_tac x = "Suc (Suc 0) + stp" in exI, + simp only: abc_steps_add, simp) + done + from k1 and k2 show + "\ stp. abc_steps_l (0, lm @ x # y # suf_lm) + [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp + = (3, lm @ Suc x # 0 # suf_lm)" + apply(erule_tac exE, erule_tac exE) + apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add) + done +qed + +lemma [simp]: + "length lm = Suc n \ + (abc_steps_l (length ap, lm @ x # Suc y # suf_lm) + (ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length ap)]) + (Suc (Suc (Suc 0)))) + = (length ap, lm @ Suc x # y # suf_lm)" +apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps + abc_lm_v.simps list_update_append nth_append abc_lm_s.simps) +done + +lemma switch_para_inv: + assumes bp_def:"bp = ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto ss]" + and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" + "ss = length ap" + "length lm = Suc n" + shows " \stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp = + (0, lm @ (x + y) # 0 # suf_lm)" +apply(induct y arbitrary: x) +apply(rule_tac x = "Suc 0" in exI, + simp add: bp_def empty.simps abc_steps_l.simps + abc_fetch.simps h abc_step_l.simps + abc_lm_v.simps list_update_append nth_append + abc_lm_s.simps) +proof - + fix y x + assume ind: + "\x. \stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp = + (0, lm @ (x + y) # 0 # suf_lm)" + show "\stp. abc_steps_l (ss, lm @ x # Suc y # suf_lm) bp stp = + (0, lm @ (x + Suc y) # 0 # suf_lm)" + apply(insert ind[of "Suc x"], erule_tac exE) + apply(rule_tac x = "Suc (Suc (Suc 0)) + stp" in exI, + simp only: abc_steps_add bp_def h) + apply(simp add: h) + done +qed + +lemma [simp]: + "length lm = rs_pos \ Suc (Suc rs_pos) < a_md \ 0 < rs_pos \ + a_md - Suc 0 < Suc (Suc (Suc (a_md + length suf_lm - + Suc (Suc (Suc 0)))))" +apply(arith) +done + +lemma [simp]: + "Suc (Suc rs_pos) < a_md \ 0 < rs_pos \ + \ a_md - Suc 0 < rs_pos - Suc 0" +apply(arith) +done + +lemma [simp]: + "Suc (Suc rs_pos) < a_md \ 0 < rs_pos \ + \ a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))" +apply(arith) +done + +lemma butlast_append_last: "lm \ [] \ lm = butlast lm @ [last lm]" +apply(auto) +done + +lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md) + \ (Suc (Suc rs_pos)) < a_md" +apply(simp add: rec_ci.simps) +apply(case_tac "rec_ci f", simp) +apply(case_tac "rec_ci g", simp) +apply(arith) +done + +(* +lemma pr_para_ge_suc0: "rec_calc_rel (Pr n f g) lm xs \ 0 < n" +apply(erule calc_pr_reverse, simp, simp) +done +*) + +lemma ci_pr_para_eq: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md) + \ rs_pos = Suc n" +apply(simp add: rec_ci.simps) +apply(case_tac "rec_ci g", case_tac "rec_ci f", simp) +done + +lemma [intro]: + "\rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm xs\ + \ length lm = rs_pos" +apply(simp add: rec_ci.simps rec_ci_z_def) +apply(erule_tac calc_z_reverse, simp) +done + +lemma [intro]: + "\rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm xs\ + \ length lm = rs_pos" +apply(simp add: rec_ci.simps rec_ci_s_def) +apply(erule_tac calc_s_reverse, simp) +done + +lemma [intro]: + "\rec_ci (recf.id nat1 nat2) = (aprog, rs_pos, a_md); + rec_calc_rel (recf.id nat1 nat2) lm xs\ \ length lm = rs_pos" +apply(simp add: rec_ci.simps rec_ci_id.simps) +apply(erule_tac calc_id_reverse, simp) +done + +lemma [intro]: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_calc_rel (Cn n f gs) lm xs\ \ length lm = rs_pos" +apply(erule_tac calc_cn_reverse, simp) +apply(simp add: rec_ci.simps) +apply(case_tac "rec_ci f", simp) +done + +lemma [intro]: + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_calc_rel (Pr n f g) lm xs\ \ length lm = rs_pos" +apply(erule_tac calc_pr_reverse, simp) +apply(drule_tac ci_pr_para_eq, simp, simp) +apply(drule_tac ci_pr_para_eq, simp) +done + +lemma [intro]: + "\rec_ci (Mn n f) = (aprog, rs_pos, a_md); + rec_calc_rel (Mn n f) lm xs\ \ length lm = rs_pos" +apply(erule_tac calc_mn_reverse) +apply(simp add: rec_ci.simps) +apply(case_tac "rec_ci f", simp) +done + +lemma para_pattern: + "\rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm xs\ + \ length lm = rs_pos" +apply(case_tac f, auto) +done + +lemma ci_pr_g_paras: + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci g = (a, aa, ba); + rec_calc_rel (Pr n f g) (lm @ [x]) rs; x > 0\ \ + aa = Suc rs_pos " +apply(erule calc_pr_reverse, simp) +apply(subgoal_tac "length (args @ [k, rk]) = aa", simp) +apply(subgoal_tac "rs_pos = Suc n", simp) +apply(simp add: ci_pr_para_eq) +apply(erule para_pattern, simp) +done + +lemma ci_pr_g_md_less: + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci g = (a, aa, ba)\ \ ba < a_md" +apply(simp add: rec_ci.simps) +apply(case_tac "rec_ci f", auto) +done + +lemma [intro]: "rec_ci z = (ap, rp, ad) \ rp < ad" + by(simp add: rec_ci.simps) + +lemma [intro]: "rec_ci s = (ap, rp, ad) \ rp < ad" + by(simp add: rec_ci.simps) + +lemma [intro]: "rec_ci (recf.id nat1 nat2) = (ap, rp, ad) \ rp < ad" + by(simp add: rec_ci.simps) + +lemma [intro]: "rec_ci (Cn n f gs) = (ap, rp, ad) \ rp < ad" +apply(simp add: rec_ci.simps) +apply(case_tac "rec_ci f", simp) +done + +lemma [intro]: "rec_ci (Pr n f g) = (ap, rp, ad) \ rp < ad" +apply(simp add: rec_ci.simps) +by(case_tac "rec_ci f", case_tac "rec_ci g", auto) + +lemma [intro]: "rec_ci (Mn n f) = (ap, rp, ad) \ rp < ad" +apply(simp add: rec_ci.simps) +apply(case_tac "rec_ci f", simp) +apply(arith) +done + +lemma ci_ad_ge_paras: "rec_ci f = (ap, rp, ad) \ ad > rp" +apply(case_tac f, auto) +done + +lemma [elim]: "\a [+] b = []; a \ [] \ b \ []\ \ RR" +apply(auto simp: abc_append.simps abc_shift.simps) +done + +lemma [intro]: "rec_ci z = ([], aa, ba) \ False" +by(simp add: rec_ci.simps rec_ci_z_def) + +lemma [intro]: "rec_ci s = ([], aa, ba) \ False" +by(auto simp: rec_ci.simps rec_ci_s_def addition.simps) + +lemma [intro]: "rec_ci (id m n) = ([], aa, ba) \ False" +by(auto simp: rec_ci.simps rec_ci_id.simps addition.simps) + +lemma [intro]: "rec_ci (Cn n f gs) = ([], aa, ba) \ False" +apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_append.simps) +apply(simp add: abc_shift.simps empty.simps) +done + +lemma [intro]: "rec_ci (Pr n f g) = ([], aa, ba) \ False" +apply(simp add: rec_ci.simps) +apply(case_tac "rec_ci f", case_tac "rec_ci g") +by(auto) + +lemma [intro]: "rec_ci (Mn n f) = ([], aa, ba) \ False" +apply(case_tac "rec_ci f", auto simp: rec_ci.simps) +done + +lemma rec_ci_not_null: "rec_ci g = (a, aa, ba) \ a \ []" +by(case_tac g, auto) + +lemma calc_pr_g_def: + "\rec_calc_rel (Pr rs_pos f g) (lm @ [Suc x]) rsa; + rec_calc_rel (Pr rs_pos f g) (lm @ [x]) rsxa\ + \ rec_calc_rel g (lm @ [x, rsxa]) rsa" +apply(erule_tac calc_pr_reverse, simp, simp) +apply(subgoal_tac "rsxa = rk", simp) +apply(erule_tac rec_calc_inj, auto) +done + +lemma ci_pr_md_def: + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\ + \ a_md = Suc (max (n + 3) (max bc ba))" +by(simp add: rec_ci.simps) + +lemma ci_pr_f_paras: + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_calc_rel (Pr n f g) lm rs; + rec_ci f = (ab, ac, bc)\ \ ac = rs_pos - Suc 0" +apply(subgoal_tac "\rs. rec_calc_rel f (butlast lm) rs", + erule_tac exE) +apply(drule_tac f = f and lm = "butlast lm" in para_pattern, + simp, simp) +apply(drule_tac para_pattern, simp) +apply(subgoal_tac "lm \ []", simp) +apply(erule_tac calc_pr_reverse, simp, simp) +apply(erule calc_pr_zero_ex) +done + +lemma ci_pr_md_ge_f: "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci f = (ab, ac, bc)\ \ Suc bc \ a_md" +apply(case_tac "rec_ci g") +apply(simp add: rec_ci.simps, auto) +done + +lemma ci_pr_md_ge_g: "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci g = (ab, ac, bc)\ \ bc < a_md" +apply(case_tac "rec_ci f") +apply(simp add: rec_ci.simps, auto) +done + +lemma rec_calc_rel_def0: + "\rec_calc_rel (Pr n f g) lm rs; rec_calc_rel f (butlast lm) rsa\ + \ rec_calc_rel (Pr n f g) (butlast lm @ [0]) rsa" + apply(rule_tac calc_pr_zero, simp) +apply(erule_tac calc_pr_reverse, simp, simp, simp) +done + +lemma [simp]: "length (empty m n) = 3" +by (auto simp: empty.simps) +(* +lemma + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_calc_rel (Pr n f g) lm rs; + rec_ci g = (a, aa, ba); + rec_ci f = (ab, ac, bc)\ +\ \ap bp cp. aprog = ap [+] bp [+] cp \ length ap = 3 + length ab \ bp = recursive.empty (n - Suc 0) n 3" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "recursive.empty (n - Suc 0) (max (Suc (Suc n)) (max bc ba)) 3 [+] ab" in exI, simp) +apply(rule_tac x = "([Dec (max (Suc (Suc n)) (max bc ba)) (length a + 7)] [+] a [+] + [Inc (n - Suc 0), Dec n 3, Goto (Suc 0)]) @ [Dec (Suc n) 0, Inc n, Goto (length a + 4)]" in exI, simp) +apply(auto simp: abc_append_commute) +done + +lemma "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\ + \ \ap bp cp. aprog = ap [+] bp [+] cp \ length ap = 3 \ bp = ab" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "recursive.empty (n - Suc 0) (max (Suc (Suc n)) (max bc ba)) 3" in exI, simp) +apply(rule_tac x = "recursive.empty (n - Suc 0) n 3 [+] + ([Dec (max (Suc (Suc n)) (max bc ba)) (length a + 7)] [+] a + [+] [Inc (n - Suc 0), Dec n 3, Goto (Suc 0)]) @ [Dec (Suc n) 0, Inc n, Goto (length a + 4)]" in exI, auto) +apply(simp add: abc_append_commute) +done +*) + +lemma [simp]: "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\ + \ rs_pos = Suc n" +apply(simp add: ci_pr_para_eq) +done + + +lemma [simp]: "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\ + \ length lm = Suc n" +apply(subgoal_tac "rs_pos = Suc n", rule_tac para_pattern, simp, simp) +apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps) +done + +lemma [simp]: "rec_ci (Pr n f g) = (a, rs_pos, a_md) \ Suc (Suc n) < a_md" +apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps) +apply arith +done + +lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md) \ 0 < rs_pos" +apply(case_tac "rec_ci f", case_tac "rec_ci g") +apply(simp add: rec_ci.simps) +done + +lemma [simp]: "Suc (Suc rs_pos) < a_md \ + butlast lm @ (last lm - xa) # (rsa::nat) # 0 # 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm = + butlast lm @ (last lm - xa) # rsa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm" +apply(simp add: exp_ind_def[THEN sym]) +done + +lemma pr_cycle_part_ind: + assumes g_ind: + "\lm rs suf_lm. rec_calc_rel g lm rs \ + \stp. abc_steps_l (0, lm @ 0\<^bsup>ba - aa\<^esup> @ suf_lm) a stp = + (length a, lm @ rs # 0\<^bsup>ba - Suc aa\<^esup> @ suf_lm)" + and ap_def: + "ap = ([Dec (a_md - Suc 0) (length a + 7)] [+] + (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)])) @ + [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" + and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Pr n f g) + (butlast lm @ [last lm - Suc xa]) rsxa" + "Suc xa \ last lm" + "rec_ci g = (a, aa, ba)" + "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsa" + "lm \ []" + shows + "\stp. abc_steps_l + (0, butlast lm @ (last lm - Suc xa) # rsxa # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm) ap stp = + (0, butlast lm @ (last lm - xa) # rsa + # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm)" +proof - + have k1: "\stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # + rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm) ap stp = + (length a + 4, butlast lm @ (last lm - xa) # 0 # rsa # + 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm)" + apply(simp add: ap_def, rule_tac abc_add_exc1) + apply(rule_tac as = "Suc 0" and + bm = "butlast lm @ (last lm - Suc xa) # + rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm" in abc_append_exc2, + auto) + proof - + show + "\astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa + # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm) + [Dec (a_md - Suc 0)(length a + 7)] astp = + (Suc 0, butlast lm @ (last lm - Suc xa) # + rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm)" + apply(rule_tac x = "Suc 0" in exI, + simp add: abc_steps_l.simps abc_step_l.simps + abc_fetch.simps) + apply(subgoal_tac "length lm = Suc n \ rs_pos = Suc n \ + a_md > Suc (Suc rs_pos)") + apply(simp add: abc_lm_v.simps nth_append abc_lm_s.simps) + apply(insert nth_append[of + "(last lm - Suc xa) # rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup>" + "Suc xa # suf_lm" "(a_md - rs_pos)"], simp) + apply(simp add: list_update_append del: list_update.simps) + apply(insert list_update_append[of "(last lm - Suc xa) # rsxa # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup>" + "Suc xa # suf_lm" "a_md - rs_pos" "xa"], simp) + apply(case_tac a_md, simp, simp) + apply(insert h, simp) + apply(insert para_pattern[of "Pr n f g" aprog rs_pos a_md + "(butlast lm @ [last lm - Suc xa])" rsxa], simp) + done + next + show "\bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # + rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm) (a [+] + [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)]) bstp = + (3 + length a, butlast lm @ (last lm - xa) # 0 # rsa # + 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm)" + apply(rule_tac as = "length a" and + bm = "butlast lm @ (last lm - Suc xa) # rsxa # rsa # + 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm" + in abc_append_exc2, simp_all) + proof - + from h have j1: "aa = Suc rs_pos \ a_md > ba \ ba > Suc rs_pos" + apply(insert h) + apply(insert ci_pr_g_paras[of n f g aprog rs_pos + a_md a aa ba "butlast lm" "last lm - xa" rsa], simp) + apply(drule_tac ci_pr_md_ge_g, auto) + apply(erule_tac ci_ad_ge_paras) + done + from h have j2: "rec_calc_rel g (butlast lm @ + [last lm - Suc xa, rsxa]) rsa" + apply(rule_tac calc_pr_g_def, simp, simp) + done + from j1 and j2 show + "\astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # + rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm) a astp = + (length a, butlast lm @ (last lm - Suc xa) # rsxa # rsa + # 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm)" + apply(insert g_ind[of + "butlast lm @ (last lm - Suc xa) # [rsxa]" rsa + "0\<^bsup>a_md - ba - Suc 0 \<^esup> @ xa # suf_lm"], simp, auto) + apply(simp add: exponent_add_iff) + apply(rule_tac x = stp in exI, simp add: numeral_3_eq_3) + done + next + from h have j3: "length lm = rs_pos \ rs_pos > 0" + apply(rule_tac conjI) + apply(drule_tac lm = "(butlast lm @ [last lm - Suc xa])" + and xs = rsxa in para_pattern, simp, simp, simp) + done + from h have j4: "Suc (last lm - Suc xa) = last lm - xa" + apply(case_tac "last lm", simp, simp) + done + from j3 and j4 show + "\bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa # + rsa # 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm) + [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)] bstp = + (3, butlast lm @ (last lm - xa) # 0 # rsa # + 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm)" + apply(insert pr_cycle_part_middle_inv[of "butlast lm" + "rs_pos - Suc 0" "(last lm - Suc xa)" rsxa + "rsa # 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm"], simp) + done + qed + qed + from h have k2: + "\stp. abc_steps_l (length a + 4, butlast lm @ (last lm - xa) # 0 + # rsa # 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm) ap stp = + (0, butlast lm @ (last lm - xa) # rsa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm)" + apply(insert switch_para_inv[of ap + "([Dec (a_md - Suc 0) (length a + 7)] [+] + (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)]))" + n "length a + 4" f g aprog rs_pos a_md + "butlast lm @ [last lm - xa]" 0 rsa + "0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm"]) + apply(simp add: h ap_def) + apply(subgoal_tac "length lm = Suc n \ Suc (Suc rs_pos) < a_md", + simp) + apply(insert h, simp) + apply(frule_tac lm = "(butlast lm @ [last lm - Suc xa])" + and xs = rsxa in para_pattern, simp, simp) + done + from k1 and k2 show "?thesis" + apply(auto) + apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add) + done +qed + +lemma ci_pr_ex1: + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci g = (a, aa, ba); + rec_ci f = (ab, ac, bc)\ +\ \ap bp. length ap = 6 + length ab \ + aprog = ap [+] bp \ + bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] (a [+] + [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)])) @ + [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "recursive.empty n (max (Suc (Suc (Suc n))) + (max bc ba)) [+] ab [+] recursive.empty n (Suc n)" in exI, + simp) +apply(auto simp add: abc_append_commute add3_Suc) +done + +lemma pr_cycle_part: + "\\lm rs suf_lm. rec_calc_rel g lm rs \ + \stp. abc_steps_l (0, lm @ 0\<^bsup>ba - aa\<^esup> @ suf_lm) a stp = + (length a, lm @ rs # 0\<^bsup>ba - Suc aa\<^esup> @ suf_lm); + rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_calc_rel (Pr n f g) lm rs; + rec_ci g = (a, aa, ba); + rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rsx; + rec_ci f = (ab, ac, bc); + lm \ []; + x \ last lm\ \ + \stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) # + rsx # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ x # suf_lm) aprog stp = + (6 + length ab, butlast lm @ last lm # rs # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm)" +proof - + assume g_ind: + "\lm rs suf_lm. rec_calc_rel g lm rs \ + \stp. abc_steps_l (0, lm @ 0\<^bsup>ba - aa\<^esup> @ suf_lm) a stp = + (length a, lm @ rs # 0\<^bsup>ba - Suc aa\<^esup> @ suf_lm)" + and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Pr n f g) lm rs" + "rec_ci g = (a, aa, ba)" + "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rsx" + "lm \ []" + "x \ last lm" + "rec_ci f = (ab, ac, bc)" + from h show + "\stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) # + rsx # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ x # suf_lm) aprog stp = + (6 + length ab, butlast lm @ last lm # rs # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm)" + proof(induct x arbitrary: rsx, simp_all) + fix rsxa + assume "rec_calc_rel (Pr n f g) lm rsxa" + "rec_calc_rel (Pr n f g) lm rs" + from h and this have "rs = rsxa" + apply(subgoal_tac "lm \ [] \ rs_pos = Suc n", simp) + apply(rule_tac rec_calc_inj, simp, simp) + apply(simp) + done + thus "\stp. abc_steps_l (6 + length ab, butlast lm @ last lm # + rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm) aprog stp = + (6 + length ab, butlast lm @ last lm # rs # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm)" + by(rule_tac x = 0 in exI, simp add: abc_steps_l.simps) + next + fix xa rsxa + assume ind: + "\rsx. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsx + \ \stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - xa) # + rsx # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm) aprog stp = + (6 + length ab, butlast lm @ last lm # rs # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm)" + and g: "rec_calc_rel (Pr n f g) + (butlast lm @ [last lm - Suc xa]) rsxa" + "Suc xa \ last lm" + "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Pr n f g) lm rs" + "rec_ci g = (a, aa, ba)" + "rec_ci f = (ab, ac, bc)" "lm \ []" + from g have k1: + "\ rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rs" + apply(rule_tac rs = rs in calc_pr_less_ex, simp, simp) + done + from g and this show + "\stp. abc_steps_l (6 + length ab, + butlast lm @ (last lm - Suc xa) # rsxa # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm) aprog stp = + (6 + length ab, butlast lm @ last lm # rs # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm)" + proof(erule_tac exE) + fix rsa + assume k2: "rec_calc_rel (Pr n f g) + (butlast lm @ [last lm - xa]) rsa" + from g and k2 have + "\stp. abc_steps_l (6 + length ab, butlast lm @ + (last lm - Suc xa) # rsxa # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm) aprog stp + = (6 + length ab, butlast lm @ (last lm - xa) # rsa # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm)" + proof - + from g have k2_1: + "\ ap bp. length ap = 6 + length ab \ + aprog = ap [+] bp \ + bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] + (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, + Goto (Suc 0)])) @ + [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" + apply(rule_tac ci_pr_ex1, auto) + done + from k2_1 and k2 and g show "?thesis" + proof(erule_tac exE, erule_tac exE) + fix ap bp + assume + "length ap = 6 + length ab \ + aprog = ap [+] bp \ bp = + ([Dec (a_md - Suc 0) (length a + 7)] [+] + (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, + Goto (Suc 0)])) @ + [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" + from g and this and k2 and g_ind show "?thesis" + apply(insert abc_append_exc3[of + "butlast lm @ (last lm - Suc xa) # rsxa # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm" bp 0 + "butlast lm @ (last lm - xa) # rsa # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm" "length ap" ap], + simp) + apply(subgoal_tac + "\stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) + # rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # + suf_lm) bp stp = + (0, butlast lm @ (last lm - xa) # rsa # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm)", + simp, erule_tac conjE, erule conjE) + apply(erule pr_cycle_part_ind, auto) + done + qed + qed + from g and k2 and this show "?thesis" + apply(erule_tac exE) + apply(insert ind[of rsa], simp) + apply(erule_tac exE) + apply(rule_tac x = "stp + stpa" in exI, + simp add: abc_steps_add) + done + qed + qed +qed + +lemma ci_pr_length: + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci g = (a, aa, ba); + rec_ci f = (ab, ac, bc)\ + \ length aprog = 13 + length ab + length a" +apply(auto simp: rec_ci.simps) +done + +thm empty.simps +term max +fun empty_inv :: "nat \ nat list \ nat \ nat \ nat list \ bool" + where + "empty_inv (as, lm) m n initlm = + (let plus = initlm ! m + initlm ! n in + length initlm > max m n \ m \ n \ + (if as = 0 then \ k l. lm = initlm[m := k, n := l] \ + k + l = plus \ k \ initlm ! m + else if as = 1 then \ k l. lm = initlm[m := k, n := l] + \ k + l + 1 = plus \ k < initlm ! m + else if as = 2 then \ k l. lm = initlm[m := k, n := l] + \ k + l = plus \ k \ initlm ! m + else if as = 3 then lm = initlm[m := 0, n := plus] + else False))" + +fun empty_stage1 :: "nat \ nat list \ nat \ nat" + where + "empty_stage1 (as, lm) m = + (if as = 3 then 0 + else 1)" + +fun empty_stage2 :: "nat \ nat list \ nat \ nat" + where + "empty_stage2 (as, lm) m = (lm ! m)" + +fun empty_stage3 :: "nat \ nat list \ nat \ nat" + where + "empty_stage3 (as, lm) m = (if as = 1 then 3 + else if as = 2 then 2 + else if as = 0 then 1 + else 0)" + + + +fun empty_measure :: "((nat \ nat list) \ nat) \ (nat \ nat \ nat)" + where + "empty_measure ((as, lm), m) = + (empty_stage1 (as, lm) m, empty_stage2 (as, lm) m, + empty_stage3 (as, lm) m)" + +definition lex_pair :: "((nat \ nat) \ nat \ nat) set" + where + "lex_pair = less_than <*lex*> less_than" + +definition lex_triple :: + "((nat \ (nat \ nat)) \ (nat \ (nat \ nat))) set" + where + "lex_triple \ less_than <*lex*> lex_pair" + +definition empty_LE :: + "(((nat \ nat list) \ nat) \ ((nat \ nat list) \ nat)) set" + where + "empty_LE \ (inv_image lex_triple empty_measure)" + +lemma wf_lex_triple: "wf lex_triple" + by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def) + +lemma wf_empty_le[intro]: "wf empty_LE" +by(auto intro:wf_inv_image wf_lex_triple simp: empty_LE_def) + +declare empty_inv.simps[simp del] + +lemma empty_inv_init: +"\m < length initlm; n < length initlm; m \ n\ \ + empty_inv (0, initlm) m n initlm" +apply(simp add: abc_steps_l.simps empty_inv.simps) +apply(rule_tac x = "initlm ! m" in exI, + rule_tac x = "initlm ! n" in exI, simp) +done + +lemma [simp]: "abc_fetch 0 (recursive.empty m n) = Some (Dec m 3)" +apply(simp add: empty.simps abc_fetch.simps) +done + +lemma [simp]: "abc_fetch (Suc 0) (recursive.empty m n) = + Some (Inc n)" +apply(simp add: empty.simps abc_fetch.simps) +done + +lemma [simp]: "abc_fetch 2 (recursive.empty m n) = Some (Goto 0)" +apply(simp add: empty.simps abc_fetch.simps) +done + +lemma [simp]: "abc_fetch 3 (recursive.empty m n) = None" +apply(simp add: empty.simps abc_fetch.simps) +done + +lemma [simp]: + "\m \ n; m < length initlm; n < length initlm; + k + l = initlm ! m + initlm ! n; k \ initlm ! m; 0 < k\ + \ \ka la. initlm[m := k, n := l, m := k - Suc 0] = + initlm[m := ka, n := la] \ + Suc (ka + la) = initlm ! m + initlm ! n \ + ka < initlm ! m" +apply(rule_tac x = "k - Suc 0" in exI, rule_tac x = l in exI, + simp, auto) +apply(subgoal_tac + "initlm[m := k, n := l, m := k - Suc 0] = + initlm[n := l, m := k, m := k - Suc 0]") +apply(simp add: list_update_overwrite ) +apply(simp add: list_update_swap) +apply(simp add: list_update_swap) +done + +lemma [simp]: + "\m \ n; m < length initlm; n < length initlm; + Suc (k + l) = initlm ! m + initlm ! n; + k < initlm ! m\ + \ \ka la. initlm[m := k, n := l, n := Suc l] = + initlm[m := ka, n := la] \ + ka + la = initlm ! m + initlm ! n \ + ka \ initlm ! m" +apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto) +done + +lemma [simp]: + "\length initlm > max m n; m \ n\ \ + \na. \ (\(as, lm) m. as = 3) + (abc_steps_l (0, initlm) (recursive.empty m n) na) m \ + empty_inv (abc_steps_l (0, initlm) + (recursive.empty m n) na) m n initlm \ + empty_inv (abc_steps_l (0, initlm) + (recursive.empty m n) (Suc na)) m n initlm \ + ((abc_steps_l (0, initlm) (recursive.empty m n) (Suc na), m), + abc_steps_l (0, initlm) (recursive.empty m n) na, m) \ empty_LE" +apply(rule allI, rule impI, simp add: abc_steps_ind) +apply(case_tac "(abc_steps_l (0, initlm) (recursive.empty m n) na)", + simp) +apply(auto split:if_splits simp add:abc_steps_l.simps empty_inv.simps) +apply(auto simp add: empty_LE_def lex_triple_def lex_pair_def + abc_step_l.simps abc_steps_l.simps + empty_inv.simps abc_lm_v.simps abc_lm_s.simps + split: if_splits ) +apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, simp) +done + +lemma empty_inv_halt: + "\length initlm > max m n; m \ n\ \ + \ stp. (\ (as, lm). as = 3 \ + empty_inv (as, lm) m n initlm) + (abc_steps_l (0::nat, initlm) (empty m n) stp)" +apply(insert halt_lemma2[of empty_LE + "\ ((as, lm), m). as = (3::nat)" + "\ stp. (abc_steps_l (0, initlm) (recursive.empty m n) stp, m)" + "\ ((as, lm), m). empty_inv (as, lm) m n initlm"]) +apply(insert wf_empty_le, simp add: empty_inv_init abc_steps_zero) +apply(erule_tac exE) +apply(rule_tac x = na in exI) +apply(case_tac "(abc_steps_l (0, initlm) (recursive.empty m n) na)", + simp, auto) +done + +lemma empty_halt_cond: + "\m \ n; empty_inv (a, b) m n lm; a = 3\ \ + b = lm[n := lm ! m + lm ! n, m := 0]" +apply(simp add: empty_inv.simps, auto) +apply(simp add: list_update_swap) +done + +lemma empty_ex: + "\length lm > max m n; m \ n\ \ + \ stp. abc_steps_l (0::nat, lm) (empty m n) stp + = (3, (lm[n := (lm ! m + lm ! n)])[m := 0::nat])" +apply(drule empty_inv_halt, simp, erule_tac exE) +apply(rule_tac x = stp in exI) +apply(case_tac "abc_steps_l (0, lm) (recursive.empty m n) stp", + simp) +apply(erule_tac empty_halt_cond, auto) +done + +lemma [simp]: + "\a_md = Suc (max (Suc (Suc n)) (max bc ba)); + length lm = rs_pos \ rs_pos = n \ n > 0\ + \ n - Suc 0 < length lm + + (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm) \ + Suc (Suc n) < length lm + (Suc (max (Suc (Suc n)) (max bc ba)) - + rs_pos + length suf_lm) \ bc < length lm + (Suc (max (Suc (Suc n)) + (max bc ba)) - rs_pos + length suf_lm) \ ba < length lm + + (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm)" +apply(arith) +done + +lemma [simp]: + "\a_md = Suc (max (Suc (Suc n)) (max bc ba)); + length lm = rs_pos \ rs_pos = n \ n > 0\ + \ n - Suc 0 < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \ + Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \ + bc < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \ + ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))" +apply(arith) +done + +lemma [simp]: "n - Suc 0 \ max (Suc (Suc n)) (max bc ba)" +apply(arith) +done + +lemma [simp]: + "a_md \ Suc bc \ rs_pos > 0 \ bc \ rs_pos \ + bc - (rs_pos - Suc 0) + a_md - Suc bc = Suc (a_md - rs_pos - Suc 0)" +apply(arith) +done + +lemma [simp]: "length lm = n \ rs_pos = n \ 0 < rs_pos \ + Suc rs_pos < a_md + \ n - Suc 0 < Suc (Suc (a_md + length suf_lm - Suc (Suc 0))) + \ n < Suc (Suc (a_md + length suf_lm - Suc (Suc 0)))" +apply(arith) +done + +lemma [simp]: "length lm = n \ rs_pos = n \ 0 < rs_pos \ + Suc rs_pos < a_md \ n - Suc 0 \ n" +by arith + +lemma ci_pr_ex2: + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_calc_rel (Pr n f g) lm rs; + rec_ci g = (a, aa, ba); + rec_ci f = (ab, ac, bc)\ + \ \ap bp. aprog = ap [+] bp \ + ap = empty n (max (Suc (Suc (Suc n))) (max bc ba))" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "(ab [+] (recursive.empty n (Suc n) [+] + ([Dec (max (n + 3) (max bc ba)) (length a + 7)] + [+] (a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ + [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]))" in exI, auto) +apply(simp add: abc_append_commute add3_Suc) +done + +lemma [simp]: + "max (Suc (Suc (Suc n))) (max bc ba) - n < + Suc (max (Suc (Suc (Suc n))) (max bc ba)) - n" +apply(arith) +done +lemma exp_nth[simp]: "n < m \ a\<^bsup>m\<^esup> ! n = a" +apply(simp add: exponent_def) +done + +lemma [simp]: "length lm = n \ rs_pos = n \ 0 < n \ + lm[n - Suc 0 := 0::nat] = butlast lm @ [0]" +apply(auto) +apply(insert list_update_append[of "butlast lm" "[last lm]" + "length lm - Suc 0" "0"], simp) +done + +lemma [simp]: "\length lm = n; 0 < n\ \ lm ! (n - Suc 0) = last lm" +apply(insert nth_append[of "butlast lm" "[last lm]" "n - Suc 0"], + simp) +apply(insert butlast_append_last[of lm], auto) +done +lemma exp_suc_iff: "a\<^bsup>b\<^esup> @ [a] = a\<^bsup>b + Suc 0\<^esup>" +apply(simp add: exponent_def rep_ind del: replicate.simps) +done + +lemma less_not_less[simp]: "n > 0 \ \ n < n - Suc 0" +by auto + +lemma [simp]: + "Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \ + bc < Suc (length suf_lm + max (Suc (Suc n)) + (max bc ba)) \ + ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))" + by arith + +lemma [simp]: "length lm = n \ rs_pos = n \ n > 0 \ +(lm @ 0\<^bsup>Suc (max (Suc (Suc n)) (max bc ba)) - n\<^esup> @ suf_lm) + [max (Suc (Suc n)) (max bc ba) := + (lm @ 0\<^bsup>Suc (max (Suc (Suc n)) (max bc ba)) - n\<^esup> @ suf_lm) ! (n - Suc 0) + + (lm @ 0\<^bsup>Suc (max (Suc (Suc n)) (max bc ba)) - n\<^esup> @ suf_lm) ! + max (Suc (Suc n)) (max bc ba), n - Suc 0 := 0::nat] + = butlast lm @ 0 # 0\<^bsup>max (Suc (Suc n)) (max bc ba) - n\<^esup> @ last lm # suf_lm" +apply(simp add: nth_append exp_nth list_update_append) +apply(insert list_update_append[of "0\<^bsup>(max (Suc (Suc n)) (max bc ba)) - n\<^esup>" + "[0]" "max (Suc (Suc n)) (max bc ba) - n" "last lm"], simp) +apply(simp add: exp_suc_iff Suc_diff_le del: list_update.simps) +done + +lemma exp_eq: "(a = b) = (c\<^bsup>a\<^esup> = c\<^bsup>b\<^esup>)" +apply(auto simp: exponent_def) +done + +lemma [simp]: + "\length lm = n; 0 < n; Suc n < a_md\ \ + (butlast lm @ rsa # 0\<^bsup>a_md - Suc n\<^esup> @ last lm # suf_lm) + [n := (butlast lm @ rsa # 0\<^bsup>a_md - Suc n\<^esup> @ last lm # suf_lm) ! + (n - Suc 0) + (butlast lm @ rsa # (0::nat)\<^bsup>a_md - Suc n\<^esup> @ + last lm # suf_lm) ! n, n - Suc 0 := 0] + = butlast lm @ 0 # rsa # 0\<^bsup>a_md - Suc (Suc n)\<^esup> @ last lm # suf_lm" +apply(simp add: nth_append exp_nth list_update_append) +apply(case_tac "a_md - Suc n", simp, simp add: exponent_def) +done + +lemma [simp]: + "Suc (Suc rs_pos) \ a_md \ length lm = rs_pos \ 0 < rs_pos + \ a_md - Suc 0 < + Suc (Suc (Suc (a_md + length suf_lm - Suc (Suc (Suc 0)))))" +by arith + +lemma [simp]: + "Suc (Suc rs_pos) \ a_md \ length lm = rs_pos \ 0 < rs_pos \ + \ a_md - Suc 0 < rs_pos - Suc 0" +by arith + +lemma [simp]: "Suc (Suc rs_pos) \ a_md \ + \ a_md - Suc 0 < rs_pos - Suc 0" +by arith + +lemma [simp]: "\Suc (Suc rs_pos) \ a_md\ \ + \ a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))" +by arith + +lemma [simp]: + "Suc (Suc rs_pos) \ a_md \ length lm = rs_pos \ 0 < rs_pos + \ (abc_lm_v (butlast lm @ last lm # rs # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ + 0 # suf_lm) (a_md - Suc 0) = 0 \ + abc_lm_s (butlast lm @ last lm # rs # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ + 0 # suf_lm) (a_md - Suc 0) 0 = + lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) \ + abc_lm_v (butlast lm @ last lm # rs # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ + 0 # suf_lm) (a_md - Suc 0) = 0" +apply(simp add: abc_lm_v.simps nth_append abc_lm_s.simps) +apply(insert nth_append[of "last lm # rs # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup>" + "0 # suf_lm" "(a_md - rs_pos)"], auto) +apply(simp only: exp_suc_iff) +apply(subgoal_tac "a_md - Suc 0 < a_md + length suf_lm", simp) +apply(case_tac "lm = []", auto) +done + +lemma pr_prog_ex[simp]: "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\ + \ \cp. aprog = recursive.empty n (max (n + 3) + (max bc ba)) [+] cp" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "(ab [+] (recursive.empty n (Suc n) [+] + ([Dec (max (n + 3) (max bc ba)) (length a + 7)] + [+] (a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) + @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]))" in exI) +apply(auto simp: abc_append_commute) +done + +lemma [simp]: "empty m n \ []" +by (simp add: empty.simps) +(* +lemma [simp]: "\rs_pos = n; 0 < rs_pos ; Suc rs_pos < a_md\ \ + n - Suc 0 < a_md + length suf_lm" +by arith +*) +lemma [intro]: + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci f = (ab, ac, bc)\ \ + \ap. (\cp. aprog = ap [+] ab [+] cp) \ length ap = 3" +apply(case_tac "rec_ci g", simp add: rec_ci.simps) +apply(rule_tac x = "empty n + (max (n + 3) (max bc c))" in exI, simp) +apply(rule_tac x = "recursive.empty n (Suc n) [+] + ([Dec (max (n + 3) (max bc c)) (length a + 7)] + [+] a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) + @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" in exI, + auto) +apply(simp add: abc_append_commute) +done + +lemma [intro]: + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci g = (a, aa, ba); + rec_ci f = (ab, ac, bc)\ \ + \ap. (\cp. aprog = ap [+] recursive.empty n (Suc n) [+] cp) + \ length ap = 3 + length ab" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "recursive.empty n (max (n + 3) + (max bc ba)) [+] ab" in exI, simp) +apply(rule_tac x = "([Dec (max (n + 3) (max bc ba)) + (length a + 7)] [+] a [+] + [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ + [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" in exI) +apply(auto simp: abc_append_commute) +done + +(* +lemma [simp]: + "n - Suc 0 < Suc (max (Suc (Suc n)) (max bc ba) + length suf_lm) \ + Suc n < max (Suc (Suc n)) (max bc ba) + length suf_lm \ + bc < Suc (max (Suc (Suc n)) (max bc ba) + length suf_lm) \ + ba < Suc (max (Suc (Suc n)) (max bc ba) + length suf_lm)" +by arith +*) + +lemma [intro]: + "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); + rec_ci g = (a, aa, ba); + rec_ci f = (ab, ac, bc)\ + \ \ap. (\cp. aprog = ap [+] ([Dec (a_md - Suc 0) (length a + 7)] + [+] (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, + Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), + Goto (length a + 4)] [+] cp) \ + length ap = 6 + length ab" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "recursive.empty n + (max (n + 3) (max bc ba)) [+] ab [+] + recursive.empty n (Suc n)" in exI, simp) +apply(rule_tac x = "[]" in exI, auto) +apply(simp add: abc_append_commute) +done + +(* +lemma [simp]: "\rs_pos = n; 0 < rs_pos ; Suc rs_pos < a_md\ \ + n - Suc 0 < Suc (Suc (a_md + length suf_lm - 2)) \ + n < Suc (Suc (a_md + length suf_lm - 2))" +by arith +*) + +lemma [simp]: + "n < Suc (max (n + 3) (max bc ba) + length suf_lm) \ + Suc (Suc n) < max (n + 3) (max bc ba) + length suf_lm \ + bc < Suc (max (n + 3) (max bc ba) + length suf_lm) \ + ba < Suc (max (n + 3) (max bc ba) + length suf_lm)" +by arith + +lemma [simp]: "n \ max (n + (3::nat)) (max bc ba)" +by arith + +lemma [simp]:"length lm = Suc n \ lm[n := (0::nat)] = butlast lm @ [0]" +apply(subgoal_tac "\ xs x. lm = xs @ [x]", auto simp: list_update_append) +apply(rule_tac x = "butlast lm" in exI, rule_tac x = "last lm" in exI) +apply(case_tac lm, auto) +done + +lemma [simp]: "length lm = Suc n \ lm ! n =last lm" +apply(subgoal_tac "lm \ []") +apply(simp add: last_conv_nth, case_tac lm, simp_all) +done + +lemma [simp]: "length lm = Suc n \ + (lm @ (0::nat)\<^bsup>max (n + 3) (max bc ba) - n\<^esup> @ suf_lm) + [max (n + 3) (max bc ba) := (lm @ 0\<^bsup>max (n + 3) (max bc ba) - n\<^esup> @ suf_lm) ! n + + (lm @ 0\<^bsup>max (n + 3) (max bc ba) - n\<^esup> @ suf_lm) ! max (n + 3) (max bc ba), n := 0] + = butlast lm @ 0 # 0\<^bsup>max (n + 3) (max bc ba) - Suc n\<^esup> @ last lm # suf_lm" +apply(auto simp: list_update_append nth_append) +apply(subgoal_tac "(0\<^bsup>max (n + 3) (max bc ba) - n\<^esup>) = 0\<^bsup>max (n + 3) (max bc ba) - Suc n\<^esup> @ [0::nat]") +apply(simp add: list_update_append) +apply(simp add: exp_suc_iff) +done + +lemma [simp]: "Suc (Suc n) < a_md \ + n < Suc (Suc (a_md + length suf_lm - 2)) \ + n < Suc (a_md + length suf_lm - 2)" +by(arith) + +lemma [simp]: "\length lm = Suc n; Suc (Suc n) < a_md\ + \(butlast lm @ (rsa::nat) # 0\<^bsup>a_md - Suc (Suc n)\<^esup> @ last lm # suf_lm) + [Suc n := (butlast lm @ rsa # 0\<^bsup>a_md - Suc (Suc n)\<^esup> @ last lm # suf_lm) ! n + + (butlast lm @ rsa # 0\<^bsup>a_md - Suc (Suc n)\<^esup> @ last lm # suf_lm) ! Suc n, n := 0] + = butlast lm @ 0 # rsa # 0\<^bsup>a_md - Suc (Suc (Suc n))\<^esup> @ last lm # suf_lm" +apply(auto simp: list_update_append) +apply(subgoal_tac "(0\<^bsup>a_md - Suc (Suc n)\<^esup>) = (0::nat) # (0\<^bsup>a_md - Suc (Suc (Suc n))\<^esup>)", simp add: nth_append) +apply(simp add: exp_ind_def[THEN sym]) +done + +lemma pr_case: + assumes nf_ind: + "\ lm rs suf_lm. rec_calc_rel f lm rs \ + \stp. abc_steps_l (0, lm @ 0\<^bsup>bc - ac\<^esup> @ suf_lm) ab stp = + (length ab, lm @ rs # 0\<^bsup>bc - Suc ac\<^esup> @ suf_lm)" + and ng_ind: "\ lm rs suf_lm. rec_calc_rel g lm rs \ + \stp. abc_steps_l (0, lm @ 0\<^bsup>ba - aa\<^esup> @ suf_lm) a stp = + (length a, lm @ rs # 0\<^bsup>ba - Suc aa\<^esup> @ suf_lm)" + and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" "rec_calc_rel (Pr n f g) lm rs" + "rec_ci g = (a, aa, ba)" "rec_ci f = (ab, ac, bc)" + shows "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" +proof - + from h have k1: "\ stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp + = (3, butlast lm @ 0 # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ last lm # suf_lm)" + proof - + have "\bp cp. aprog = bp [+] cp \ bp = empty n + (max (n + 3) (max bc ba))" + apply(insert h, simp) + apply(erule pr_prog_ex, auto) + done + thus "?thesis" + apply(erule_tac exE, erule_tac exE, simp) + apply(subgoal_tac + "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) + ([] [+] recursive.empty n + (max (n + 3) (max bc ba)) [+] cp) stp = + (0 + 3, butlast lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ + last lm # suf_lm)", simp) + apply(rule_tac abc_append_exc1, simp_all) + apply(insert empty_ex[of "n" "(max (n + 3) + (max bc ba))" "lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm"], simp) + apply(subgoal_tac "a_md = Suc (max (n + 3) (max bc ba))", + simp) + apply(subgoal_tac "length lm = Suc n \ rs_pos = Suc n", simp) + apply(insert h) + apply(simp add: para_pattern ci_pr_para_eq) + apply(rule ci_pr_md_def, auto) + done + qed + from h have k2: + "\ stp. abc_steps_l (3, butlast lm @ 0 # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ + last lm # suf_lm) aprog stp + = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + proof - + from h have k2_1: "\ rs. rec_calc_rel f (butlast lm) rs" + apply(erule_tac calc_pr_zero_ex) + done + thus "?thesis" + proof(erule_tac exE) + fix rsa + assume k2_2: "rec_calc_rel f (butlast lm) rsa" + from h and k2_2 have k2_2_1: + "\ stp. abc_steps_l (3, butlast lm @ 0 # 0\<^bsup>a_md - rs_pos - 1\<^esup> + @ last lm # suf_lm) aprog stp + = (3 + length ab, butlast lm @ rsa # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ + last lm # suf_lm)" + proof - + from h have j1: " + \ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = 3 \ + bp = ab" + apply(auto) + done + from h have j2: "ac = rs_pos - 1" + apply(drule_tac ci_pr_f_paras, simp, auto) + done + from h and j2 have j3: "a_md \ Suc bc \ rs_pos > 0 \ bc \ rs_pos" + apply(rule_tac conjI) + apply(erule_tac ab = ab and ac = ac in ci_pr_md_ge_f, simp) + apply(rule_tac context_conjI) + apply(simp_all add: rec_ci.simps) + apply(drule_tac ci_ad_ge_paras, drule_tac ci_ad_ge_paras) + apply(arith) + done + from j1 and j2 show "?thesis" + apply(auto simp del: abc_append_commute) + apply(rule_tac abc_append_exc1, simp_all) + apply(insert nf_ind[of "butlast lm" "rsa" + "0\<^bsup>a_md - bc - Suc 0\<^esup> @ last lm # suf_lm"], + simp add: k2_2 j2, erule_tac exE) + apply(simp add: exponent_add_iff j3) + apply(rule_tac x = "stp" in exI, simp) + done + qed + from h have k2_2_2: + "\ stp. abc_steps_l (3 + length ab, butlast lm @ rsa # + 0\<^bsup>a_md - rs_pos - 1\<^esup> @ last lm # suf_lm) aprog stp + = (6 + length ab, butlast lm @ 0 # rsa # + 0\<^bsup>a_md - rs_pos - 2\<^esup> @ last lm # suf_lm)" + proof - + from h have "\ ap bp cp. aprog = ap [+] bp [+] cp \ + length ap = 3 + length ab \ bp = recursive.empty n (Suc n)" + by auto + thus "?thesis" + proof(erule_tac exE, erule_tac exE, erule_tac exE, + erule_tac exE) + fix ap cp bp apa + assume "aprog = ap [+] bp [+] cp \ length ap = 3 + + length ab \ bp = recursive.empty n (Suc n)" + thus "?thesis" + apply(simp del: abc_append_commute) + apply(subgoal_tac + "\stp. abc_steps_l (3 + length ab, + butlast lm @ rsa # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ + last lm # suf_lm) (ap [+] + recursive.empty n (Suc n) [+] cp) stp = + ((3 + length ab) + 3, butlast lm @ 0 # rsa # + 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ last lm # suf_lm)", simp) + apply(rule_tac abc_append_exc1, simp_all) + apply(insert empty_ex[of n "Suc n" + "butlast lm @ rsa # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ + last lm # suf_lm"], simp) + apply(subgoal_tac "length lm = Suc n \ rs_pos = Suc n \ a_md > Suc (Suc n)", simp) + apply(insert h, simp) + done + qed + qed + from h have k2_3: "lm \ []" + apply(rule_tac calc_pr_para_not_null, simp) + done + from h and k2_2 and k2_3 have k2_2_3: + "\ stp. abc_steps_l (6 + length ab, butlast lm @ + (last lm - last lm) # rsa # + 0\<^bsup>a_md - (Suc (Suc rs_pos))\<^esup> @ last lm # suf_lm) aprog stp + = (6 + length ab, butlast lm @ last lm # rs # + 0\<^bsup>a_md - Suc (Suc (rs_pos))\<^esup> @ 0 # suf_lm)" + apply(rule_tac x = "last lm" and g = g in pr_cycle_part, auto) + apply(rule_tac ng_ind, simp) + apply(rule_tac rec_calc_rel_def0, simp, simp) + done + from h have k2_2_4: + "\ stp. abc_steps_l (6 + length ab, + butlast lm @ last lm # rs # 0\<^bsup>a_md - rs_pos - 2\<^esup> @ + 0 # suf_lm) aprog stp + = (13 + length ab + length a, + lm @ rs # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + proof - + from h have + "\ ap bp cp. aprog = ap [+] bp [+] cp \ + length ap = 6 + length ab \ + bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] + (a [+] [Inc (rs_pos - Suc 0), + Dec rs_pos 3, Goto (Suc 0)])) @ + [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" + by auto + thus "?thesis" + apply(auto) + apply(subgoal_tac + "\stp. abc_steps_l (6 + length ab, butlast lm @ + last lm # rs # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm) + (ap [+] ([Dec (a_md - Suc 0) (length a + 7)] [+] + (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, + Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), + Goto (length a + 4)] [+] cp) stp = + (6 + length ab + (length a + 7) , + lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)", simp) + apply(subgoal_tac "13 + (length ab + length a) = + 13 + length ab + length a", simp) + apply(arith) + apply(rule abc_append_exc1, simp_all) + apply(rule_tac x = "Suc 0" in exI, + simp add: abc_steps_l.simps abc_fetch.simps + nth_append abc_append_nth abc_step_l.simps) + apply(subgoal_tac "a_md > Suc (Suc rs_pos) \ + length lm = rs_pos \ rs_pos > 0", simp) + apply(insert h, simp) + apply(subgoal_tac "rs_pos = Suc n", simp, simp) + done + qed + from h have k2_2_5: "length aprog = 13 + length ab + length a" + apply(rule_tac ci_pr_length, simp_all) + done + from k2_2_1 and k2_2_2 and k2_2_3 and k2_2_4 and k2_2_5 + show "?thesis" + apply(auto) + apply(rule_tac x = "stp + stpa + stpb + stpc" in exI, + simp add: abc_steps_add) + done + qed + qed + from k1 and k2 show + "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp + = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + apply(erule_tac exE) + apply(erule_tac exE) + apply(rule_tac x = "stp + stpa" in exI) + apply(simp add: abc_steps_add) + done +qed + +thm rec_calc_rel.induct + +lemma eq_switch: "x = y \ y = x" +by simp + +lemma [simp]: + "\rec_ci f = (a, aa, ba); + rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ \ \bp. aprog = a @ bp" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "[Dec (Suc n) (length a + 5), + Dec (Suc n) (length a + 3), Goto (Suc (length a)), + Inc n, Goto 0]" in exI, auto) +done + +lemma ci_mn_para_eq[simp]: + "rec_ci (Mn n f) = (aprog, rs_pos, a_md) \ rs_pos = n" +apply(case_tac "rec_ci f", simp add: rec_ci.simps) +done +(* +lemma [simp]: "\rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md); rec_calc_rel (Mn n f) lm rs\ \ aa = Suc rs_pos" +apply(rule_tac calc_mn_reverse, simp) +apply(insert para_pattern [of f a aa ba "lm @ [rs]" 0], simp) +apply(subgoal_tac "rs_pos = length lm", simp) +apply(drule_tac ci_mn_para_eq, simp) +done +*) +lemma [simp]: "rec_ci f = (a, aa, ba) \ aa < ba" +apply(simp add: ci_ad_ge_paras) +done + +lemma [simp]: "\rec_ci f = (a, aa, ba); + rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ + \ ba \ a_md" +apply(simp add: rec_ci.simps) +by arith + +lemma mn_calc_f: + assumes ind: + "\aprog a_md rs_pos rs suf_lm lm. + \rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\ + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp + = (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + and h: "rec_ci f = (a, aa, ba)" + "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" + "rec_calc_rel f (lm @ [x]) rsx" + "aa = Suc n" + shows "\stp. abc_steps_l (0, lm @ x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) + aprog stp = (length a, + lm @ x # rsx # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ suf_lm)" +proof - + from h have k1: "\ ap bp. aprog = ap @ bp \ ap = a" + by simp + from h have k2: "rs_pos = n" + apply(erule_tac ci_mn_para_eq) + done + from h and k1 and k2 show "?thesis" + + proof(erule_tac exE, erule_tac exE, simp, + rule_tac abc_add_exc1, auto) + fix bp + show + "\astp. abc_steps_l (0, lm @ x # 0\<^bsup>a_md - Suc n\<^esup> @ suf_lm) a astp + = (length a, lm @ x # rsx # 0\<^bsup>a_md - Suc (Suc n)\<^esup> @ suf_lm)" + apply(insert ind[of a "Suc n" ba "lm @ [x]" rsx + "0\<^bsup>a_md - ba\<^esup> @ suf_lm"], simp add: exponent_add_iff h k2) + apply(subgoal_tac "ba > aa \ a_md \ ba \ aa = Suc n", + insert h, auto) + done + qed +qed +thm rec_ci.simps + +fun mn_ind_inv :: + "nat \ nat list \ nat \ nat \ nat \ nat list \ nat list \ bool" + where + "mn_ind_inv (as, lm') ss x rsx suf_lm lm = + (if as = ss then lm' = lm @ x # rsx # suf_lm + else if as = ss + 1 then + \y. (lm' = lm @ x # y # suf_lm) \ y \ rsx + else if as = ss + 2 then + \y. (lm' = lm @ x # y # suf_lm) \ y \ rsx + else if as = ss + 3 then lm' = lm @ x # 0 # suf_lm + else if as = ss + 4 then lm' = lm @ Suc x # 0 # suf_lm + else if as = 0 then lm' = lm @ Suc x # 0 # suf_lm + else False +)" + +fun mn_stage1 :: "nat \ nat list \ nat \ nat \ nat" + where + "mn_stage1 (as, lm) ss n = + (if as = 0 then 0 + else if as = ss + 4 then 1 + else if as = ss + 3 then 2 + else if as = ss + 2 \ as = ss + 1 then 3 + else if as = ss then 4 + else 0 +)" + +fun mn_stage2 :: "nat \ nat list \ nat \ nat \ nat" + where + "mn_stage2 (as, lm) ss n = + (if as = ss + 1 \ as = ss + 2 then (lm ! (Suc n)) + else 0)" + +fun mn_stage3 :: "nat \ nat list \ nat \ nat \ nat" + where + "mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)" + + +fun mn_measure :: "((nat \ nat list) \ nat \ nat) \ + (nat \ nat \ nat)" + where + "mn_measure ((as, lm), ss, n) = + (mn_stage1 (as, lm) ss n, mn_stage2 (as, lm) ss n, + mn_stage3 (as, lm) ss n)" + +definition mn_LE :: "(((nat \ nat list) \ nat \ nat) \ + ((nat \ nat list) \ nat \ nat)) set" + where "mn_LE \ (inv_image lex_triple mn_measure)" + +thm halt_lemma2 +lemma wf_mn_le[intro]: "wf mn_LE" +by(auto intro:wf_inv_image wf_lex_triple simp: mn_LE_def) + +declare mn_ind_inv.simps[simp del] + +lemma mn_inv_init: + "mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0) + (length a) x rsx suf_lm lm" +apply(simp add: mn_ind_inv.simps abc_steps_zero) +done + +lemma mn_halt_init: + "rec_ci f = (a, aa, ba) \ + \ (\(as, lm') (ss, n). as = 0) + (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0) + (length a, n)" +apply(simp add: abc_steps_zero) +apply(erule_tac rec_ci_not_null) +done + +thm rec_ci.simps +lemma [simp]: + "\rec_ci f = (a, aa, ba); + rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ + \ abc_fetch (length a) aprog = + Some (Dec (Suc n) (length a + 5))" +apply(simp add: rec_ci.simps abc_fetch.simps, + erule_tac conjE, erule_tac conjE, simp) +apply(drule_tac eq_switch, drule_tac eq_switch, simp) +done + +lemma [simp]: "\rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ + \ abc_fetch (Suc (length a)) aprog = Some (Dec (Suc n) (length a + 3))" +apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp) +apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append) +done + +lemma [simp]: + "\rec_ci f = (a, aa, ba); + rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ + \ abc_fetch (Suc (Suc (length a))) aprog = + Some (Goto (length a + 1))" +apply(simp add: rec_ci.simps abc_fetch.simps, + erule_tac conjE, erule_tac conjE, simp) +apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append) +done + +lemma [simp]: + "\rec_ci f = (a, aa, ba); + rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ + \ abc_fetch (length a + 3) aprog = Some (Inc n)" +apply(simp add: rec_ci.simps abc_fetch.simps, + erule_tac conjE, erule_tac conjE, simp) +apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append) +done + +lemma [simp]: "\rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ + \ abc_fetch (length a + 4) aprog = Some (Goto 0)" +apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp) +apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append) +done + +lemma [simp]: + "0 < rsx + \ \y. (lm @ x # rsx # suf_lm)[Suc (length lm) := rsx - Suc 0] + = lm @ x # y # suf_lm \ y \ rsx" +apply(case_tac rsx, simp, simp) +apply(rule_tac x = nat in exI, simp add: list_update_append) +done + +lemma [simp]: + "\y \ rsx; 0 < y\ + \ \ya. (lm @ x # y # suf_lm)[Suc (length lm) := y - Suc 0] + = lm @ x # ya # suf_lm \ ya \ rsx" +apply(case_tac y, simp, simp) +apply(rule_tac x = nat in exI, simp add: list_update_append) +done + +lemma mn_halt_lemma: + "\rec_ci f = (a, aa, ba); + rec_ci (Mn n f) = (aprog, rs_pos, a_md); + 0 < rsx; length lm = n\ + \ + \na. \ (\(as, lm') (ss, n). as = 0) + (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na) + (length a, n) + \ mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) + aprog na) (length a) x rsx suf_lm lm +\ mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog + (Suc na)) (length a) x rsx suf_lm lm + \ ((abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog (Suc na), + length a, n), + abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na, + length a, n) \ mn_LE" +apply(rule allI, rule impI, simp add: abc_steps_ind) +apply(case_tac "(abc_steps_l (length a, lm @ x # rsx # suf_lm) + aprog na)", simp) +apply(auto split:if_splits simp add:abc_steps_l.simps + mn_ind_inv.simps abc_steps_zero) +apply(auto simp add: mn_LE_def lex_triple_def lex_pair_def + abc_step_l.simps abc_steps_l.simps mn_ind_inv.simps + abc_lm_v.simps abc_lm_s.simps nth_append + split: if_splits) +apply(drule_tac rec_ci_not_null, simp) +done + +lemma mn_halt: + "\rec_ci f = (a, aa, ba); + rec_ci (Mn n f) = (aprog, rs_pos, a_md); + 0 < rsx; length lm = n\ + \ \ stp. (\ (as, lm'). (as = 0 \ + mn_ind_inv (as, lm') (length a) x rsx suf_lm lm)) + (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp)" +apply(insert wf_mn_le) +apply(insert halt_lemma2[of mn_LE + "\ ((as, lm'), ss, n). as = 0" + "\ stp. (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp, + length a, n)" + "\ ((as, lm'), ss, n). mn_ind_inv (as, lm') ss x rsx suf_lm lm"], + simp) +apply(simp add: mn_halt_init mn_inv_init) +apply(drule_tac x = x and suf_lm = suf_lm in mn_halt_lemma, auto) +apply(rule_tac x = n in exI, + case_tac "(abc_steps_l (length a, lm @ x # rsx # suf_lm) + aprog n)", simp) +done + +lemma [simp]: "Suc rs_pos < a_md \ + Suc (a_md - Suc (Suc rs_pos)) = a_md - Suc rs_pos" +by arith + +term rec_ci +(* +lemma [simp]: "\rec_ci (Mn n f) = (aprog, rs_pos, a_md); rec_calc_rel (Mn n f) lm rs\ \ Suc rs_pos < a_md" +apply(case_tac "rec_ci f") +apply(subgoal_tac "c > b \ b = Suc rs_pos \ a_md \ c") +apply(arith, auto) +done +*) +lemma mn_ind_step: + assumes ind: + "\aprog a_md rs_pos rs suf_lm lm. + \rec_ci f = (aprog, rs_pos, a_md); + rec_calc_rel f lm rs\ \ + \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp + = (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + and h: "rec_ci f = (a, aa, ba)" + "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" + "rec_calc_rel f (lm @ [x]) rsx" + "rsx > 0" + "Suc rs_pos < a_md" + "aa = Suc rs_pos" + shows "\stp. abc_steps_l (0, lm @ x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) + aprog stp = (0, lm @ Suc x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" +thm abc_add_exc1 +proof - + have k1: + "\ stp. abc_steps_l (0, lm @ x # 0\<^bsup>a_md - Suc (rs_pos)\<^esup> @ suf_lm) + aprog stp = + (length a, lm @ x # rsx # 0\<^bsup>a_md - Suc (Suc rs_pos) \<^esup>@ suf_lm)" + apply(insert h) + apply(auto intro: mn_calc_f ind) + done + from h have k2: "length lm = n" + apply(subgoal_tac "rs_pos = n") + apply(drule_tac para_pattern, simp, simp, simp) + done + from h have k3: "a_md > (Suc rs_pos)" + apply(simp) + done + from k2 and h and k3 have k4: + "\ stp. abc_steps_l (length a, + lm @ x # rsx # 0\<^bsup>a_md - Suc (Suc rs_pos) \<^esup>@ suf_lm) aprog stp = + (0, lm @ Suc x # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + apply(frule_tac x = x and + suf_lm = "0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ suf_lm" in mn_halt, auto) + apply(rule_tac x = "stp" in exI, + simp add: mn_ind_inv.simps rec_ci_not_null exponent_def) + apply(simp only: replicate.simps[THEN sym], simp) + done + + from k1 and k4 show "?thesis" + apply(auto) + apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add) + done +qed + +lemma [simp]: + "\rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md); + rec_calc_rel (Mn n f) lm rs\ \ aa = Suc rs_pos" +apply(rule_tac calc_mn_reverse, simp) +apply(insert para_pattern [of f a aa ba "lm @ [rs]" 0], simp) +apply(subgoal_tac "rs_pos = length lm", simp) +apply(drule_tac ci_mn_para_eq, simp) +done + +lemma [simp]: "\rec_ci (Mn n f) = (aprog, rs_pos, a_md); + rec_calc_rel (Mn n f) lm rs\ \ Suc rs_pos < a_md" +apply(case_tac "rec_ci f") +apply(subgoal_tac "c > b \ b = Suc rs_pos \ a_md \ c") +apply(arith, auto) +done + +lemma mn_ind_steps: + assumes ind: + "\aprog a_md rs_pos rs suf_lm lm. + \rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\ \ + \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + and h: "rec_ci f = (a, aa, ba)" + "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Mn n f) lm rs" + "rec_calc_rel f (lm @ [rs]) 0" + "\x v. rec_calc_rel f (lm @ [x]) v \ 0 < v)" + "n = length lm" + "x \ rs" + shows "\stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) + aprog stp = (0, lm @ x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" +apply(insert h, induct x, + rule_tac x = 0 in exI, simp add: abc_steps_zero, simp) +proof - + fix x + assume k1: + "\stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) + aprog stp = (0, lm @ x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + and k2: "rec_ci (Mn (length lm) f) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Mn (length lm) f) lm rs" + "rec_calc_rel f (lm @ [rs]) 0" + "\x v. rec_calc_rel f (lm @ [x]) v \ v > 0)" + "n = length lm" + "Suc x \ rs" + "rec_ci f = (a, aa, ba)" + hence k2: + "\stp. abc_steps_l (0, lm @ x # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm) aprog + stp = (0, lm @ Suc x # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + apply(erule_tac x = x in allE) + apply(auto) + apply(rule_tac x = x in mn_ind_step) + apply(rule_tac ind, auto) + done + from k1 and k2 show + "\stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) + aprog stp = (0, lm @ Suc x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + apply(auto) + apply(rule_tac x = "stp + stpa" in exI, simp only: abc_steps_add) + done +qed + +lemma [simp]: +"\rec_ci f = (a, aa, ba); + rec_ci (Mn n f) = (aprog, rs_pos, a_md); + rec_calc_rel (Mn n f) lm rs; + length lm = n\ + \ abc_lm_v (lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) (Suc n) = 0" +apply(auto simp: abc_lm_v.simps nth_append) +done + +lemma [simp]: + "\rec_ci f = (a, aa, ba); + rec_ci (Mn n f) = (aprog, rs_pos, a_md); + rec_calc_rel (Mn n f) lm rs; + length lm = n\ + \ abc_lm_s (lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) (Suc n) 0 = + lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm" +apply(auto simp: abc_lm_s.simps list_update_append) +done + +lemma mn_length: + "\rec_ci f = (a, aa, ba); + rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ + \ length aprog = length a + 5" +apply(simp add: rec_ci.simps, erule_tac conjE) +apply(drule_tac eq_switch, drule_tac eq_switch, simp) +done + +lemma mn_final_step: + assumes ind: + "\aprog a_md rs_pos rs suf_lm lm. + \rec_ci f = (aprog, rs_pos, a_md); + rec_calc_rel f lm rs\ \ + \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + and h: "rec_ci f = (a, aa, ba)" + "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Mn n f) lm rs" + "rec_calc_rel f (lm @ [rs]) 0" + shows "\stp. abc_steps_l (0, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) + aprog stp = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" +proof - + from h and ind have k1: + "\stp. abc_steps_l (0, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) + aprog stp = (length a, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + thm mn_calc_f + apply(insert mn_calc_f[of f a aa ba n aprog + rs_pos a_md lm rs 0 suf_lm], simp) + apply(subgoal_tac "aa = Suc n", simp add: exponent_cons_iff) + apply(subgoal_tac "rs_pos = n", simp, simp) + done + from h have k2: "length lm = n" + apply(subgoal_tac "rs_pos = n") + apply(drule_tac f = "Mn n f" in para_pattern, simp, simp, simp) + done + from h and k2 have k3: + "\stp. abc_steps_l (length a, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) + aprog stp = (length a + 5, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + apply(rule_tac x = "Suc 0" in exI, + simp add: abc_step_l.simps abc_steps_l.simps) + done + from h have k4: "length aprog = length a + 5" + apply(simp add: mn_length) + done + from k1 and k3 and k4 show "?thesis" + apply(auto) + apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add) + done +qed + +lemma mn_case: + assumes ind: + "\aprog a_md rs_pos rs suf_lm lm. + \rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\ \ + \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + and h: "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Mn n f) lm rs" + shows "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp + = (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" +apply(case_tac "rec_ci f", simp) +apply(insert h, rule_tac calc_mn_reverse, simp) +proof - + fix a b c v + assume h: "rec_ci f = (a, b, c)" + "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Mn n f) lm rs" + "rec_calc_rel f (lm @ [rs]) 0" + "\xv. rec_calc_rel f (lm @ [x]) v \ 0 < v" + "n = length lm" + hence k1: + "\stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) aprog + stp = (0, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + thm mn_ind_steps + apply(auto intro: mn_ind_steps ind) + done + from h have k2: + "\stp. abc_steps_l (0, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) aprog + stp = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + apply(auto intro: mn_final_step ind) + done + from k1 and k2 show + "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + apply(auto, insert h) + apply(subgoal_tac "Suc rs_pos < a_md") + apply(rule_tac x = "stp + stpa" in exI, + simp only: abc_steps_add exponent_cons_iff, simp, simp) + done +qed + +lemma z_rs: "rec_calc_rel z lm rs \ rs = 0" +apply(rule_tac calc_z_reverse, auto) +done + +lemma z_case: + "\rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm rs\ + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" +apply(simp add: rec_ci.simps rec_ci_z_def, auto) +apply(rule_tac x = "Suc 0" in exI, simp add: abc_steps_l.simps + abc_fetch.simps abc_step_l.simps z_rs) +done +thm addition.simps + +thm addition.simps +thm rec_ci_s_def +fun addition_inv :: "nat \ nat list \ nat \ nat \ nat \ + nat list \ bool" + where + "addition_inv (as, lm') m n p lm = + (let sn = lm ! n in + let sm = lm ! m in + lm ! p = 0 \ + (if as = 0 then \ x. x \ lm ! m \ lm' = lm[m := x, + n := (sn + sm - x), p := (sm - x)] + else if as = 1 then \ x. x < lm ! m \ lm' = lm[m := x, + n := (sn + sm - x - 1), p := (sm - x - 1)] + else if as = 2 then \ x. x < lm ! m \ lm' = lm[m := x, + n := (sn + sm - x), p := (sm - x - 1)] + else if as = 3 then \ x. x < lm ! m \ lm' = lm[m := x, + n := (sn + sm - x), p := (sm - x)] + else if as = 4 then \ x. x \ lm ! m \ lm' = lm[m := x, + n := (sn + sm), p := (sm - x)] + else if as = 5 then \ x. x < lm ! m \ lm' = lm[m := x, + n := (sn + sm), p := (sm - x - 1)] + else if as = 6 then \ x. x < lm ! m \ lm' = + lm[m := Suc x, n := (sn + sm), p := (sm - x - 1)] + else if as = 7 then lm' = lm[m := sm, n := (sn + sm)] + else False))" + +fun addition_stage1 :: "nat \ nat list \ nat \ nat \ nat" + where + "addition_stage1 (as, lm) m p = + (if as = 0 \ as = 1 \ as = 2 \ as = 3 then 2 + else if as = 4 \ as = 5 \ as = 6 then 1 + else 0)" + +fun addition_stage2 :: "nat \ nat list \ nat \ nat \ nat" + where + "addition_stage2 (as, lm) m p = + (if 0 \ as \ as \ 3 then lm ! m + else if 4 \ as \ as \ 6 then lm ! p + else 0)" + +fun addition_stage3 :: "nat \ nat list \ nat \ nat \ nat" + where + "addition_stage3 (as, lm) m p = + (if as = 1 then 4 + else if as = 2 then 3 + else if as = 3 then 2 + else if as = 0 then 1 + else if as = 5 then 2 + else if as = 6 then 1 + else if as = 4 then 0 + else 0)" + +fun addition_measure :: "((nat \ nat list) \ nat \ nat) \ + (nat \ nat \ nat)" + where + "addition_measure ((as, lm), m, p) = + (addition_stage1 (as, lm) m p, + addition_stage2 (as, lm) m p, + addition_stage3 (as, lm) m p)" + +definition addition_LE :: "(((nat \ nat list) \ nat \ nat) \ + ((nat \ nat list) \ nat \ nat)) set" + where "addition_LE \ (inv_image lex_triple addition_measure)" + +lemma [simp]: "wf addition_LE" +by(simp add: wf_inv_image wf_lex_triple addition_LE_def) + +declare addition_inv.simps[simp del] + +lemma addition_inv_init: + "\m \ n; max m n < p; length lm > p; lm ! p = 0\ \ + addition_inv (0, lm) m n p lm" +apply(simp add: addition_inv.simps) +apply(rule_tac x = "lm ! m" in exI, simp) +done + +thm addition.simps + +lemma [simp]: "abc_fetch 0 (addition m n p) = Some (Dec m 4)" +by(simp add: abc_fetch.simps addition.simps) + +lemma [simp]: "abc_fetch (Suc 0) (addition m n p) = Some (Inc n)" +by(simp add: abc_fetch.simps addition.simps) + +lemma [simp]: "abc_fetch 2 (addition m n p) = Some (Inc p)" +by(simp add: abc_fetch.simps addition.simps) + +lemma [simp]: "abc_fetch 3 (addition m n p) = Some (Goto 0)" +by(simp add: abc_fetch.simps addition.simps) + +lemma [simp]: "abc_fetch 4 (addition m n p) = Some (Dec p 7)" +by(simp add: abc_fetch.simps addition.simps) + +lemma [simp]: "abc_fetch 5 (addition m n p) = Some (Inc m)" +by(simp add: abc_fetch.simps addition.simps) + +lemma [simp]: "abc_fetch 6 (addition m n p) = Some (Goto 4)" +by(simp add: abc_fetch.simps addition.simps) + +lemma [simp]: + "\m \ n; p < length lm; lm ! p = 0; m < p; n < p; x \ lm ! m; 0 < x\ + \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ + \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ + \ \xam \ n; p < length lm; lm ! p = (0::nat); m < p; n < p; x < lm ! m\ + \ \xa\lm ! m. lm[m := x, n := lm ! n + lm ! m - x, + p := lm ! m - x] = + lm[m := xa, n := lm ! n + lm ! m - xa, + p := lm ! m - xa]" +apply(rule_tac x = x in exI, simp) +done + +lemma [simp]: + "\m \ n; p < length lm; lm ! p = 0; m < p; n < p; + x \ lm ! m; lm ! m \ x\ + \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ + \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ + \ \xa\lm ! m. lm[m := Suc x, n := lm ! n + lm ! m, + p := lm ! m - Suc x] + = lm[m := xa, n := lm ! n + lm ! m, p := lm ! m - xa]" +apply(rule_tac x = "Suc x" in exI, simp) +done + +lemma addition_halt_lemma: + "\m \ n; max m n < p; length lm > p; lm ! p = 0\ \ + \na. \ (\(as, lm') (m, p). as = 7) + (abc_steps_l (0, lm) (addition m n p) na) (m, p) \ + addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm +\ addition_inv (abc_steps_l (0, lm) (addition m n p) + (Suc na)) m n p lm + \ ((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p), + abc_steps_l (0, lm) (addition m n p) na, m, p) \ addition_LE" +apply(rule allI, rule impI, simp add: abc_steps_ind) +apply(case_tac "(abc_steps_l (0, lm) (addition m n p) na)", simp) +apply(auto split:if_splits simp add: addition_inv.simps + abc_steps_zero) +apply(simp_all add: abc_steps_l.simps abc_steps_zero) +apply(auto simp add: addition_LE_def lex_triple_def lex_pair_def + abc_step_l.simps addition_inv.simps + abc_lm_v.simps abc_lm_s.simps nth_append + split: if_splits) +apply(rule_tac x = x in exI, simp) +done + +lemma addition_ex: + "\m \ n; max m n < p; length lm > p; lm ! p = 0\ \ + \ stp. (\ (as, lm'). as = 7 \ addition_inv (as, lm') m n p lm) + (abc_steps_l (0, lm) (addition m n p) stp)" +apply(insert halt_lemma2[of addition_LE + "\ ((as, lm'), m, p). as = 7" + "\ stp. (abc_steps_l (0, lm) (addition m n p) stp, m, p)" + "\ ((as, lm'), m, p). addition_inv (as, lm') m n p lm"], + simp add: abc_steps_zero addition_inv_init) +apply(drule_tac addition_halt_lemma, simp, simp, simp, + simp, erule_tac exE) +apply(rule_tac x = na in exI, + case_tac "(abc_steps_l (0, lm) (addition m n p) na)", auto) +done + +lemma [simp]: "length (addition m n p) = 7" +by (simp add: addition.simps) + +lemma [elim]: "addition 0 (Suc 0) 2 = [] \ RR" +by(simp add: addition.simps) + +lemma [simp]: "(0\<^bsup>2\<^esup>)[0 := n] = [n, 0::nat]" +apply(subgoal_tac "2 = Suc 1", + simp only: replicate.simps exponent_def) +apply(auto) +done + +lemma [simp]: + "\stp. abc_steps_l (0, n # 0\<^bsup>2\<^esup> @ suf_lm) + (addition 0 (Suc 0) 2 [+] [Inc (Suc 0)]) stp = + (8, n # Suc n # 0 # suf_lm)" +apply(rule_tac bm = "n # n # 0 # suf_lm" in abc_append_exc2, auto) +apply(insert addition_ex[of 0 "Suc 0" 2 "n # 0\<^bsup>2\<^esup> @ suf_lm"], + simp add: nth_append numeral_2_eq_2, erule_tac exE) +apply(rule_tac x = stp in exI, + case_tac "(abc_steps_l (0, n # 0\<^bsup>2\<^esup> @ suf_lm) + (addition 0 (Suc 0) 2) stp)", + simp add: addition_inv.simps nth_append list_update_append numeral_2_eq_2) +apply(simp add: nth_append numeral_2_eq_2, erule_tac exE) +apply(rule_tac x = "Suc 0" in exI, + simp add: abc_steps_l.simps abc_fetch.simps + abc_steps_zero abc_step_l.simps abc_lm_s.simps abc_lm_v.simps) +done + +lemma s_case: + "\rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm rs\ + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" +apply(simp add: rec_ci.simps rec_ci_s_def, auto) +apply(rule_tac calc_s_reverse, auto) +done + +lemma [simp]: + "\n < length lm; lm ! n = rs\ + \ \stp. abc_steps_l (0, lm @ 0 # 0 #suf_lm) + (addition n (length lm) (Suc (length lm))) stp + = (7, lm @ rs # 0 # suf_lm)" +apply(insert addition_ex[of n "length lm" + "Suc (length lm)" "lm @ 0 # 0 # suf_lm"]) +apply(simp add: nth_append, erule_tac exE) +apply(rule_tac x = stp in exI) +apply(case_tac "abc_steps_l (0, lm @ 0 # 0 # suf_lm) (addition n (length lm) + (Suc (length lm))) stp", simp) +apply(simp add: addition_inv.simps) +apply(insert nth_append[of lm "0 # 0 # suf_lm" "n"], simp) +done + +lemma [simp]: "0\<^bsup>2\<^esup> = [0, 0::nat]" +apply(auto simp: exponent_def numeral_2_eq_2) +done + +lemma id_case: + "\rec_ci (id m n) = (aprog, rs_pos, a_md); + rec_calc_rel (id m n) lm rs\ + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" +apply(simp add: rec_ci.simps rec_ci_id.simps, auto) +apply(rule_tac calc_id_reverse, simp, simp) +done + +lemma list_tl_induct: + "\P []; \a list. P list \ P (list @ [a::'a])\ \ + P ((list::'a list))" +apply(case_tac "length list", simp) +proof - + fix nat + assume ind: "\a list. P list \ P (list @ [a])" + and h: "length list = Suc nat" "P []" + from h show "P list" + proof(induct nat arbitrary: list, case_tac lista, simp, simp) + fix lista a listaa + from h show "P [a]" + by(insert ind[of "[]"], simp add: h) + next + fix nat list + assume nind: "\list. \length list = Suc nat; P []\ \ P list" + and g: "length (list:: 'a list) = Suc (Suc nat)" + from g show "P (list::'a list)" + apply(insert nind[of "butlast list"], simp add: h) + apply(insert ind[of "butlast list" "last list"], simp) + apply(subgoal_tac "butlast list @ [last list] = list", simp) + apply(case_tac "list::'a list", simp, simp) + done + qed +qed + +thm list.induct + +lemma nth_eq_butlast_nth: "\length ys > Suc k\ \ + ys ! k = butlast ys ! k" +apply(subgoal_tac "\ xs y. ys = xs @ [y]", auto simp: nth_append) +apply(rule_tac x = "butlast ys" in exI, rule_tac x = "last ys" in exI) +apply(case_tac "ys = []", simp, simp) +done + +lemma [simp]: +"\\k + \ \k(ap, pos, n)\map rec_ci list. length ap) + 3 * length list " +apply(induct list arbitrary: pstr, simp, simp) +apply(case_tac "rec_ci a", simp) +done + +lemma [simp]: "Suc n \ pstr \ pstr + x - n > 0" +by arith + +lemma [simp]: + "\Suc (pstr + length list) \ a_md; + length ys = Suc (length list); + length lm = n; + Suc n \ pstr\ + \ (ys ! length list # 0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @ + 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm) ! + (pstr + length list - n) = (0 :: nat)" +apply(insert nth_append[of "ys ! length list # 0\<^bsup>pstr - Suc n\<^esup> @ + butlast ys" "0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm" + "(pstr + length list - n)"], simp add: nth_append) +done + +lemma [simp]: + "\Suc (pstr + length list) \ a_md; + length ys = Suc (length list); + length lm = n; + Suc n \ pstr\ + \ (lm @ last ys # 0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @ + 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm)[pstr + length list := + last ys, n := 0] = + lm @ 0::nat\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm" +apply(insert list_update_length[of + "lm @ last ys # 0\<^bsup>pstr - Suc n\<^esup> @ butlast ys" 0 + "0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm" "last ys"], simp) +apply(simp add: exponent_cons_iff) +apply(insert list_update_length[of "lm" + "last ys" "0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @ + last ys # 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm" 0], simp) +apply(simp add: exponent_cons_iff) +apply(case_tac "ys = []", simp_all add: append_butlast_last_id) +done + + +lemma cn_merge_gs_ex: + "\\x aprog a_md rs_pos rs suf_lm lm. + \x \ set gs; rec_ci x = (aprog, rs_pos, a_md); + rec_calc_rel x lm rs\ + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp + = (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm); + pstr + length gs\ a_md; + \k Max (set (Suc n # map (\(aprog, p, n). n) (map rec_ci gs)))\ + \ \ stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suf_lm) + (cn_merge_gs (map rec_ci gs) pstr) stp + = (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) gs) + + 3 * length gs, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - (pstr + length gs)\<^esup> @ suf_lm)" +apply(induct gs arbitrary: ys rule: list_tl_induct) +apply(simp add: exponent_add_iff, simp) +proof - + fix a list ys + assume ind: "\x aprog a_md rs_pos rs suf_lm lm. + \x = a \ x \ set list; rec_ci x = (aprog, rs_pos, a_md); + rec_calc_rel x lm rs\ + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + and ind2: + "\ys. \\x aprog a_md rs_pos rs suf_lm lm. + \x \ set list; rec_ci x = (aprog, rs_pos, a_md); + rec_calc_rel x lm rs\ + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp + = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm); + \k + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suf_lm) + (cn_merge_gs (map rec_ci list) pstr) stp = + (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) list) + + 3 * length list, + lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm)" + and h: "Suc (pstr + length list) \ a_md" + "\k pstr \ (\(aprog, p, n). n) (rec_ci a) \ pstr \ + (\a\set list. (\(aprog, p, n). n) (rec_ci a) \ pstr)" + from h have k1: + "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suf_lm) + (cn_merge_gs (map rec_ci list) pstr) stp = + (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) list) + + 3 * length list, lm @ 0\<^bsup>pstr - n\<^esup> @ butlast ys @ + 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm) " + apply(rule_tac ind2) + apply(rule_tac ind, auto) + done + from k1 and h show + "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suf_lm) + (cn_merge_gs (map rec_ci list @ [rec_ci a]) pstr) stp = + (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) list) + + (\(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list), + lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm)" + apply(simp add: cn_merge_gs_tl_app) + thm abc_append_exc2 + apply(rule_tac as = + "(\(ap, pos, n)\map rec_ci list. length ap) + 3 * length list" + and bm = "lm @ 0\<^bsup>pstr - n\<^esup> @ butlast ys @ + 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm" + and bs = "(\(ap, pos, n). length ap) (rec_ci a) + 3" + and bm' = "lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ + suf_lm" in abc_append_exc2, simp) + apply(simp add: cn_merge_gs_length) + proof - + from h show + "\bstp. abc_steps_l (0, lm @ 0\<^bsup>pstr - n\<^esup> @ butlast ys @ + 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm) + ((\(gprog, gpara, gn). gprog [+] recursive.empty gpara + (pstr + length list)) (rec_ci a)) bstp = + ((\(ap, pos, n). length ap) (rec_ci a) + 3, + lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm)" + apply(case_tac "rec_ci a", simp) + apply(rule_tac as = "length aa" and + bm = "lm @ (ys ! (length list)) # + 0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @ 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm" + and bs = "3" and bm' = "lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ + 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm" in abc_append_exc2) + proof - + fix aa b c + assume g: "rec_ci a = (aa, b, c)" + from h and g have k2: "b = n" + apply(erule_tac x = "length list" in allE, simp) + apply(subgoal_tac "length lm = b", simp) + apply(rule para_pattern, simp, simp) + done + from h and g and this show + "\astp. abc_steps_l (0, lm @ 0\<^bsup>pstr - n\<^esup> @ butlast ys @ + 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm) aa astp = + (length aa, lm @ ys ! length list # 0\<^bsup>pstr - Suc n\<^esup> @ + butlast ys @ 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm)" + apply(subgoal_tac "c \ Suc n") + apply(insert ind[of a aa b c lm "ys ! length list" + "0\<^bsup>pstr - c\<^esup> @ butlast ys @ 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm"], simp) + apply(erule_tac x = "length list" in allE, + simp add: exponent_add_iff) + apply(rule_tac Suc_leI, rule_tac ci_ad_ge_paras, simp) + done + next + fix aa b c + show "length aa = length aa" by simp + next + fix aa b c + assume "rec_ci a = (aa, b, c)" + from h and this show + "\bstp. abc_steps_l (0, lm @ ys ! length list # + 0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @ 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm) + (recursive.empty b (pstr + length list)) bstp = + (3, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm)" + apply(insert empty_ex [of b "pstr + length list" + "lm @ ys ! length list # 0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @ + 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm"], simp) + apply(subgoal_tac "b = n") + apply(simp add: nth_append split: if_splits) + apply(erule_tac x = "length list" in allE, simp) + apply(drule para_pattern, simp, simp) + done + next + fix aa b c + show "3 = length (recursive.empty b (pstr + length list))" + by simp + next + fix aa b aaa ba + show "length aa + 3 = length aa + 3" by simp + next + fix aa b c + show "empty b (pstr + length list) \ []" + by(simp add: empty.simps) + qed + next + show "(\(ap, pos, n). length ap) (rec_ci a) + 3 = + length ((\(gprog, gpara, gn). gprog [+] + recursive.empty gpara (pstr + length list)) (rec_ci a))" + by(case_tac "rec_ci a", simp) + next + show "listsum (map ((\(ap, pos, n). length ap) \ rec_ci) list) + + (\(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list)= + (\(ap, pos, n)\map rec_ci list. length ap) + 3 * length list + + ((\(ap, pos, n). length ap) (rec_ci a) + 3)" by simp + next + show "(\(gprog, gpara, gn). gprog [+] + recursive.empty gpara (pstr + length list)) (rec_ci a) \ []" + by(case_tac "rec_ci a", + simp add: abc_append.simps abc_shift.simps) + qed +qed + +declare drop_abc_lm_v_simp[simp del] + +lemma [simp]: "length (mv_boxes aa ba n) = 3*n" +by(induct n, auto simp: mv_boxes.simps) + +lemma exp_suc: "a\<^bsup>Suc b\<^esup> = a\<^bsup>b\<^esup> @ [a]" +by(simp add: exponent_def rep_ind del: replicate.simps) + +lemma [simp]: + "\Suc n \ ba - aa; length lm2 = Suc n; + length lm3 = ba - Suc (aa + n)\ + \ (last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = (0::nat)" +proof - + assume h: "Suc n \ ba - aa" + and g: "length lm2 = Suc n" "length lm3 = ba - Suc (aa + n)" + from h and g have k: "ba - aa = Suc (length lm3 + n)" + by arith + from k show + "(last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = 0" + apply(simp, insert g) + apply(simp add: nth_append) + done +qed + +lemma [simp]: "length lm1 = aa \ + (lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (aa + n) = last lm2" +apply(simp add: nth_append) +done + +lemma [simp]: "\Suc n \ ba - aa; aa < ba\ \ + (ba < Suc (aa + (ba - Suc (aa + n) + n))) = False" +apply arith +done + +lemma [simp]: "\Suc n \ ba - aa; aa < ba; length lm1 = aa; + length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\ + \ (lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba + n) = 0" +using nth_append[of "lm1 @ 0\'a\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2" + "(0\'a) # lm4" "ba + n"] +apply(simp) +done + +lemma [simp]: + "\Suc n \ ba - aa; aa < ba; length lm1 = aa; length lm2 = Suc n; + length lm3 = ba - Suc (aa + n)\ + \ (lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2 @ (0::nat) # lm4) + [ba + n := last lm2, aa + n := 0] = + lm1 @ 0 # 0\<^bsup>n\<^esup> @ lm3 @ lm2 @ lm4" +using list_update_append[of "lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2" "0 # lm4" + "ba + n" "last lm2"] +apply(simp) +apply(simp add: list_update_append) +apply(simp only: exponent_cons_iff exp_suc, simp) +apply(case_tac lm2, simp, simp) +done + + +lemma mv_boxes_ex: + "\n \ ba - aa; ba > aa; length lm1 = aa; + length (lm2::nat list) = n; length lm3 = ba - aa - n\ + \ \ stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<^bsup>n\<^esup> @ lm4) + (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<^bsup>n\<^esup> @ lm3 @ lm2 @ lm4)" +apply(induct n arbitrary: lm2 lm3 lm4, simp) +apply(rule_tac x = 0 in exI, simp add: abc_steps_zero, + simp add: mv_boxes.simps del: exp_suc_iff) +apply(rule_tac as = "3 *n" and bm = "lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ + butlast lm2 @ 0 # lm4" in abc_append_exc2, simp_all) +apply(simp only: exponent_cons_iff, simp only: exp_suc, simp) +proof - + fix n lm2 lm3 lm4 + assume ind: + "\lm2 lm3 lm4. \length lm2 = n; length lm3 = ba - (aa + n)\ \ + \stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<^bsup>n\<^esup> @ lm4) + (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<^bsup>n\<^esup> @ lm3 @ lm2 @ lm4)" + and h: "Suc n \ ba - aa" "aa < ba" "length (lm1::nat list) = aa" + "length (lm2::nat list) = Suc n" + "length (lm3::nat list) = ba - Suc (aa + n)" + from h show + "\astp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<^bsup>n\<^esup> @ 0 # lm4) + (mv_boxes aa ba n) astp = + (3 * n, lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4)" + apply(insert ind[of "butlast lm2" "last lm2 # lm3" "0 # lm4"], + simp) + apply(subgoal_tac "lm1 @ butlast lm2 @ last lm2 # lm3 @ 0\<^bsup>n\<^esup> @ + 0 # lm4 = lm1 @ lm2 @ lm3 @ 0\<^bsup>n\<^esup> @ 0 # lm4", simp, simp) + apply(case_tac "lm2 = []", simp, simp) + done +next + fix n lm2 lm3 lm4 + assume h: "Suc n \ ba - aa" + "aa < ba" + "length (lm1::nat list) = aa" + "length (lm2::nat list) = Suc n" + "length (lm3::nat list) = ba - Suc (aa + n)" + thus " \bstp. abc_steps_l (0, lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ + butlast lm2 @ 0 # lm4) + (recursive.empty (aa + n) (ba + n)) bstp + = (3, lm1 @ 0 # 0\<^bsup>n\<^esup> @ lm3 @ lm2 @ lm4)" + apply(insert empty_ex[of "aa + n" "ba + n" + "lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4"], simp) + done +qed +(* +lemma [simp]: "\Suc n \ aa - ba; + ba < aa; + length lm2 = aa - Suc (ba + n)\ + \ ((0::nat) # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4) ! (aa - ba) + = last lm3" +proof - + assume h: "Suc n \ aa - ba" + and g: " ba < aa" "length lm2 = aa - Suc (ba + n)" + from h and g have k: "aa - ba = Suc (length lm2 + n)" + by arith + thus "((0::nat) # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4) ! (aa - ba) = last lm3" + apply(simp, simp add: nth_append) + done +qed +*) + +lemma [simp]: "\Suc n \ aa - ba; ba < aa; length lm1 = ba; + length lm2 = aa - Suc (ba + n); length lm3 = Suc n\ + \ (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4) ! (aa + n) = last lm3" +using nth_append[of "lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup>" "last lm3 # lm4" "aa + n"] +apply(simp) +done + +lemma [simp]: "\Suc n \ aa - ba; ba < aa; length lm1 = ba; + length lm2 = aa - Suc (ba + n); length lm3 = Suc n\ + \ (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4) ! (ba + n) = 0" +apply(simp add: nth_append) +done + + +lemma [simp]: "\Suc n \ aa - ba; ba < aa; length lm1 = ba; + length lm2 = aa - Suc (ba + n); length lm3 = Suc n\ + \ (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4)[ba + n := last lm3, aa + n := 0] + = lm1 @ lm3 @ lm2 @ 0 # 0\<^bsup>n\<^esup> @ lm4" +using list_update_append[of "lm1 @ butlast lm3" "(0\'a) # lm2 @ 0\'a\<^bsup>n\<^esup> @ last lm3 # lm4"] +apply(simp) +using list_update_append[of "lm1 @ butlast lm3 @ last lm3 # lm2 @ 0\'a\<^bsup>n\<^esup>" + "last lm3 # lm4" "aa + n" "0"] +apply(simp) +apply(simp only: exp_ind_def[THEN sym] exp_suc, simp) +apply(case_tac lm3, simp, simp) +done + + +lemma mv_boxes_ex2: + "\n \ aa - ba; + ba < aa; + length (lm1::nat list) = ba; + length (lm2::nat list) = aa - ba - n; + length (lm3::nat list) = n\ + \ \ stp. abc_steps_l (0, lm1 @ 0\<^bsup>n\<^esup> @ lm2 @ lm3 @ lm4) + (mv_boxes aa ba n) stp = + (3 * n, lm1 @ lm3 @ lm2 @ 0\<^bsup>n\<^esup> @ lm4)" +apply(induct n arbitrary: lm2 lm3 lm4, simp) +apply(rule_tac x = 0 in exI, simp add: abc_steps_zero, + simp add: mv_boxes.simps del: exp_suc_iff) +apply(rule_tac as = "3 *n" and bm = "lm1 @ butlast lm3 @ 0 # lm2 @ + 0\<^bsup>n\<^esup> @ last lm3 # lm4" in abc_append_exc2, simp_all) +apply(simp only: exponent_cons_iff, simp only: exp_suc, simp) +proof - + fix n lm2 lm3 lm4 + assume ind: +"\lm2 lm3 lm4. \length lm2 = aa - (ba + n); length lm3 = n\ \ + \stp. abc_steps_l (0, lm1 @ 0\<^bsup>n\<^esup> @ lm2 @ lm3 @ lm4) + (mv_boxes aa ba n) stp = + (3 * n, lm1 @ lm3 @ lm2 @ 0\<^bsup>n\<^esup> @ lm4)" + and h: "Suc n \ aa - ba" + "ba < aa" + "length (lm1::nat list) = ba" + "length (lm2::nat list) = aa - Suc (ba + n)" + "length (lm3::nat list) = Suc n" + from h show + "\astp. abc_steps_l (0, lm1 @ 0\<^bsup>n\<^esup> @ 0 # lm2 @ lm3 @ lm4) + (mv_boxes aa ba n) astp = + (3 * n, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4)" + apply(insert ind[of "0 # lm2" "butlast lm3" "last lm3 # lm4"], + simp) + apply(subgoal_tac + "lm1 @ 0\<^bsup>n\<^esup> @ 0 # lm2 @ butlast lm3 @ last lm3 # lm4 = + lm1 @ 0\<^bsup>n\<^esup> @ 0 # lm2 @ lm3 @ lm4", simp, simp) + apply(case_tac "lm3 = []", simp, simp) + done +next + fix n lm2 lm3 lm4 + assume h: + "Suc n \ aa - ba" + "ba < aa" + "length lm1 = ba" + "length (lm2::nat list) = aa - Suc (ba + n)" + "length (lm3::nat list) = Suc n" + thus + "\bstp. abc_steps_l (0, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup> @ + last lm3 # lm4) + (recursive.empty (aa + n) (ba + n)) bstp = + (3, lm1 @ lm3 @ lm2 @ 0 # 0\<^bsup>n\<^esup> @ lm4)" + apply(insert empty_ex[of "aa + n" "ba + n" "lm1 @ butlast lm3 @ + 0 # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4"], simp) + done +qed + +lemma cn_merge_gs_len: + "length (cn_merge_gs (map rec_ci gs) pstr) = + (\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs" +apply(induct gs arbitrary: pstr, simp, simp) +apply(case_tac "rec_ci a", simp) +done + +lemma [simp]: "n < pstr \ + Suc (pstr + length ys - n) = Suc (pstr + length ys) - n" +by arith + +lemma save_paras': + "\length lm = n; pstr > n; a_md > pstr + length ys + n\ + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ + 0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm) + (mv_boxes 0 (pstr + Suc (length ys)) n) stp + = (3 * n, 0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" +thm mv_boxes_ex +apply(insert mv_boxes_ex[of n "pstr + Suc (length ys)" 0 "[]" "lm" + "0\<^bsup>pstr - n\<^esup> @ ys @ [0]" "0\<^bsup>a_md - pstr - length ys - n - Suc 0\<^esup> @ suf_lm"], simp) +apply(erule_tac exE, rule_tac x = stp in exI, + simp add: exponent_add_iff) +apply(simp only: exponent_cons_iff, simp) +done + +lemma [simp]: + "(max ba (Max (insert ba (((\(aprog, p, n). n) o rec_ci) ` set gs)))) + = (Max (insert ba (((\(aprog, p, n). n) o rec_ci) ` set gs)))" +apply(rule min_max.sup_absorb2, auto) +done + +lemma [simp]: + "((\(aprog, p, n). n) ` rec_ci ` set gs) = + (((\(aprog, p, n). n) o rec_ci) ` set gs)" +apply(induct gs) +apply(simp, simp) +done + +lemma ci_cn_md_def: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_ci f = (a, aa, ba)\ + \ a_md = max (Suc n) (Max (insert ba (((\(aprog, p, n). n) o + rec_ci) ` set gs))) + Suc (length gs) + n" +apply(simp add: rec_ci.simps, auto) +done + +lemma save_paras_prog_ex: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_ci f = (a, aa, ba); + pstr = Max (set (Suc n # ba # map (\(aprog, p, n). n) + (map rec_ci (f # gs))))\ + \ \ap bp cp. + aprog = ap [+] bp [+] cp \ + length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + + 3 * length gs \ bp = mv_boxes 0 (pstr + Suc (length gs)) n" +apply(simp add: rec_ci.simps) +apply(rule_tac x = + "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) o rec_ci) ` set gs))))" in exI, + simp add: cn_merge_gs_len) +apply(rule_tac x = + "mv_boxes (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) + 0 (length gs) [+] a [+]recursive.empty aa (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] + empty_boxes (length gs) [+] recursive.empty (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] + mv_boxes (Suc (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) + ` set gs))) + length gs)) 0 n" in exI, auto) +apply(simp add: abc_append_commute) +done + +lemma save_paras: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rs_pos = n; + \k(aprog, p, n). n) + (map rec_ci (f # gs))))\ + \ \stp. abc_steps_l ((\(ap, pos, n)\map rec_ci gs. length ap) + + 3 * length gs, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ + 0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm) aprog stp = + ((\(ap, pos, n)\map rec_ci gs. length ap) + + 3 * length gs + 3 * n, + 0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" +proof - + assume h: + "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" + "rs_pos = n" + "\k(aprog, p, n). n) + (map rec_ci (f # gs))))" + from h and g have k1: + "\ ap bp cp. aprog = ap [+] bp [+] cp \ + length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + + 3 *length gs \ bp = mv_boxes 0 (pstr + Suc (length ys)) n" + apply(drule_tac save_paras_prog_ex, auto) + done + from h have k2: + "\ stp. abc_steps_l (0, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ + 0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm) + (mv_boxes 0 (pstr + Suc (length ys)) n) stp = + (3 * n, 0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" + apply(rule_tac save_paras', simp, simp_all add: g) + apply(drule_tac a = a and aa = aa and ba = ba in + ci_cn_md_def, simp, simp) + done + from k1 show + "\stp. abc_steps_l ((\(ap, pos, n)\map rec_ci gs. length ap) + + 3 * length gs, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ + 0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm) aprog stp = + ((\(ap, pos, n)\map rec_ci gs. length ap) + + 3 * length gs + 3 * n, + 0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" + proof(erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE) + fix ap bp apa cp + assume "aprog = ap [+] bp [+] cp \ length ap = + (\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs + \ bp = mv_boxes 0 (pstr + Suc (length ys)) n" + from this and k2 show "?thesis" + apply(simp) + apply(rule_tac abc_append_exc1, simp, simp, simp) + done + qed +qed + +lemma ci_cn_para_eq: + "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md) \ rs_pos = n" +apply(simp add: rec_ci.simps, case_tac "rec_ci f", simp) +done + +lemma calc_gs_prog_ex: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_ci f = (a, aa, ba); + Max (set (Suc n # ba # map (\(aprog, p, n). n) + (map rec_ci (f # gs)))) = pstr\ + \ \ap bp. aprog = ap [+] bp \ + ap = cn_merge_gs (map rec_ci gs) pstr" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "mv_boxes 0 (Suc (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) n [+] + mv_boxes (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs) [+] + a [+] recursive.empty aa (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] + empty_boxes (length gs) [+] recursive.empty (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] + mv_boxes (Suc (max (Suc n) (Max + (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) 0 n" + in exI) +apply(auto simp: abc_append_commute) +done + +lemma cn_calc_gs: + assumes ind: + "\x aprog a_md rs_pos rs suf_lm lm. + \x \ set gs; + rec_ci x = (aprog, rs_pos, a_md); + rec_calc_rel x lm rs\ + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" + "\k(aprog, p, n). n) + (map rec_ci (f # gs)))) = pstr" + shows + "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + ((\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs, + lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md -pstr - length ys\<^esup> @ suf_lm) " +proof - + from h have k1: + "\ ap bp. aprog = ap [+] bp \ ap = + cn_merge_gs (map rec_ci gs) pstr" + by(erule_tac calc_gs_prog_ex, auto) + from h have j1: "rs_pos = n" + by(simp add: ci_cn_para_eq) + from h have j2: "a_md \ pstr" + by(drule_tac a = a and aa = aa and ba = ba in + ci_cn_md_def, simp, simp) + from h have j3: "pstr > n" + by(auto) + from j1 and j2 and j3 and h have k2: + "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) + (cn_merge_gs (map rec_ci gs) pstr) stp + = ((\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs, + lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm)" + apply(simp) + apply(rule_tac cn_merge_gs_ex, rule_tac ind, simp, simp, auto) + apply(drule_tac a = a and aa = aa and ba = ba in + ci_cn_md_def, simp, simp) + apply(rule min_max.le_supI2, auto) + done + from k1 show "?thesis" + proof(erule_tac exE, erule_tac exE, simp) + fix ap bp + from k2 show + "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) + (cn_merge_gs (map rec_ci gs) pstr [+] bp) stp = + (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) gs) + + 3 * length gs, + lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - (pstr + length ys)\<^esup> @ suf_lm)" + apply(insert abc_append_exc1[of + "lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm" + "(cn_merge_gs (map rec_ci gs) pstr)" + "length (cn_merge_gs (map rec_ci gs) pstr)" + "lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm" 0 + "[]" bp], simp add: cn_merge_gs_len) + done + qed +qed + +lemma reset_new_paras': + "\length lm = n; + pstr > 0; + a_md \ pstr + length ys + n; + pstr > length ys\ \ + \stp. abc_steps_l (0, 0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ + suf_lm) (mv_boxes pstr 0 (length ys)) stp = + (3 * length ys, ys @ 0\<^bsup>pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" +thm mv_boxes_ex2 +apply(insert mv_boxes_ex2[of "length ys" "pstr" 0 "[]" + "0\<^bsup>pstr - length ys\<^esup>" "ys" + "0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm"], + simp add: exponent_add_iff) +done + +lemma [simp]: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_calc_rel f ys rs; rec_ci f = (a, aa, ba); + pstr = Max (set (Suc n # ba # map (\(aprog, p, n). n) + (map rec_ci (f # gs))))\ + \ length ys < pstr" +apply(subgoal_tac "length ys = aa", simp) +apply(subgoal_tac "aa < ba \ ba \ pstr", + rule basic_trans_rules(22), auto) +apply(rule min_max.le_supI2) +apply(auto) +apply(erule_tac para_pattern, simp) +done + +lemma reset_new_paras_prog_ex: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_ci f = (a, aa, ba); + Max (set (Suc n # ba # map (\(aprog, p, n). n) + (map rec_ci (f # gs)))) = pstr\ + \ \ ap bp cp. aprog = ap [+] bp [+] cp \ + length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + + 3 *length gs + 3 * n \ bp = mv_boxes pstr 0 (length gs)" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] + mv_boxes 0 (Suc (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) n" in exI, + simp add: cn_merge_gs_len) +apply(rule_tac x = "a [+] + recursive.empty aa (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] + empty_boxes (length gs) [+] recursive.empty + (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n + [+] mv_boxes (Suc (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) 0 n" in exI, + auto simp: abc_append_commute) +done + + +lemma reset_new_paras: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rs_pos = n; + \k(aprog, p, n). n) + (map rec_ci (f # gs))))\ +\ \stp. abc_steps_l ((\(ap, pos, n)\map rec_ci gs. length ap) + + 3 * length gs + 3 * n, + 0\<^bsup>pstr \<^esup>@ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) aprog stp = + ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + 3 * n, + ys @ 0\<^bsup>pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" +proof - + assume h: + "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" + "rs_pos = n" + "length ys = aa" + "\k(aprog, p, n). n) + (map rec_ci (f # gs))))" + thm rec_ci.simps + from h and g have k1: + "\ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = + (\(ap, pos, n)\map rec_ci gs. length ap) + + 3 *length gs + 3 * n \ bp = mv_boxes pstr 0 (length ys)" + by(drule_tac reset_new_paras_prog_ex, auto) + from h have k2: + "\ stp. abc_steps_l (0, 0\<^bsup>pstr \<^esup>@ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ + suf_lm) (mv_boxes pstr 0 (length ys)) stp = + (3 * (length ys), + ys @ 0\<^bsup>pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" + apply(rule_tac reset_new_paras', simp) + apply(simp add: g) + apply(drule_tac a = a and aa = aa and ba = ba in ci_cn_md_def, + simp, simp add: g, simp) + apply(subgoal_tac "length gs = aa \ aa < ba \ ba \ pstr", arith, + simp add: para_pattern) + apply(insert g, auto intro: min_max.le_supI2) + done + from k1 show + "\stp. abc_steps_l ((\(ap, pos, n)\map rec_ci gs. length ap) + 3 + * length gs + 3 * n, 0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ + suf_lm) aprog stp = + ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + + 3 * n, ys @ 0\<^bsup>pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" + proof(erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE) + fix ap bp apa cp + assume "aprog = ap [+] bp [+] cp \ length ap = + (\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs + + 3 * n \ bp = mv_boxes pstr 0 (length ys)" + from this and k2 show "?thesis" + apply(simp) + apply(drule_tac as = + "(\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs + + 3 * n" and ap = ap and cp = cp in abc_append_exc1, auto) + apply(rule_tac x = stp in exI, simp add: h) + using h + apply(simp) + done + qed +qed + +thm rec_ci.simps + +lemma calc_f_prog_ex: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_ci f = (a, aa, ba); + Max (set (Suc n # ba # map (\(aprog, p, n). n) + (map rec_ci (f # gs)))) = pstr\ + \ \ ap bp cp. aprog = ap [+] bp [+] cp \ + length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + + 6 *length gs + 3 * n \ bp = a" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] + mv_boxes 0 (Suc (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) n [+] + mv_boxes (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs)" in exI, + simp add: cn_merge_gs_len) +apply(rule_tac x = "recursive.empty aa (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] + empty_boxes (length gs) [+] recursive.empty (max (Suc n) ( + Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] + mv_boxes (Suc (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) 0 n" in exI, + auto simp: abc_append_commute) +done + +lemma calc_cn_f: + assumes ind: + "\x aprog a_md rs_pos rs suf_lm lm. + \x \ set (f # gs); + rec_ci x = (aprog, rs_pos, a_md); + rec_calc_rel x lm rs\ + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Cn n f gs) lm rs" + "length ys = length gs" + "rec_calc_rel f ys rs" + "length lm = n" + "rec_ci f = (a, aa, ba)" + and p: "pstr = Max (set (Suc n # ba # map (\(aprog, p, n). n) + (map rec_ci (f # gs))))" + shows "\stp. abc_steps_l + ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + 3 * n, + ys @ 0\<^bsup>pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) aprog stp = + ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + + 3 * n + length a, + ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" +proof - + from h have k1: + "\ ap bp cp. aprog = ap [+] bp [+] cp \ + length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + + 6 *length gs + 3 * n \ bp = a" + by(drule_tac calc_f_prog_ex, auto) + from h and k1 show "?thesis" + proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE) + fix ap bp apa cp + assume + "aprog = ap [+] bp [+] cp \ + length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + + 6 * length gs + 3 * n \ bp = a" + from h and this show "?thesis" + apply(simp, rule_tac abc_append_exc1, simp_all) + apply(insert ind[of f "a" aa ba ys rs + "0\<^bsup>pstr - ba + length gs \<^esup> @ 0 # lm @ + 0\<^bsup>a_md - Suc (pstr + length gs + n)\<^esup> @ suf_lm"], simp) + apply(subgoal_tac "ba > aa \ aa = length gs\ pstr \ ba", simp) + apply(simp add: exponent_add_iff) + apply(case_tac pstr, simp add: p) + apply(simp only: exp_suc, simp) + apply(rule conjI, rule ci_ad_ge_paras, simp, rule conjI) + apply(subgoal_tac "length ys = aa", simp, + rule para_pattern, simp, simp) + apply(insert p, simp) + apply(auto intro: min_max.le_supI2) + done + qed +qed +(* +lemma [simp]: + "\pstr + length ys + n \ a_md; ys \ []\ \ + pstr < a_md + length suf_lm" +apply(case_tac "length ys", simp) +apply(arith) +done +*) +lemma [simp]: + "pstr > length ys + \ (ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @ + 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) ! pstr = (0::nat)" +apply(simp add: nth_append) +done + +(* +lemma [simp]: "\length ys < pstr; pstr - length ys = Suc x\ + \ pstr - Suc (length ys) = x" +by arith +*) +lemma [simp]: "pstr > length ys \ + (ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) + [pstr := rs, length ys := 0] = + ys @ 0\<^bsup>pstr - length ys\<^esup> @ (rs::nat) # 0\<^bsup>length ys\<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm" +apply(auto simp: list_update_append) +apply(case_tac "pstr - length ys",simp_all) +using list_update_length[of + "0\<^bsup>pstr - Suc (length ys)\<^esup>" "0" "0\<^bsup>length ys\<^esup> @ lm @ + 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm" rs] +apply(simp only: exponent_cons_iff exponent_add_iff, simp) +apply(subgoal_tac "pstr - Suc (length ys) = nat", simp, simp) +done + +lemma save_rs': + "\pstr > length ys\ + \ \stp. abc_steps_l (0, ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @ + 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) + (recursive.empty (length ys) pstr) stp = + (3, ys @ 0\<^bsup>pstr - (length ys)\<^esup> @ rs # + 0\<^bsup>length ys \<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" +apply(insert empty_ex[of "length ys" pstr + "ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @ 0\<^bsup>a_md - Suc(pstr + length ys + n)\<^esup> @ suf_lm"], + simp) +done + + +lemma save_rs_prog_ex: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_ci f = (a, aa, ba); + Max (set (Suc n # ba # map (\(aprog, p, n). n) + (map rec_ci (f # gs)))) = pstr\ + \ \ ap bp cp. aprog = ap [+] bp [+] cp \ + length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + + 6 *length gs + 3 * n + length a + \ bp = empty aa pstr" +apply(simp add: rec_ci.simps) +apply(rule_tac x = + "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) + [+] mv_boxes 0 (Suc (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) n [+] + mv_boxes (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) + 0 (length gs) [+] a" + in exI, simp add: cn_merge_gs_len) +apply(rule_tac x = + "empty_boxes (length gs) [+] + recursive.empty (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] + mv_boxes (Suc (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + + length gs)) 0 n" in exI, + auto simp: abc_append_commute) +done + +lemma save_rs: + assumes h: + "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Cn n f gs) lm rs" + "\k(aprog, p, n). n) + (map rec_ci (f # gs))))" + shows "\stp. abc_steps_l + ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + + 3 * n + length a, ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @ + 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) aprog stp = + ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + + 3 * n + length a + 3, + ys @ 0\<^bsup>pstr - length ys \<^esup> @ rs # 0\<^bsup>length ys\<^esup> @ lm @ + 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" +proof - + thm rec_ci.simps + from h and pdef have k1: + "\ ap bp cp. aprog = ap [+] bp [+] cp \ + length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + + 6 *length gs + 3 * n + length a \ bp = empty (length ys) pstr " + apply(subgoal_tac "length ys = aa") + apply(drule_tac a = a and aa = aa and ba = ba in save_rs_prog_ex, + simp, simp, simp) + by(rule_tac para_pattern, simp, simp) + from k1 show + "\stp. abc_steps_l + ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + 3 * n + + length a, ys @ rs # 0\<^bsup>pstr \<^esup>@ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> + @ suf_lm) aprog stp = + ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + 3 * n + + length a + 3, ys @ 0\<^bsup>pstr - length ys\<^esup> @ rs # + 0\<^bsup>length ys\<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" + proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE) + fix ap bp apa cp + assume "aprog = ap [+] bp [+] cp \ length ap = + (\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + + 3 * n + length a \ bp = recursive.empty (length ys) pstr" + thus"?thesis" + apply(simp, rule_tac abc_append_exc1, simp_all) + apply(rule_tac save_rs', insert h) + apply(subgoal_tac "length gs = aa \ pstr \ ba \ ba > aa", + arith) + apply(simp add: para_pattern, insert pdef, auto) + apply(rule_tac min_max.le_supI2, simp) + done + qed +qed + +lemma [simp]: "length (empty_boxes n) = 2*n" +apply(induct n, simp, simp) +done + +lemma empty_step_ex: "length lm = n \ + \stp. abc_steps_l (0, lm @ Suc x # suf_lm) [Dec n 2, Goto 0] stp + = (0, lm @ x # suf_lm)" +apply(rule_tac x = "Suc (Suc 0)" in exI, + simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps + abc_lm_v.simps abc_lm_s.simps nth_append list_update_append) +done + +lemma empty_box_ex: + "\length lm = n\ \ + \ stp. abc_steps_l (0, lm @ x # suf_lm) [Dec n 2, Goto 0] stp = + (Suc (Suc 0), lm @ 0 # suf_lm)" +apply(induct x) +apply(rule_tac x = "Suc 0" in exI, + simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps + abc_lm_v.simps nth_append abc_lm_s.simps, simp) +apply(drule_tac x = x and suf_lm = suf_lm in empty_step_ex, + erule_tac exE, erule_tac exE) +apply(rule_tac x = "stpa + stp" in exI, simp add: abc_steps_add) +done + +lemma [simp]: "drop n lm = a # list \ list = drop (Suc n) lm" +apply(induct n arbitrary: lm a list, simp) +apply(case_tac "lm", simp, simp) +done + +lemma empty_boxes_ex: "\length lm \ n\ + \ \stp. abc_steps_l (0, lm) (empty_boxes n) stp = + (2*n, 0\<^bsup>n\<^esup> @ drop n lm)" +apply(induct n, simp, simp) +apply(rule_tac abc_append_exc2, auto) +apply(case_tac "drop n lm", simp, simp) +proof - + fix n stp a list + assume h: "Suc n \ length lm" "drop n lm = a # list" + thus "\bstp. abc_steps_l (0, 0\<^bsup>n\<^esup> @ a # list) [Dec n 2, Goto 0] bstp = + (Suc (Suc 0), 0 # 0\<^bsup>n\<^esup> @ drop (Suc n) lm)" + apply(insert empty_box_ex[of "0\<^bsup>n\<^esup>" n a list], simp, erule_tac exE) + apply(rule_tac x = stp in exI, simp, simp only: exponent_cons_iff) + apply(simp add: exponent_def rep_ind del: replicate.simps) + done +qed + + +lemma empty_paras_prog_ex: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_ci f = (a, aa, ba); + Max (set (Suc n # ba # map (\(aprog, p, n). n) + (map rec_ci (f # gs)))) = pstr\ + \ \ ap bp cp. aprog = ap [+] bp [+] cp \ + length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + + 6 *length gs + 3 * n + length a + 3 \ bp = empty_boxes (length gs)" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] + mv_boxes 0 (Suc (max (Suc n) (Max + (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) n + [+] mv_boxes (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs) [+] + a [+] recursive.empty aa (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))))" + in exI, simp add: cn_merge_gs_len) +apply(rule_tac x = " recursive.empty (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] + mv_boxes (Suc (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) 0 n" in exI, + auto simp: abc_append_commute) +done + +lemma empty_paras: + assumes h: + "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Cn n f gs) lm rs" + "\k(aprog, p, n). n) + (map rec_ci (f # gs))))" + and starts: "ss = (\(ap, pos, n)\map rec_ci gs. length ap) + + 6 * length gs + 3 * n + length a + 3" + shows "\stp. abc_steps_l + (ss, ys @ 0\<^bsup>pstr - length ys\<^esup> @ rs # 0\<^bsup>length ys\<^esup> + @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) aprog stp = + (ss + 2 * length gs, 0\<^bsup>pstr\<^esup> @ rs # 0\<^bsup>length ys \<^esup> @ lm @ + 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" +proof - + from h and pdef and starts have k1: + "\ ap bp cp. aprog = ap [+] bp [+] cp \ + length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + + 6 *length gs + 3 * n + length a + 3 + \ bp = empty_boxes (length ys)" + by(drule_tac empty_paras_prog_ex, auto) + from h have j1: "aa < ba" + by(simp add: ci_ad_ge_paras) + from h have j2: "length gs = aa" + by(drule_tac f = f in para_pattern, simp, simp) + from h and pdef have j3: "ba \ pstr" + apply simp + apply(rule_tac min_max.le_supI2, simp) + done + from k1 show "?thesis" + proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE) + fix ap bp apa cp + assume "aprog = ap [+] bp [+] cp \ + length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + + 6 * length gs + 3 * n + length a + 3 \ + bp = empty_boxes (length ys)" + thus"?thesis" + apply(simp, rule_tac abc_append_exc1, simp_all add: starts h) + apply(insert empty_boxes_ex[of + "length gs" "ys @ 0\<^bsup>pstr - (length gs)\<^esup> @ rs # + 0\<^bsup>length gs\<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length gs + n)\<^esup> @ suf_lm"], + simp add: h) + apply(erule_tac exE, rule_tac x = stp in exI, + simp add: exponent_def replicate.simps[THEN sym] + replicate_add[THEN sym] del: replicate.simps) + apply(subgoal_tac "pstr >(length gs)", simp) + apply(subgoal_tac "ba > aa \ length gs = aa \ pstr \ ba", simp) + apply(simp add: j1 j2 j3) + done + qed +qed + +(* +lemma [simp]: " n < pstr \ + (0\<^bsup>pstr\<^esup>)[n := rs] @ [0::nat] = 0\<^bsup>n\<^esup> @ rs # 0\<^bsup>pstr - n\<^esup>" +apply(insert list_update_length[of "0\<^bsup>n\<^esup>" 0 "0\<^bsup>pstr - Suc n\<^esup>" rs]) +apply(insert exponent_cons_iff[of "0::nat" "pstr - Suc n" "[]"], simp) +apply(insert exponent_add_iff[of "0::nat" n "pstr - n" "[]"], simp) +apply(case_tac "pstr - n", simp, simp only: exp_suc, simp) +apply(subgoal_tac "pstr - Suc n = nat", simp) +by arith +*) + +lemma restore_rs_prog_ex: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_ci f = (a, aa, ba); + Max (set (Suc n # ba # map (\(aprog, p, n). n) + (map rec_ci (f # gs)))) = pstr; + ss = (\(ap, pos, n)\map rec_ci gs. length ap) + + 8 * length gs + 3 * n + length a + 3\ + \ \ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = ss \ + bp = empty pstr n" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] + mv_boxes 0 (Suc (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) + \ rec_ci) ` set gs))) + length gs)) n [+] + mv_boxes (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs) [+] + a [+] recursive.empty aa (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] + empty_boxes (length gs)" in exI, simp add: cn_merge_gs_len) +apply(rule_tac x = "mv_boxes (Suc (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + + length gs)) 0 n" + in exI, auto simp: abc_append_commute) +done + +lemma exp_add: "a\<^bsup>b+c\<^esup> = a\<^bsup>b\<^esup> @ a\<^bsup>c\<^esup>" +apply(simp add: exponent_def replicate_add) +done + +lemma [simp]: "n < pstr \ (0\<^bsup>pstr\<^esup>)[n := rs] @ [0::nat] = 0\<^bsup>n\<^esup> @ rs # 0\<^bsup>pstr - n\<^esup>" +using list_update_length[of "0\<^bsup>n\<^esup>" "0::nat" "0\<^bsup>pstr - Suc n\<^esup>" rs] +apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym] exp_suc[THEN sym]) +done + +lemma restore_rs: + assumes h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Cn n f gs) lm rs" + "\k(aprog, p, n). n) + (map rec_ci (f # gs))))" + and starts: "ss = (\(ap, pos, n)\map rec_ci gs. length ap) + + 8 * length gs + 3 * n + length a + 3" + shows "\stp. abc_steps_l + (ss, 0\<^bsup>pstr\<^esup> @ rs # 0\<^bsup>length ys \<^esup> @ lm @ + 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) aprog stp = + (ss + 3, 0\<^bsup>n\<^esup> @ rs # 0\<^bsup>pstr - n\<^esup> @ 0\<^bsup>length ys \<^esup> @ lm @ + 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)" +proof - + from h and pdef and starts have k1: + "\ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = ss \ + bp = empty pstr n" + by(drule_tac restore_rs_prog_ex, auto) + from k1 show "?thesis" + proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE) + fix ap bp apa cp + assume "aprog = ap [+] bp [+] cp \ length ap = ss \ + bp = recursive.empty pstr n" + thus"?thesis" + apply(simp, rule_tac abc_append_exc1, simp_all add: starts h) + apply(insert empty_ex[of pstr n "0\<^bsup>pstr\<^esup> @ rs # 0\<^bsup>length gs\<^esup> @ + lm @ 0\<^bsup>a_md - Suc (pstr + length gs + n)\<^esup> @ suf_lm"], simp) + apply(subgoal_tac "pstr > n", simp) + apply(erule_tac exE, rule_tac x = stp in exI, + simp add: nth_append list_update_append) + apply(simp add: pdef) + done + qed +qed + +lemma [simp]:"xs \ [] \ length xs \ Suc 0" +by(case_tac xs, auto) + +lemma [simp]: "n < max (Suc n) (max ba (Max (((\(aprog, p, n). n) o + rec_ci) ` set gs)))" +by(simp) + +lemma restore_paras_prog_ex: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_ci f = (a, aa, ba); + Max (set (Suc n # ba # map (\(aprog, p, n). n) + (map rec_ci (f # gs)))) = pstr; + ss = (\(ap, pos, n)\map rec_ci gs. length ap) + + 8 * length gs + 3 * n + length a + 6\ + \ \ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = ss \ + bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n" +apply(simp add: rec_ci.simps) +apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) + [+] mv_boxes 0 (Suc (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + + length gs)) n [+] mv_boxes (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs) [+] + a [+] recursive.empty aa (max (Suc n) + (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] + empty_boxes (length gs) [+] + recursive.empty (max (Suc n) (Max (insert ba + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n" in exI, simp add: cn_merge_gs_len) +apply(rule_tac x = "[]" in exI, auto simp: abc_append_commute) +done + +lemma restore_paras: + assumes h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Cn n f gs) lm rs" + "\k(aprog, p, n). n) + (map rec_ci (f # gs))))" + and starts: "ss = (\(ap, pos, n)\map rec_ci gs. length ap) + + 8 * length gs + 3 * n + length a + 6" + shows "\stp. abc_steps_l (ss, 0\<^bsup>n\<^esup> @ rs # 0\<^bsup>pstr - n+ length ys\<^esup> @ + lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) + aprog stp = (ss + 3 * n, lm @ rs # 0\<^bsup>a_md - Suc n\<^esup> @ suf_lm)" +proof - + thm rec_ci.simps + from h and pdef and starts have k1: + "\ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = ss \ + bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n" + by(drule_tac restore_paras_prog_ex, auto) + from k1 show "?thesis" + proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE) + fix ap bp apa cp + assume "aprog = ap [+] bp [+] cp \ length ap = ss \ + bp = mv_boxes (pstr + Suc (length gs)) 0 n" + thus"?thesis" + apply(simp, rule_tac abc_append_exc1, simp_all add: starts h) + apply(insert mv_boxes_ex2[of n "pstr + Suc (length gs)" 0 "[]" + "rs # 0\<^bsup>pstr - n + length gs\<^esup>" "lm" + "0\<^bsup>a_md - Suc (pstr + length gs + n)\<^esup> @ suf_lm"], simp) + apply(subgoal_tac "pstr > n \ + a_md > pstr + length gs + n \ length lm = n" , simp add: exponent_add_iff h) + using h pdef + apply(simp) + apply(frule_tac a = a and + aa = aa and ba = ba in ci_cn_md_def, simp, simp) + apply(subgoal_tac "length lm = rs_pos", + simp add: ci_cn_para_eq, erule_tac para_pattern, simp) + done + qed +qed + +lemma ci_cn_length: + "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); + rec_calc_rel (Cn n f gs) lm rs; + rec_ci f = (a, aa, ba)\ + \ length aprog = (\(ap, pos, n)\map rec_ci gs. length ap) + + 8 * length gs + 6 * n + length a + 6" +apply(simp add: rec_ci.simps, auto simp: cn_merge_gs_len) +done + + +lemma cn_case: + assumes ind: + "\x aprog a_md rs_pos rs suf_lm lm. + \x \ set (f # gs); + rec_ci x = (aprog, rs_pos, a_md); + rec_calc_rel x lm rs\ + \ \stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Cn n f gs) lm rs" + + shows "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp + = (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" +apply(insert h, case_tac "rec_ci f", rule_tac calc_cn_reverse, simp) +proof - + fix a b c ys + let ?pstr = "Max (set (Suc n # c # (map (\(aprog, p, n). n) + (map rec_ci (f # gs)))))" + let ?gs_len = "listsum (map (\ (ap, pos, n). length ap) + (map rec_ci (gs)))" + assume g: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" + "rec_calc_rel (Cn n f gs) lm rs" + "\k stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (?gs_len + 3 * length gs, lm @ 0\<^bsup>?pstr - n\<^esup> @ ys @ + 0\<^bsup>a_md - ?pstr - length ys\<^esup> @ suf_lm)" + apply(rule_tac a = a and aa = b and ba = c in cn_calc_gs) + apply(rule_tac ind, auto) + done + thm rec_ci.simps + from g have k2: + "\ stp. abc_steps_l (?gs_len + 3 * length gs, lm @ + 0\<^bsup>?pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - ?pstr - length ys\<^esup> @ suf_lm) aprog stp = + (?gs_len + 3 * length gs + 3 * n, 0\<^bsup>?pstr\<^esup> @ ys @ 0 # lm @ + 0\<^bsup>a_md - Suc (?pstr + length ys + n )\<^esup> @ suf_lm)" + thm save_paras + apply(erule_tac ba = c in save_paras, auto intro: ci_cn_para_eq) + done + from g have k3: + "\ stp. abc_steps_l (?gs_len + 3 * length gs + 3 * n, + 0\<^bsup>?pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup> @ suf_lm) aprog stp = + (?gs_len + 6 * length gs + 3 * n, + ys @ 0\<^bsup>?pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup> @ suf_lm)" + apply(erule_tac ba = c in reset_new_paras, + auto intro: ci_cn_para_eq) + using para_pattern[of f a b c ys rs] + apply(simp) + done + from g have k4: + "\stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n, + ys @ 0\<^bsup>?pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup> @ suf_lm) aprog stp = + (?gs_len + 6 * length gs + 3 * n + length a, + ys @ rs # 0\<^bsup>?pstr \<^esup> @ lm @ 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup> @ suf_lm)" + apply(rule_tac ba = c in calc_cn_f, rule_tac ind, auto) + done +thm rec_ci.simps + from g h have k5: + "\ stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + length a, + ys @ rs # 0\<^bsup>?pstr \<^esup>@ lm @ 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup> @ suf_lm) + aprog stp = + (?gs_len + 6 * length gs + 3 * n + length a + 3, + ys @ 0\<^bsup>?pstr - length ys\<^esup> @ rs # 0\<^bsup>length ys\<^esup> @ lm @ + 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup> @ suf_lm)" + apply(rule_tac save_rs, auto simp: h) + done + thm rec_ci.simps + thm empty_boxes.simps + from g have k6: + "\ stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + + length a + 3, ys @ 0\<^bsup>?pstr - length ys\<^esup> @ rs # 0\<^bsup>length ys\<^esup> @ lm @ + 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup> @ suf_lm) + aprog stp = + (?gs_len + 8 * length gs + 3 *n + length a + 3, + 0\<^bsup>?pstr \<^esup> @ rs # 0\<^bsup>length ys\<^esup> @ lm @ + 0\<^bsup>a_md -Suc (?pstr + length ys + n)\<^esup> @ suf_lm)" + apply(drule_tac suf_lm = suf_lm in empty_paras, auto) + apply(rule_tac x = stp in exI, simp) + done + from g have k7: + "\ stp. abc_steps_l (?gs_len + 8 * length gs + 3 *n + + length a + 3, 0\<^bsup>?pstr \<^esup> @ rs # 0\<^bsup>length ys\<^esup> @ lm @ + 0\<^bsup>a_md -Suc (?pstr + length ys + n)\<^esup> @ suf_lm) aprog stp = + (?gs_len + 8 * length gs + 3 * n + length a + 6, + 0\<^bsup>n\<^esup> @ rs # 0\<^bsup>?pstr - n\<^esup> @ 0\<^bsup>length ys\<^esup> @ lm @ + 0\<^bsup>a_md -Suc (?pstr + length ys + n) \<^esup> @ suf_lm)" + apply(drule_tac suf_lm = suf_lm in restore_rs, auto) + apply(rule_tac x = stp in exI, simp) + done + from g have k8: "\ stp. abc_steps_l (?gs_len + 8 * length gs + + 3 * n + length a + 6, + 0\<^bsup>n\<^esup> @ rs # 0\<^bsup>?pstr - n\<^esup> @ 0\<^bsup>length ys\<^esup> @ lm @ + 0\<^bsup>a_md -Suc (?pstr + length ys + n) \<^esup> @ suf_lm) aprog stp = + (?gs_len + 8 * length gs + 6 * n + length a + 6, + lm @ rs # 0\<^bsup>a_md - Suc n \<^esup>@ suf_lm)" + apply(drule_tac suf_lm = suf_lm in restore_paras, auto) + apply(simp add: exponent_add_iff) + apply(rule_tac x = stp in exI, simp) + done + from g have j1: + "length aprog = ?gs_len + 8 * length gs + 6 * n + length a + 6" + by(drule_tac a = a and aa = b and ba = c in ci_cn_length, + simp, simp, simp) + from g have j2: "rs_pos = n" + by(simp add: ci_cn_para_eq) + from k1 and k2 and k3 and k4 and k5 and k6 and k7 and k8 + and j1 and j2 show + "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = + (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)" + apply(auto) + apply(rule_tac x = "stp + stpa + stpb + stpc + + stpd + stpe + stpf + stpg" in exI, simp add: abc_steps_add) + done +qed + +text {* + Correctness of the complier (terminate case), which says if the execution of + a recursive function @{text "recf"} terminates and gives result, then + the Abacus program compiled from @{text "recf"} termintes and gives the same result. + Additionally, to facilitate induction proof, we append @{text "anything"} to the + end of Abacus memory. +*} + +lemma aba_rec_equality: + "\rec_ci recf = (ap, arity, fp); + rec_calc_rel recf args r\ + \ (\ stp. (abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp) = + (length ap, args@[r]@0\<^bsup>fp - arity - 1\<^esup> @ anything))" +apply(induct arbitrary: ap fp arity r anything args + rule: rec_ci.induct) +prefer 5 +proof(case_tac "rec_ci g", case_tac "rec_ci f", simp) + fix n f g ap fp arity r anything args a b c aa ba ca + assume f_ind: + "\ap fp arity r anything args. + \aa = ap \ ba = arity \ ca = fp; rec_calc_rel f args r\ \ + \stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp = + (length ap, args @ r # 0\<^bsup>fp - Suc arity\<^esup> @ anything)" + and g_ind: + "\x xa y xb ya ap fp arity r anything args. + \x = (aa, ba, ca); xa = aa \ y = (ba, ca); xb = ba \ ya = ca; + a = ap \ b = arity \ c = fp; rec_calc_rel g args r\ + \ \stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp = + (length ap, args @ r # 0\<^bsup>fp - Suc arity\<^esup> @ anything)" + and h: "rec_ci (Pr n f g) = (ap, arity, fp)" + "rec_calc_rel (Pr n f g) args r" + "rec_ci g = (a, b, c)" + "rec_ci f = (aa, ba, ca)" + from h have nf_ind: + "\ args r anything. rec_calc_rel f args r \ + \stp. abc_steps_l (0, args @ 0\<^bsup>ca - ba\<^esup> @ anything) aa stp = + (length aa, args @ r # 0\<^bsup>ca - Suc ba\<^esup> @ anything)" + and ng_ind: + "\ args r anything. rec_calc_rel g args r \ + \stp. abc_steps_l (0, args @ 0\<^bsup>c - b\<^esup> @ anything) a stp = + (length a, args @ r # 0\<^bsup>c - Suc b \<^esup> @ anything)" + apply(insert f_ind[of aa ba ca], simp) + apply(insert g_ind[of "(aa, ba, ca)" aa "(ba, ca)" ba ca a b c], + simp) + done + from nf_ind and ng_ind and h show + "\stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp = + (length ap, args @ r # 0\<^bsup>fp - Suc arity\<^esup> @ anything)" + apply(auto intro: nf_ind ng_ind pr_case) + done +next + fix ap fp arity r anything args + assume h: + "rec_ci z = (ap, arity, fp)" "rec_calc_rel z args r" + thus "\stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp = + (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)" + by (rule_tac z_case) +next + fix ap fp arity r anything args + assume h: + "rec_ci s = (ap, arity, fp)" + "rec_calc_rel s args r" + thus + "\stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp = + (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)" + by(erule_tac s_case, simp) +next + fix m n ap fp arity r anything args + assume h: "rec_ci (id m n) = (ap, arity, fp)" + "rec_calc_rel (id m n) args r" + thus "\stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp + = (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)" + by(erule_tac id_case) +next + fix n f gs ap fp arity r anything args + assume ind: "\x ap fp arity r anything args. + \x \ set (f # gs); + rec_ci x = (ap, arity, fp); + rec_calc_rel x args r\ + \ \stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp = + (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)" + and h: "rec_ci (Cn n f gs) = (ap, arity, fp)" + "rec_calc_rel (Cn n f gs) args r" + from h show + "\stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) + ap stp = (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)" + apply(rule_tac cn_case, rule_tac ind, auto) + done +next + fix n f ap fp arity r anything args + assume ind: + "\ap fp arity r anything args. + \rec_ci f = (ap, arity, fp); rec_calc_rel f args r\ \ + \stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp = + (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)" + and h: "rec_ci (Mn n f) = (ap, arity, fp)" + "rec_calc_rel (Mn n f) args r" + from h show + "\stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp = + (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)" + apply(rule_tac mn_case, rule_tac ind, auto) + done +qed + + +thm abc_append_state_in_exc +lemma abc_append_uhalt1: + "\\ stp. (\ (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp); + p = ap [+] bp [+] cp\ + \ \ stp. (\ (ss, e). ss < length p) + (abc_steps_l (length ap, lm) p stp)" +apply(auto) +apply(erule_tac x = stp in allE, auto) +apply(frule_tac ap = ap and cp = cp in abc_append_state_in_exc, auto) +done + + +lemma abc_append_unhalt2: + "\abc_steps_l (0, am) ap stp = (length ap, lm); bp \ []; + \ stp. (\ (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp); + p = ap [+] bp [+] cp\ + \ \ stp. (\ (ss, e). ss < length p) (abc_steps_l (0, am) p stp)" +proof - + assume h: + "abc_steps_l (0, am) ap stp = (length ap, lm)" + "bp \ []" + "\ stp. (\ (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp)" + "p = ap [+] bp [+] cp" + have "\ stp. (abc_steps_l (0, am) p stp) = (length ap, lm)" + using h + thm abc_add_exc1 + apply(simp add: abc_append.simps) + apply(rule_tac abc_add_exc1, auto) + done + from this obtain stpa where g1: + "(abc_steps_l (0, am) p stpa) = (length ap, lm)" .. + moreover have g2: "\ stp. (\ (ss, e). ss < length p) + (abc_steps_l (length ap, lm) p stp)" + using h + apply(erule_tac abc_append_uhalt1, simp) + done + moreover from g1 and g2 have + "\ stp. (\ (ss, e). ss < length p) + (abc_steps_l (0, am) p (stpa + stp))" + apply(simp add: abc_steps_add) + done + thus "\ stp. (\ (ss, e). ss < length p) + (abc_steps_l (0, am) p stp)" + apply(rule_tac allI, auto) + apply(case_tac "stp \ stpa") + apply(erule_tac x = "stp - stpa" in allE, simp) + proof - + fix stp a b + assume g3: "abc_steps_l (0, am) p stp = (a, b)" + "\ stpa \ stp" + thus "a < length p" + using g1 h + apply(case_tac "a < length p", simp, simp) + apply(subgoal_tac "\ d. stpa = stp + d") + using abc_state_keep[of p a b "stpa - stp"] + apply(erule_tac exE, simp add: abc_steps_add) + apply(rule_tac x = "stpa - stp" in exI, simp) + done + qed +qed + +text {* + Correctness of the complier (non-terminating case for Mn). There are many cases when a + recursive function does not terminate. For the purpose of Uiversal Turing Machine, we only + need to prove the case for @{text "Mn"} and @{text "Cn"}. + This lemma is for @{text "Mn"}. For @{text "Mn n f"}, this lemma describes what + happens when @{text "f"} always terminates but always does not return zero, so that + @{text "Mn"} has to loop forever. + *} + +lemma Mn_unhalt: + assumes mn_rf: "rf = Mn n f" + and compiled_mnrf: "rec_ci rf = (aprog, rs_pos, a_md)" + and compiled_f: "rec_ci f = (aprog', rs_pos', a_md')" + and args: "length lm = n" + and unhalt_condition: "\ y. (\ rs. rec_calc_rel f (lm @ [y]) rs \ rs \ 0)" + shows "\ stp. case abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) + aprog stp of (ss, e) \ ss < length aprog" + using mn_rf compiled_mnrf compiled_f args unhalt_condition +proof(rule_tac allI) + fix stp + assume h: "rf = Mn n f" + "rec_ci rf = (aprog, rs_pos, a_md)" + "rec_ci f = (aprog', rs_pos', a_md')" + "\y. \rs. rec_calc_rel f (lm @ [y]) rs \ rs \ 0" "length lm = n" + thm mn_ind_step + have "\stpa \ stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) aprog stpa + = (0, lm @ stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + proof(induct stp, auto) + show "\stpa. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) + aprog stpa = (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps) + done + next + fix stp stpa + assume g1: "stp \ stpa" + and g2: "abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) + aprog stpa + = (0, lm @ stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + have "\rs. rec_calc_rel f (lm @ [stp]) rs \ rs \ 0" + using h + apply(erule_tac x = stp in allE, simp) + done + from this obtain rs where g3: + "rec_calc_rel f (lm @ [stp]) rs \ rs \ 0" .. + hence "\ stpb. abc_steps_l (0, lm @ stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ + suf_lm) aprog stpb + = (0, lm @ Suc stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + using h + apply(rule_tac mn_ind_step) + apply(rule_tac aba_rec_equality, simp, simp) + proof - + show "rec_ci f = ((aprog', rs_pos', a_md'))" using h by simp + next + show "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" using h by simp + next + show "rec_calc_rel f (lm @ [stp]) rs" using g3 by simp + next + show "0 < rs" using g3 by simp + next + show "Suc rs_pos < a_md" + using g3 h + apply(auto) + apply(frule_tac f = f in para_pattern, simp, simp) + apply(simp add: rec_ci.simps, auto) + apply(subgoal_tac "Suc (length lm) < a_md'") + apply(arith) + apply(simp add: ci_ad_ge_paras) + done + next + show "rs_pos' = Suc rs_pos" + using g3 h + apply(auto) + apply(frule_tac f = f in para_pattern, simp, simp) + apply(simp add: rec_ci.simps) + done + qed + thus "\stpa\Suc stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ + suf_lm) aprog stpa + = (0, lm @ Suc stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" + using g2 + apply(erule_tac exE) + apply(case_tac "stpb = 0", simp add: abc_steps_l.simps) + apply(rule_tac x = "stpa + stpb" in exI, simp add: + abc_steps_add) + using g1 + apply(arith) + done + qed + from this obtain stpa where + "stp \ stpa \ abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) + aprog stpa = (0, lm @ stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" .. + thus "case abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp + of (ss, e) \ ss < length aprog" + apply(case_tac "abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog + stp", simp, case_tac "a \ length aprog", + simp, simp) + apply(subgoal_tac "\ d. stpa = stp + d", erule_tac exE) + apply(subgoal_tac "lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm = lm @ 0 # + 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm", simp add: abc_steps_add) + apply(frule_tac as = a and lm = b and stp = d in abc_state_keep, + simp) + using h + apply(simp add: rec_ci.simps, simp, + simp only: exp_ind_def[THEN sym]) + apply(case_tac rs_pos, simp, simp) + apply(rule_tac x = "stpa - stp" in exI, simp, simp) + done +qed + + +lemma abc_append_cons_eq[intro!]: + "\ap = bp; cp = dp\ \ ap [+] cp = bp [+] dp" +by simp + +lemma cn_merge_gs_split: + "\i < length gs; rec_ci (gs!i) = (ga, gb, gc)\ \ + cn_merge_gs (map rec_ci gs) p = + cn_merge_gs (map rec_ci (take i gs)) p [+] ga [+] + empty gb (p + i) [+] + cn_merge_gs (map rec_ci (drop (Suc i) gs)) (p + Suc i)" +apply(induct i arbitrary: gs p, case_tac gs, simp, simp) +apply(case_tac gs, simp, case_tac "rec_ci a", + simp add: abc_append_commute[THEN sym]) +done + +text {* + Correctness of the complier (non-terminating case for Mn). There are many cases when a + recursive function does not terminate. For the purpose of Uiversal Turing Machine, we only + need to prove the case for @{text "Mn"} and @{text "Cn"}. + This lemma is for @{text "Cn"}. For @{text "Cn f g1 g2 \gi, gi+1, \ gn"}, this lemma describes what + happens when every one of @{text "g1, g2, \ gi"} terminates, but + @{text "gi+1"} does not terminate, so that whole function @{text "Cn f g1 g2 \gi, gi+1, \ gn"} + does not terminate. + *} + +lemma cn_gi_uhalt: + assumes cn_recf: "rf = Cn n f gs" + and compiled_cn_recf: "rec_ci rf = (aprog, rs_pos, a_md)" + and args_length: "length lm = n" + and exist_unhalt_recf: "i < length gs" "gi = gs ! i" + and complied_unhalt_recf: "rec_ci gi = (ga, gb, gc)" "gb = n" + and all_halt_before_gi: "\ j < i. (\ rs. rec_calc_rel (gs!j) lm rs)" + and unhalt_condition: "\ slm. \ stp. case abc_steps_l (0, lm @ 0\<^bsup>gc - gb\<^esup> @ slm) + ga stp of (se, e) \ se < length ga" + shows " \ stp. case abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suflm) aprog + stp of (ss, e) \ ss < length aprog" + using cn_recf compiled_cn_recf args_length exist_unhalt_recf complied_unhalt_recf + all_halt_before_gi unhalt_condition +proof(case_tac "rec_ci f", simp) + fix a b c + assume h1: "rf = Cn n f gs" + "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" + "length lm = n" + "gi = gs ! i" + "rec_ci (gs!i) = (ga, n, gc)" + "gb = n" "rec_ci f = (a, b, c)" + and h2: "\jrs. rec_calc_rel (gs ! j) lm rs" + "i < length gs" + and ind: + "\ slm. \ stp. case abc_steps_l (0, lm @ 0\<^bsup>gc - n\<^esup> @ slm) ga stp of (se, e) \ se < length ga" + have h3: "rs_pos = n" + using h1 + by(rule_tac ci_cn_para_eq, simp) + let ?ggs = "take i gs" + have "\ ys. (length ys = i \ + (\ k < i. rec_calc_rel (?ggs ! k) lm (ys ! k)))" + using h2 + apply(induct i, simp, simp) + apply(erule_tac exE) + apply(erule_tac x = ia in allE, simp) + apply(erule_tac exE) + apply(rule_tac x = "ys @ [x]" in exI, simp add: nth_append, auto) + apply(subgoal_tac "k = length ys", simp, simp) + done + from this obtain ys where g1: + "(length ys = i \ (\ k < i. rec_calc_rel (?ggs ! k) + lm (ys ! k)))" .. + let ?pstr = "Max (set (Suc n # c # map (\(aprog, p, n). n) + (map rec_ci (f # gs))))" + have "\stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suflm) + (cn_merge_gs (map rec_ci ?ggs) ?pstr) stp = + (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) ?ggs) + + 3 * length ?ggs, lm @ 0\<^bsup>?pstr - n\<^esup> @ ys @ 0\<^bsup>a_md -(?pstr + length ?ggs)\<^esup> @ + suflm) " + apply(rule_tac cn_merge_gs_ex) + apply(rule_tac aba_rec_equality, simp, simp) + using h1 + apply(simp add: rec_ci.simps, auto) + using g1 + apply(simp) + using h2 g1 + apply(simp) + apply(rule_tac min_max.le_supI2) + apply(rule_tac Max_ge, simp, simp, rule_tac disjI2) + apply(subgoal_tac "aa \ set gs", simp) + using h2 + apply(rule_tac A = "set (take i gs)" in subsetD, + simp add: set_take_subset, simp) + done + thm cn_merge_gs.simps + from this obtain stpa where g2: + "abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suflm) + (cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa = + (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) ?ggs) + + 3 * length ?ggs, lm @ 0\<^bsup>?pstr - n\<^esup> @ ys @ 0\<^bsup>a_md -(?pstr + length ?ggs)\<^esup> @ + suflm)" .. + moreover have + "\ cp. aprog = (cn_merge_gs + (map rec_ci ?ggs) ?pstr) [+] ga [+] cp" + using h1 + apply(simp add: rec_ci.simps) + apply(rule_tac x = "empty n (?pstr + i) [+] + (cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?pstr + Suc i)) + [+]mv_boxes 0 (Suc (max (Suc n) (Max (insert c + (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + + length gs)) n [+] mv_boxes (max (Suc n) (Max (insert c + (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs) [+] + a [+] recursive.empty b (max (Suc n) + (Max (insert c (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] + empty_boxes (length gs) [+] recursive.empty (max (Suc n) + (Max (insert c (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] + mv_boxes (Suc (max (Suc n) (Max (insert c + (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) 0 n" in exI) + apply(simp add: abc_append_commute [THEN sym]) + apply(auto) + using cn_merge_gs_split[of i gs ga "length lm" gc + "(max (Suc (length lm)) + (Max (insert c (((\(aprog, p, n). n) \ rec_ci) ` set gs))))"] + h2 + apply(simp) + done + from this obtain cp where g3: + "aprog = (cn_merge_gs (map rec_ci ?ggs) ?pstr) [+] ga [+] cp" .. + show "\ stp. case abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suflm) + aprog stp of (ss, e) \ ss < length aprog" + proof(rule_tac abc_append_unhalt2) + show "abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suflm) ( + cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa = + (length ((cn_merge_gs (map rec_ci ?ggs) ?pstr)), + lm @ 0\<^bsup>?pstr - n\<^esup> @ ys @ 0\<^bsup>a_md -(?pstr + length ?ggs)\<^esup> @ suflm)" + using h3 g2 + apply(simp add: cn_merge_gs_length) + done + next + show "ga \ []" + using h1 + apply(simp add: rec_ci_not_null) + done + next + show "\stp. case abc_steps_l (0, lm @ 0\<^bsup>?pstr - n\<^esup> @ ys + @ 0\<^bsup>a_md - (?pstr + length (take i gs))\<^esup> @ suflm) ga stp of + (ss, e) \ ss < length ga" + using ind[of "0\<^bsup>?pstr -gc\<^esup> @ ys @ 0\<^bsup>a_md - (?pstr + length (take i gs))\<^esup> + @ suflm"] + apply(subgoal_tac "lm @ 0\<^bsup>?pstr - n\<^esup> @ ys + @ 0\<^bsup>a_md - (?pstr + length (take i gs))\<^esup> @ suflm + = lm @ 0\<^bsup>gc - n \<^esup>@ + 0\<^bsup>?pstr -gc\<^esup> @ ys @ 0\<^bsup>a_md - (?pstr + length (take i gs))\<^esup> @ suflm", simp) + apply(simp add: exponent_def replicate_add[THEN sym]) + apply(subgoal_tac "gc > n \ ?pstr \ gc") + apply(erule_tac conjE) + apply(simp add: h1) + using h1 + apply(auto) + apply(rule_tac min_max.le_supI2) + apply(rule_tac Max_ge, simp, simp) + apply(rule_tac disjI2) + using h2 + thm rev_image_eqI + apply(rule_tac x = "gs!i" in rev_image_eqI, simp, simp) + done + next + show "aprog = cn_merge_gs (map rec_ci (take i gs)) + ?pstr [+] ga [+] cp" + using g3 by simp + qed +qed + + +lemma abc_rec_halt_eq': + "\rec_ci re = (ap, ary, fp); + rec_calc_rel re args r\ + \ (\ stp. (abc_steps_l (0, args @ 0\<^bsup>fp - ary\<^esup>) ap stp) = + (length ap, args@[r]@0\<^bsup>fp - ary - 1\<^esup>))" +using aba_rec_equality[of re ap ary fp args r "[]"] +by simp + +thm abc_step_l.simps +definition dummy_abc :: "nat \ abc_inst list" +where +"dummy_abc k = [Inc k, Dec k 0, Goto 3]" + +lemma abc_rec_halt_eq'': + "\rec_ci re = (aprog, rs_pos, a_md); + rec_calc_rel re lm rs\ + \ (\ stp lm' m. (abc_steps_l (0, lm) aprog stp) = + (length aprog, lm') \ abc_list_crsp lm' (lm @ rs # 0\<^bsup>m\<^esup>))" +apply(frule_tac abc_rec_halt_eq', auto) +apply(drule_tac abc_list_crsp_steps) +apply(rule_tac rec_ci_not_null, simp) +apply(erule_tac exE, rule_tac x = stp in exI, + auto simp: abc_list_crsp_def) +done + +lemma [simp]: "length (dummy_abc (length lm)) = 3" +apply(simp add: dummy_abc_def) +done + +lemma [simp]: "dummy_abc (length lm) \ []" +apply(simp add: dummy_abc_def) +done + +lemma dummy_abc_steps_ex: + "\bstp. abc_steps_l (0, lm') (dummy_abc (length lm)) bstp = + ((Suc (Suc (Suc 0))), abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)))" +apply(rule_tac x = "Suc (Suc (Suc 0))" in exI) +apply(auto simp: abc_steps_l.simps abc_step_l.simps + dummy_abc_def abc_fetch.simps) +apply(auto simp: abc_lm_s.simps abc_lm_v.simps nth_append) +apply(simp add: butlast_append) +done + +lemma [elim]: + "lm @ rs # 0\<^bsup>m\<^esup> = lm' @ 0\<^bsup>n\<^esup> \ + \m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) = + lm @ rs # 0\<^bsup>m\<^esup>" +proof(cases "length lm' > length lm") + case True + assume h: "lm @ rs # 0\<^bsup>m\<^esup> = lm' @ 0\<^bsup>n\<^esup>" "length lm < length lm'" + hence "m \ n" + apply(drule_tac list_length) + apply(simp) + done + hence "\ d. m = d + n" + apply(rule_tac x = "m - n" in exI, simp) + done + from this obtain d where "m = d + n" .. + from h and this show "?thesis" + apply(auto simp: abc_lm_s.simps abc_lm_v.simps + exponent_def replicate_add) + done +next + case False + assume h:"lm @ rs # 0\<^bsup>m\<^esup> = lm' @ 0\<^bsup>n\<^esup>" + and g: "\ length lm < length lm'" + have "take (Suc (length lm)) (lm @ rs # 0\<^bsup>m\<^esup>) = + take (Suc (length lm)) (lm' @ 0\<^bsup>n\<^esup>)" + using h by simp + moreover have "n \ (Suc (length lm) - length lm')" + using h g + apply(drule_tac list_length) + apply(simp) + done + ultimately show + "\m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) = + lm @ rs # 0\<^bsup>m\<^esup>" + using g h + apply(simp add: abc_lm_s.simps abc_lm_v.simps + exponent_def min_def) + apply(rule_tac x = 0 in exI, + simp add:replicate_append_same replicate_Suc[THEN sym] + del:replicate_Suc) + done +qed + +lemma [elim]: + "abc_list_crsp lm' (lm @ rs # 0\<^bsup>m\<^esup>) + \ \m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) + = lm @ rs # 0\<^bsup>m\<^esup>" +apply(auto simp: abc_list_crsp_def) +apply(simp add: abc_lm_v.simps abc_lm_s.simps) +apply(rule_tac x = "m + n" in exI, + simp add: exponent_def replicate_add) +done + + +lemma abc_append_dummy_complie: + "\rec_ci recf = (ap, ary, fp); + rec_calc_rel recf args r; + length args = k\ + \ (\ stp m. (abc_steps_l (0, args) (ap [+] dummy_abc k) stp) = + (length ap + 3, args @ r # 0\<^bsup>m\<^esup>))" +apply(drule_tac abc_rec_halt_eq'', auto simp: numeral_3_eq_3) +proof - + fix stp lm' m + assume h: "rec_calc_rel recf args r" + "abc_steps_l (0, args) ap stp = (length ap, lm')" + "abc_list_crsp lm' (args @ r # 0\<^bsup>m\<^esup>)" + thm abc_append_exc2 + thm abc_lm_s.simps + have "\stp. abc_steps_l (0, args) (ap [+] + (dummy_abc (length args))) stp = (length ap + 3, + abc_lm_s lm' (length args) (abc_lm_v lm' (length args)))" + using h + apply(rule_tac bm = lm' in abc_append_exc2, + auto intro: dummy_abc_steps_ex simp: numeral_3_eq_3) + done + thus "\stp m. abc_steps_l (0, args) (ap [+] + dummy_abc (length args)) stp = (Suc (Suc (Suc (length ap))), args @ r # 0\<^bsup>m\<^esup>)" + using h + apply(erule_tac exE) + apply(rule_tac x = stpa in exI, auto) + done +qed + +lemma [simp]: "length (dummy_abc k) = 3" +apply(simp add: dummy_abc_def) +done + +lemma [simp]: "length args = k \ abc_lm_v (args @ r # 0\<^bsup>m\<^esup>) k = r " +apply(simp add: abc_lm_v.simps nth_append) +done + +lemma t_compiled_by_rec: + "\rec_ci recf = (ap, ary, fp); + rec_calc_rel recf args r; + length args = k; + ly = layout_of (ap [+] dummy_abc k); + mop_ss = start_of ly (length (ap [+] dummy_abc k)); + tp = tm_of (ap [+] dummy_abc k)\ + \ \ stp m l. steps (Suc 0, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) (tp @ (tMp k (mop_ss - 1))) stp + = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc r\<^esup> @ Bk\<^bsup>l\<^esup>)" + using abc_append_dummy_complie[of recf ap ary fp args r k] +apply(simp) +apply(erule_tac exE)+ +apply(frule_tac tprog = tp and as = "length ap + 3" and n = k + and ires = ires and rn = rn in abacus_turing_eq_halt, simp_all, simp) +apply(erule_tac exE)+ +apply(simp) +apply(rule_tac x = stpa in exI, rule_tac x = ma in exI, rule_tac x = l in exI, simp) +done + +thm tms_of.simps + +lemma [simp]: + "list_all (\(acn, s). s \ Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))) xs \ + list_all (\(acn, s). s \ Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))))) xs" +apply(induct xs, simp, simp) +apply(case_tac a, simp) +done + +(* +lemma [simp]: "t_correct (tMp n 0)" +apply(simp add: t_correct.simps tMp.simps shift_length mp_up_def iseven_def, auto) +apply(rule_tac x = "2*n + 6" in exI, simp) +apply(induct n, auto simp: mop_bef.simps) +apply(simp add: tshift.simps) +done +*) + +lemma tshift_append: "tshift (xs @ ys) n = tshift xs n @ tshift ys n" +apply(simp add: tshift.simps) +done + +lemma [simp]: "length (tMp n ss) = 4 * n + 12" +apply(auto simp: tMp.simps tshift_append shift_length mp_up_def) +done + +lemma length_tm_even[intro]: "\x. length (tm_of ap) = 2*x" +apply(subgoal_tac "t_ncorrect (tm_of ap)") +apply(simp add: t_ncorrect.simps, auto) +done + +lemma [simp]: "k < length ap \ tms_of ap ! k = + ci (layout_of ap) (start_of (layout_of ap) k) (ap ! k)" +apply(simp add: tms_of.simps tpairs_of.simps) +done + +lemma [elim]: "\k < length ap; ap ! k = Inc n; + (a, b) \ set (abacus.tshift (abacus.tshift tinc_b (2 * n)) + (start_of (layout_of ap) k - Suc 0))\ + \ b \ start_of (layout_of ap) (length ap)" +apply(subgoal_tac "b \ start_of (layout_of ap) (Suc k)") +apply(subgoal_tac "start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap) ") +apply(arith) +apply(case_tac "Suc k = length ap", simp) +apply(rule_tac start_of_le, simp) +apply(auto simp: tinc_b_def tshift.simps start_of.simps + layout_of.simps length_of.simps startof_not0) +done + +lemma findnth_le[elim]: "(a, b) \ set (abacus.tshift (findnth n) (start_of (layout_of ap) k - Suc 0)) + \ b \ Suc (start_of (layout_of ap) k + 2 * n)" +apply(induct n, simp add: findnth.simps tshift.simps) +apply(simp add: findnth.simps tshift_append, auto) +apply(auto simp: tshift.simps) +done + + +lemma [elim]: "\k < length ap; ap ! k = Inc n; (a, b) \ + set (abacus.tshift (findnth n) (start_of (layout_of ap) k - Suc 0))\ + \ b \ start_of (layout_of ap) (length ap)" +apply(subgoal_tac "b \ start_of (layout_of ap) (Suc k)") +apply(subgoal_tac "start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap) ") +apply(arith) +apply(case_tac "Suc k = length ap", simp) +apply(rule_tac start_of_le, simp) +apply(subgoal_tac "b \ start_of (layout_of ap) k + 2*n + 1 \ + start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) (Suc k)", auto) +apply(auto simp: tinc_b_def tshift.simps start_of.simps + layout_of.simps length_of.simps startof_not0) +done + +lemma start_of_eq: "length ap < as \ start_of (layout_of ap) as = start_of (layout_of ap) (length ap)" +apply(induct as, simp) +apply(case_tac "length ap < as", simp add: start_of.simps) +apply(subgoal_tac "as = length ap") +apply(simp add: start_of.simps) +apply arith +done + +lemma start_of_all_le: "start_of (layout_of ap) as \ start_of (layout_of ap) (length ap)" +apply(subgoal_tac "as > length ap \ as = length ap \ as < length ap", + auto simp: start_of_eq start_of_le) +done + +lemma [elim]: "\k < length ap; + ap ! k = Dec n e; + (a, b) \ set (abacus.tshift (findnth n) (start_of (layout_of ap) k - Suc 0))\ + \ b \ start_of (layout_of ap) (length ap)" +apply(subgoal_tac "b \ start_of (layout_of ap) k + 2*n + 1 \ + start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) (Suc k) \ + start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap)", auto) +apply(simp add: tshift.simps start_of.simps + layout_of.simps length_of.simps startof_not0) +apply(rule_tac start_of_all_le) +done + +thm length_of.simps +lemma [elim]: "\k < length ap; ap ! k = Dec n e; (a, b) \ set (abacus.tshift (abacus.tshift tdec_b (2 * n)) + (start_of (layout_of ap) k - Suc 0))\ + \ b \ start_of (layout_of ap) (length ap)" +apply(subgoal_tac "2*n + start_of (layout_of ap) k + 16 \ start_of (layout_of ap) (length ap) \ start_of (layout_of ap) k > 0") +prefer 2 +apply(subgoal_tac "2 * n + start_of (layout_of ap) k + 16 = start_of (layout_of ap) (Suc k) + \ start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap)") +apply(simp add: startof_not0, rule_tac conjI) +apply(simp add: start_of.simps layout_of.simps length_of.simps) +apply(rule_tac start_of_all_le) +apply(auto simp: tdec_b_def tshift.simps) +done + +lemma tms_any_less: "\k < length ap; (a, b) \ set (tms_of ap ! k)\ \ b \ start_of (layout_of ap) (length ap)" +apply(simp) +apply(case_tac "ap!k", simp_all add: ci.simps tshift_append, auto intro: start_of_all_le) +done +lemma concat_in: "i < length (concat xs) \ \k < length xs. concat xs ! i \ set (xs ! k)" +apply(induct xs rule: list_tl_induct, simp, simp) +apply(case_tac "i < length (concat list)", simp) +apply(erule_tac exE, rule_tac x = k in exI) +apply(simp add: nth_append) +apply(rule_tac x = "length list" in exI, simp) +apply(simp add: nth_append) +done + +lemma [simp]: "length (tms_of ap) = length ap" +apply(simp add: tms_of.simps tpairs_of.simps) +done + +lemma in_tms: "i < length (tm_of ap) \ \ k < length ap. (tm_of ap ! i) \ set (tms_of ap ! k)" +apply(simp add: tm_of.simps) +using concat_in[of i "tms_of ap"] +by simp + +lemma all_le_start_of: "list_all (\(acn, s). s \ start_of (layout_of ap) (length ap)) (tm_of ap)" +apply(simp add: list_all_length) +apply(rule_tac allI, rule_tac impI) +apply(drule_tac in_tms, auto elim: tms_any_less) +done + +lemma length_ci: "\k < length ap; length (ci ly y (ap ! k)) = 2 * qa\ + \ layout_of ap ! k = qa" +apply(case_tac "ap ! k") +apply(auto simp: layout_of.simps ci.simps + length_of.simps shift_length tinc_b_def tdec_b_def) +done + +lemma [intro]: "length (ci ly y i) mod 2 = 0" +apply(auto simp: ci.simps shift_length tinc_b_def tdec_b_def + split: abc_inst.splits) +done + +lemma [intro]: "listsum (map (length \ (\(x, y). ci ly x y)) zs) mod 2 = 0" +apply(induct zs rule: list_tl_induct, simp) +apply(case_tac a, simp) +apply(subgoal_tac "length (ci ly aa b) mod 2 = 0") +apply(auto) +done + +lemma zip_pre: + "(length ys) \ length ap \ + zip ys ap = zip ys (take (length ys) (ap::'a list))" +proof(induct ys arbitrary: ap, simp, case_tac ap, simp) + fix a ys ap aa list + assume ind: "\(ap::'a list). length ys \ length ap \ + zip ys ap = zip ys (take (length ys) ap)" + and h: "length (a # ys) \ length ap" "(ap::'a list) = aa # (list::'a list)" + from h show "zip (a # ys) ap = zip (a # ys) (take (length (a # ys)) ap)" + using ind[of list] + apply(simp) + done +qed + +lemma start_of_listsum: + "\k \ length ap; length ss = k\ \ start_of (layout_of ap) k = + Suc (listsum (map (length \ (\(x, y). ci ly x y)) (zip ss ap)) div 2)" +proof(induct k arbitrary: ss, simp add: start_of.simps, simp) + fix k ss + assume ind: "\ss. length ss = k \ start_of (layout_of ap) k = + Suc (listsum (map (length \ (\(x, y). ci ly x y)) (zip ss ap)) div 2)" + and h: "Suc k \ length ap" "length (ss::nat list) = Suc k" + have "\ ys y. ss = ys @ [y]" + using h + apply(rule_tac x = "butlast ss" in exI, + rule_tac x = "last ss" in exI) + apply(case_tac "ss = []", auto) + done + from this obtain ys y where k1: "ss = (ys::nat list) @ [y]" + by blast + from h and this have k2: + "start_of (layout_of ap) k = + Suc (listsum (map (length \ (\(x, y). ci ly x y)) (zip ys ap)) div 2)" + apply(rule_tac ind, simp) + done + have k3: "zip ys ap = zip ys (take k ap)" + using zip_pre[of ys ap] k1 h + apply(simp) + done + have k4: "(zip [y] (drop (length ys) ap)) = [(y, ap ! length ys)]" + using k1 h + apply(case_tac "drop (length ys) ap", simp) + apply(subgoal_tac "hd (drop (length ys) ap) = ap ! length ys") + apply(simp) + apply(rule_tac hd_drop_conv_nth, simp) + done + from k1 and h k2 k3 k4 show "start_of (layout_of ap) (Suc k) = + Suc (listsum (map (length \ (\(x, y). ci ly x y)) (zip ss ap)) div 2)" + apply(simp add: zip_append1 start_of.simps) + apply(subgoal_tac + "listsum (map (length \ (\(x, y). ci ly x y)) (zip ys (take k ap))) mod 2 = 0 \ + length (ci ly y (ap!k)) mod 2 = 0") + apply(auto) + apply(rule_tac length_ci, simp, simp) + done +qed + +lemma length_start_of_tm: "start_of (layout_of ap) (length ap) = Suc (length (tm_of ap) div 2)" +apply(simp add: tm_of.simps length_concat tms_of.simps tpairs_of.simps) +apply(rule_tac start_of_listsum, simp, simp) +done + +lemma tm_even: "length (tm_of ap) mod 2 = 0" +apply(subgoal_tac "t_ncorrect (tm_of ap)", auto) +apply(simp add: t_ncorrect.simps) +done + +lemma [elim]: "list_all (\(acn, s). s \ Suc q) xs + \ list_all (\(acn, s). s \ q + (2 * n + 6)) xs" +apply(simp add: list_all_length) +apply(auto) +done + +lemma [simp]: "length mp_up = 12" +apply(simp add: mp_up_def) +done + +lemma [elim]: "\na < 4 * n; tshift (mop_bef n) q ! na = (a, b)\ \ b \ q + (2 * n + 6)" +apply(induct n, simp, simp add: mop_bef.simps nth_append tshift_append shift_length) +apply(case_tac "na < 4*n", simp, simp) +apply(subgoal_tac "na = 4*n \ na = 1 + 4*n \ na = 2 + 4*n \ na = 3 + 4*n", + auto simp: shift_length) +apply(simp_all add: tshift.simps) +done + +lemma mp_up_all_le: "list_all (\(acn, s). s \ q + (2 * n + 6)) + [(R, Suc (Suc (2 * n + q))), (R, Suc (2 * n + q)), + (L, 5 + 2 * n + q), (W0, Suc (Suc (Suc (2 * n + q)))), (R, 4 + 2 * n + q), + (W0, Suc (Suc (Suc (2 * n + q)))), (R, Suc (Suc (2 * n + q))), + (W0, Suc (Suc (Suc (2 * n + q)))), (L, 5 + 2 * n + q), + (L, 6 + 2 * n + q), (R, 0), (L, 6 + 2 * n + q)]" +apply(auto) +done + + +lemma [intro]: "list_all (\(acn, s). s \ q + (2 * n + 6)) (tMp n q)" +apply(auto simp: list_all_length tMp.simps tshift_append nth_append shift_length) +apply(auto simp: tshift.simps mp_up_def) +apply(subgoal_tac "na - 4*n \ 0 \ na - 4 *n < 12", auto split: nat.splits) +apply(insert mp_up_all_le[of q n]) +apply(simp add: list_all_length) +apply(erule_tac x = "na - 4 * n" in allE, simp add: numeral_3_eq_3) +done + +lemma t_compiled_correct: + "\tp = tm_of ap; ly = layout_of ap; mop_ss = start_of ly (length ap)\ \ + t_correct (tp @ tMp n (mop_ss - Suc 0))" + using tm_even[of ap] length_start_of_tm[of ap] all_le_start_of[of ap] +apply(auto simp: t_correct.simps iseven_def) +apply(rule_tac x = "q + 2*n + 6" in exI, simp) +done + +end + + + + + + + diff -r cbb4ac6c8081 -r 1ce04eb1c8ad utm/turing_basic.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/utm/turing_basic.thy Sat Sep 29 12:38:12 2012 +0000 @@ -0,0 +1,747 @@ +theory turing_basic +imports Main +begin + +section {* Basic definitions of Turing machine *} + +(* Title: Turing machine's definition and its charater + Author: Xu Jian + Maintainer: Xu Jian +*) + +text {* +\label{description of turing machine} +*} + +section {* Basic definitions of Turing machine *} + +(* Title: Turing machine's definition and its charater + Author: Xu Jian + Maintainer: Xu Jian +*) + +text {* + Actions of Turing machine (Abbreviated TM in the following* ). +*} + +datatype taction = + -- {* Write zero *} + W0 | + -- {* Write one *} + W1 | + -- {* Move left *} + L | + -- {* Move right *} + R | + -- {* Do nothing *} + Nop + +text {* + Tape contents in every block. +*} + +datatype block = + -- {* Blank *} + Bk | + -- {* Occupied *} + Oc + +text {* + Tape is represented as a pair of lists $(L_{left}, L_{right})$, + where $L_left$, named {\em left list}, is used to represent + the tape to the left of RW-head and + $L_{right}$, named {\em right list}, is used to represent the tape + under and to the right of RW-head. +*} + +type_synonym tape = "block list \ block list" + +text {* The state of turing machine.*} +type_synonym tstate = nat + +text {* + Turing machine instruction is represented as a + pair @{text "(action, next_state)"}, + where @{text "action"} is the action to take at the current state + and @{text "next_state"} is the next state the machine is getting into + after the action. +*} +type_synonym tinst = "taction \ tstate" + +text {* + Program of Turing machine is represented as a list of Turing instructions + and the execution of the program starts from the head of the list. + *} +type_synonym tprog = "tinst list" + + +text {* + Turing machine configuration, which consists of the current state + and the tape. +*} +type_synonym t_conf = "tstate \ tape" + +fun nth_of :: "'a list \ nat \ 'a option" + where + "nth_of xs n = (if n < length xs then Some (xs!n) + else None)" + +text {* + The function used to fetech instruction out of Turing program. + *} + +fun fetch :: "tprog \ tstate \ block \ tinst" + where + "fetch p s b = (if s = 0 then (Nop, 0) else + case b of + Bk \ case nth_of p (2 * (s - 1)) of + Some i \ i + | None \ (Nop, 0) + | Oc \ case nth_of p (2 * (s - 1) +1) of + Some i \ i + | None \ (Nop, 0))" + + +fun new_tape :: "taction \ tape \ tape" +where + "new_tape action (leftn, rightn) = (case action of + W0 \ (leftn, Bk#(tl rightn)) | + W1 \ (leftn, Oc#(tl rightn)) | + L \ (if leftn = [] then (tl leftn, Bk#rightn) + else (tl leftn, (hd leftn) # rightn)) | + R \ if rightn = [] then (Bk#leftn,tl rightn) + else ((hd rightn)#leftn, tl rightn) | + Nop \ (leftn, rightn) + )" + +text {* + The one step function used to transfer Turing machine configuration. +*} +fun tstep :: "t_conf \ tprog \ t_conf" + where + "tstep c p = (let (s, l, r) = c in + let (ac, ns) = (fetch p s (case r of [] \ Bk | + x # xs \ x)) in + (ns, new_tape ac (l, r)))" + +text {* + The many-step function. +*} +fun steps :: "t_conf \ tprog \ nat \ t_conf" + where + "steps c p 0 = c" | + "steps c p (Suc n) = steps (tstep c p) p n" + +lemma tstep_red: "steps c p (Suc n) = tstep (steps c p n) p" +proof(induct n arbitrary: c) + fix c + show "steps c p (Suc 0) = tstep (steps c p 0) p" by(simp add: steps.simps) +next + fix n c + assume ind: "\ c. steps c p (Suc n) = tstep (steps c p n) p" + have "steps (tstep c p) p (Suc n) = tstep (steps (tstep c p) p n) p" + by(rule ind) + thus "steps c p (Suc (Suc n)) = tstep (steps c p (Suc n)) p" by(simp add: steps.simps) +qed + +declare Let_def[simp] option.split[split] + +definition + "iseven n \ \ x. n = 2 * x" + + +text {* + The following @{text "t_correct"} function is used to specify the wellformedness of Turing + machine. +*} +fun t_correct :: "tprog \ bool" + where + "t_correct p = (length p \ 2 \ iseven (length p) \ + list_all (\ (acn, s). s \ length p div 2) p)" + +declare t_correct.simps[simp del] + +lemma allimp: "\\x. P x \ Q x; \x. P x\ \ \x. Q x" +by(auto elim: allE) + +lemma halt_lemma: "\wf LE; \ n. (\ P (f n) \ (f (Suc n), (f n)) \ LE)\ \ \ n. P (f n)" +apply(rule exCI, drule allimp, auto) +apply(drule_tac f = f in wf_inv_image, simp add: inv_image_def) +apply(erule wf_induct, auto) +done + +lemma steps_add: "steps c t (x + y) = steps (steps c t x) t y" +by(induct x arbitrary: c, auto simp: steps.simps tstep_red) + +lemma listall_set: "list_all p t \ \ a \ set t. p a" +by(induct t, auto) + +lemma fetch_ex: "\b a. fetch T aa ab = (b, a)" +by(simp add: fetch.simps) +definition exponent :: "'a \ nat \ 'a list" ("_\<^bsup>_\<^esup>" [0, 0]100) + where "exponent x n = replicate n x" + +text {* + @{text "tinres l1 l2"} means left list @{text "l1"} is congruent with + @{text "l2"} with respect to the execution of Turing machine. + Appending Blank to the right of eigther one does not affect the + outcome of excution. +*} + +definition tinres :: "block list \ block list \ bool" + where + "tinres bx by = (\ n. bx = by@Bk\<^bsup>n\<^esup> \ by = bx @ Bk\<^bsup>n\<^esup>)" + +lemma exp_zero: "a\<^bsup>0\<^esup> = []" +by(simp add: exponent_def) +lemma exp_ind_def: "a\<^bsup>Suc x \<^esup> = a # a\<^bsup>x\<^esup>" +by(simp add: exponent_def) + +text {* + The following lemma shows the meaning of @{text "tinres"} with respect to + one step execution. + *} +lemma tinres_step: + "\tinres l l'; tstep (ss, l, r) t = (sa, la, ra); tstep (ss, l', r) t = (sb, lb, rb)\ + \ tinres la lb \ ra = rb \ sa = sb" +apply(auto simp: tstep.simps fetch.simps new_tape.simps + split: if_splits taction.splits list.splits + block.splits) +apply(case_tac [!] "t ! (2 * (ss - Suc 0))", + auto simp: exponent_def tinres_def split: if_splits taction.splits list.splits + block.splits) +apply(case_tac [!] "t ! (2 * (ss - Suc 0) + Suc 0)", + auto simp: exponent_def tinres_def split: if_splits taction.splits list.splits + block.splits) +done + +declare tstep.simps[simp del] steps.simps[simp del] + +text {* + The following lemma shows the meaning of @{text "tinres"} with respect to + many step execution. + *} +lemma tinres_steps: + "\tinres l l'; steps (ss, l, r) t stp = (sa, la, ra); steps (ss, l', r) t stp = (sb, lb, rb)\ + \ tinres la lb \ ra = rb \ sa = sb" +apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps) +apply(simp add: tstep_red) +apply(case_tac "(steps (ss, l, r) t stp)") +apply(case_tac "(steps (ss, l', r) t stp)") +proof - + fix stp sa la ra sb lb rb a b c aa ba ca + assume ind: "\sa la ra sb lb rb. \steps (ss, l, r) t stp = (sa, la, ra); + steps (ss, l', r) t stp = (sb, lb, rb)\ \ tinres la lb \ ra = rb \ sa = sb" + and h: " tinres l l'" "tstep (steps (ss, l, r) t stp) t = (sa, la, ra)" + "tstep (steps (ss, l', r) t stp) t = (sb, lb, rb)" "steps (ss, l, r) t stp = (a, b, c)" + "steps (ss, l', r) t stp = (aa, ba, ca)" + have "tinres b ba \ c = ca \ a = aa" + apply(rule_tac ind, simp_all add: h) + done + thus "tinres la lb \ ra = rb \ sa = sb" + apply(rule_tac l = b and l' = ba and r = c and ss = a + and t = t in tinres_step) + using h + apply(simp, simp, simp) + done +qed + +text {* + The following function @{text "tshift tp n"} is used to shift Turing programs + @{text "tp"} by @{text "n"} when it is going to be combined with others. + *} + +fun tshift :: "tprog \ nat \ tprog" + where + "tshift tp off = (map (\ (action, state). (action, (if state = 0 then 0 + else state + off))) tp)" + +text {* + When two Turing programs are combined, the end state (state @{text "0"}) of the one + at the prefix position needs to be connected to the start state + of the one at postfix position. If @{text "tp"} is the Turing program + to be at the prefix, @{text "change_termi_state tp"} is the transformed Turing program. + *} +fun change_termi_state :: "tprog \ tprog" + where + "change_termi_state t = + (map (\ (acn, ns). if ns = 0 then (acn, Suc ((length t) div 2)) else (acn, ns)) t)" + +text {* + @{text "t_add tp1 tp2"} is the combined Truing program. +*} + +fun t_add :: "tprog \ tprog \ tprog" ("_ |+| _" [0, 0] 100) + where + "t_add t1 t2 = ((change_termi_state t1) @ (tshift t2 ((length t1) div 2)))" + +text {* + Tests whether the current configuration is at state @{text "0"}. +*} +definition isS0 :: "t_conf \ bool" + where + "isS0 c = (let (s, l, r) = c in s = 0)" + +declare tstep.simps[simp del] steps.simps[simp del] + t_add.simps[simp del] fetch.simps[simp del] + new_tape.simps[simp del] + + +text {* + Single step execution starting from state @{text "0"} will not make any progress. +*} +lemma tstep_0: "tstep (0, tp) p = (0, tp)" +apply(simp add: tstep.simps fetch.simps new_tape.simps) +done + + +text {* + Many step executions starting from state @{text "0"} will not make any progress. +*} + +lemma steps_0: "steps (0, tp) p stp = (0, tp)" +apply(induct stp) +apply(simp add: steps.simps) +apply(simp add: tstep_red tstep_0) +done + +lemma s_keep_step: "\a \ length A div 2; tstep (a, b, c) A = (s, l, r); t_correct A\ + \ s \ length A div 2" +apply(simp add: tstep.simps fetch.simps t_correct.simps iseven_def + split: if_splits block.splits list.splits) +apply(case_tac [!] a, auto simp: list_all_length) +apply(erule_tac x = "2 * nat" in allE, auto) +apply(erule_tac x = "2 * nat" in allE, auto) +apply(erule_tac x = "Suc (2 * nat)" in allE, auto) +done + +lemma s_keep: "\steps (Suc 0, tp) A stp = (s, l, r); t_correct A\ \ s \ length A div 2" +proof(induct stp arbitrary: s l r) + case 0 thus "?case" by(auto simp: t_correct.simps steps.simps) +next + fix stp s l r + assume ind: "\s l r. \steps (Suc 0, tp) A stp = (s, l, r); t_correct A\ \ s \ length A div 2" + and h1: "steps (Suc 0, tp) A (Suc stp) = (s, l, r)" + and h2: "t_correct A" + from h1 h2 show "s \ length A div 2" + proof(simp add: tstep_red, cases "(steps (Suc 0, tp) A stp)", simp) + fix a b c + assume h3: "tstep (a, b, c) A = (s, l, r)" + and h4: "steps (Suc 0, tp) A stp = (a, b, c)" + have "a \ length A div 2" + using h2 h4 + by(rule_tac l = b and r = c in ind, auto) + thus "?thesis" + using h3 h2 + by(simp add: s_keep_step) + qed +qed + +lemma t_merge_fetch_pre: + "\fetch A s b = (ac, ns); s \ length A div 2; t_correct A; s \ 0\ \ + fetch (A |+| B) s b = (ac, if ns = 0 then Suc (length A div 2) else ns)" +apply(subgoal_tac "2 * (s - Suc 0) < length A \ Suc (2 * (s - Suc 0)) < length A") +apply(auto simp: fetch.simps t_add.simps split: if_splits block.splits) +apply(simp_all add: nth_append change_termi_state.simps) +done + +lemma [simp]: "\\ a \ length A div 2; t_correct A\ \ fetch A a b = (Nop, 0)" +apply(auto simp: fetch.simps del: nth_of.simps split: block.splits) +apply(case_tac [!] a, auto simp: t_correct.simps iseven_def) +done + +lemma [elim]: "\t_correct A; \ isS0 (tstep (a, b, c) A)\ \ a \ length A div 2" +apply(rule_tac classical, auto simp: tstep.simps new_tape.simps isS0_def) +done + +lemma [elim]: "\t_correct A; \ isS0 (tstep (a, b, c) A)\ \ 0 < a" +apply(rule_tac classical, simp add: tstep_0 isS0_def) +done + + +lemma t_merge_pre_eq_step: "\tstep (a, b, c) A = cf; t_correct A; \ isS0 cf\ + \ tstep (a, b, c) (A |+| B) = cf" +apply(subgoal_tac "a \ length A div 2 \ a \ 0") +apply(simp add: tstep.simps) +apply(case_tac "fetch A a (case c of [] \ Bk | x # xs \ x)", simp) +apply(drule_tac B = B in t_merge_fetch_pre, simp, simp, simp, simp add: isS0_def, auto) +done + +lemma t_merge_pre_eq: "\steps (Suc 0, tp) A stp = cf; \ isS0 cf; t_correct A\ + \ steps (Suc 0, tp) (A |+| B) stp = cf" +proof(induct stp arbitrary: cf) + case 0 thus "?case" by(simp add: steps.simps) +next + fix stp cf + assume ind: "\cf. \steps (Suc 0, tp) A stp = cf; \ isS0 cf; t_correct A\ + \ steps (Suc 0, tp) (A |+| B) stp = cf" + and h1: "steps (Suc 0, tp) A (Suc stp) = cf" + and h2: "\ isS0 cf" + and h3: "t_correct A" + from h1 h2 h3 show "steps (Suc 0, tp) (A |+| B) (Suc stp) = cf" + proof(simp add: tstep_red, cases "steps (Suc 0, tp) (A) stp", simp) + fix a b c + assume h4: "tstep (a, b, c) A = cf" + and h5: "steps (Suc 0, tp) A stp = (a, b, c)" + have "steps (Suc 0, tp) (A |+| B) stp = (a, b, c)" + proof(cases a) + case 0 thus "?thesis" + using h4 h2 + apply(simp add: tstep_0, cases cf, simp add: isS0_def) + done + next + case (Suc n) thus "?thesis" + using h5 h3 + apply(rule_tac ind, auto simp: isS0_def) + done + qed + thus "tstep (steps (Suc 0, tp) (A |+| B) stp) (A |+| B) = cf" + using h4 h5 h3 h2 + apply(simp) + apply(rule t_merge_pre_eq_step, auto) + done + qed +qed + +declare nth.simps[simp del] tshift.simps[simp del] change_termi_state.simps[simp del] + +lemma [simp]: "length (change_termi_state A) = length A" +by(simp add: change_termi_state.simps) + +lemma first_halt_point: "steps (Suc 0, tp) A stp = (0, tp') + \ \stp. \ isS0 (steps (Suc 0, tp) A stp) \ steps (Suc 0, tp) A (Suc stp) = (0, tp')" +proof(induct stp) + case 0 thus "?case" by(simp add: steps.simps) +next + case (Suc n) + fix stp + assume ind: "steps (Suc 0, tp) A stp = (0, tp') \ + \stp. \ isS0 (steps (Suc 0, tp) A stp) \ steps (Suc 0, tp) A (Suc stp) = (0, tp')" + and h: "steps (Suc 0, tp) A (Suc stp) = (0, tp')" + from h show "\stp. \ isS0 (steps (Suc 0, tp) A stp) \ steps (Suc 0, tp) A (Suc stp) = (0, tp')" + proof(simp add: tstep_red, cases "steps (Suc 0, tp) A stp", simp, case_tac a) + fix a b c + assume g1: "a = (0::nat)" + and g2: "tstep (a, b, c) A = (0, tp')" + and g3: "steps (Suc 0, tp) A stp = (a, b, c)" + have "steps (Suc 0, tp) A stp = (0, tp')" + using g2 g1 g3 + by(simp add: tstep_0) + hence "\ stp. \ isS0 (steps (Suc 0, tp) A stp) \ steps (Suc 0, tp) A (Suc stp) = (0, tp')" + by(rule ind) + thus "\stp. \ isS0 (steps (Suc 0, tp) A stp) \ tstep (steps (Suc 0, tp) A stp) A = (0, tp')" + apply(simp add: tstep_red) + done + next + fix a b c nat + assume g1: "steps (Suc 0, tp) A stp = (a, b, c)" + and g2: "steps (Suc 0, tp) A (Suc stp) = (0, tp')" "a= Suc nat" + thus "\stp. \ isS0 (steps (Suc 0, tp) A stp) \ tstep (steps (Suc 0, tp) A stp) A = (0, tp')" + apply(rule_tac x = stp in exI) + apply(simp add: isS0_def tstep_red) + done + qed +qed + +lemma t_merge_pre_halt_same': + "\\ isS0 (steps (Suc 0, tp) A stp) ; steps (Suc 0, tp) A (Suc stp) = (0, tp'); t_correct A\ + \ steps (Suc 0, tp) (A |+| B) (Suc stp) = (Suc (length A div 2), tp')" +proof(simp add: tstep_red, cases "steps (Suc 0, tp) A stp", simp) + fix a b c + assume h1: "\ isS0 (a, b, c)" + and h2: "tstep (a, b, c) A = (0, tp')" + and h3: "t_correct A" + and h4: "steps (Suc 0, tp) A stp = (a, b, c)" + have "steps (Suc 0, tp) (A |+| B) stp = (a, b, c)" + using h1 h4 h3 + apply(rule_tac t_merge_pre_eq, auto) + done + moreover have "tstep (a, b, c) (A |+| B) = (Suc (length A div 2), tp')" + using h2 h3 h1 h4 + apply(simp add: tstep.simps) + apply(case_tac " fetch A a (case c of [] \ Bk | x # xs \ x)", simp) + apply(drule_tac B = B in t_merge_fetch_pre, auto simp: isS0_def intro: s_keep) + done + ultimately show "tstep (steps (Suc 0, tp) (A |+| B) stp) (A |+| B) = (Suc (length A div 2), tp')" + by(simp) +qed + +text {* + When Turing machine @{text "A"} and @{text "B"} are combined and the execution + of @{text "A"} can termination within @{text "stp"} steps, + the combined machine @{text "A |+| B"} will eventually get into the starting + state of machine @{text "B"}. +*} +lemma t_merge_pre_halt_same: " + \steps (Suc 0, tp) A stp = (0, tp'); t_correct A; t_correct B\ + \ \ stp. steps (Suc 0, tp) (A |+| B) stp = (Suc (length A div 2), tp')" +proof - + assume a_wf: "t_correct A" + and b_wf: "t_correct B" + and a_ht: "steps (Suc 0, tp) A stp = (0, tp')" + have halt_point: "\ stp. \ isS0 (steps (Suc 0, tp) A stp) \ steps (Suc 0, tp) A (Suc stp) = (0, tp')" + using a_ht + by(erule_tac first_halt_point) + then obtain stp' where "\ isS0 (steps (Suc 0, tp) A stp') \ steps (Suc 0, tp) A (Suc stp') = (0, tp')".. + hence "steps (Suc 0, tp) (A |+| B) (Suc stp') = (Suc (length A div 2), tp')" + using a_wf + apply(rule_tac t_merge_pre_halt_same', auto) + done + thus "?thesis" .. +qed + +lemma fetch_0: "fetch p 0 b = (Nop, 0)" +by(simp add: fetch.simps) + +lemma [simp]: "length (tshift B x) = length B" +by(simp add: tshift.simps) + +lemma [simp]: "t_correct A \ 2 * (length A div 2) = length A" +apply(simp add: t_correct.simps iseven_def, auto) +done + +lemma t_merge_fetch_snd: + "\fetch B a b = (ac, ns); t_correct A; t_correct B; a > 0 \ + \ fetch (A |+| B) (a + length A div 2) b + = (ac, if ns = 0 then 0 else ns + length A div 2)" +apply(auto simp: fetch.simps t_add.simps split: if_splits block.splits) +apply(case_tac [!] a, simp_all) +apply(simp_all add: nth_append change_termi_state.simps tshift.simps) +done + +lemma t_merge_snd_eq_step: + "\tstep (s, l, r) B = (s', l', r'); t_correct A; t_correct B; s > 0\ + \ tstep (s + length A div 2, l, r) (A |+| B) = + (if s' = 0 then 0 else s' + length A div 2, l' ,r') " +apply(simp add: tstep.simps) +apply(cases "fetch B s (case r of [] \ Bk | x # xs \ x)") +apply(auto simp: t_merge_fetch_snd) +apply(frule_tac [!] t_merge_fetch_snd, auto) +done + +text {* + Relates the executions of TM @{text "B"}, one is when @{text "B"} is executed alone, + the other is the execution when @{text "B"} is in the combined TM. +*} +lemma t_merge_snd_eq_steps: + "\t_correct A; t_correct B; steps (s, l, r) B stp = (s', l', r'); s > 0\ + \ steps (s + length A div 2, l, r) (A |+| B) stp = + (if s' = 0 then 0 else s' + length A div 2, l', r')" +proof(induct stp arbitrary: s' l' r') + case 0 thus "?case" + by(simp add: steps.simps) +next + fix stp s' l' r' + assume ind: "\s' l' r'. \t_correct A; t_correct B; steps (s, l, r) B stp = (s', l', r'); 0 < s\ + \ steps (s + length A div 2, l, r) (A |+| B) stp = + (if s' = 0 then 0 else s' + length A div 2, l', r')" + and h1: "steps (s, l, r) B (Suc stp) = (s', l', r')" + and h2: "t_correct A" + and h3: "t_correct B" + and h4: "0 < s" + from h1 show "steps (s + length A div 2, l, r) (A |+| B) (Suc stp) + = (if s' = 0 then 0 else s' + length A div 2, l', r')" + proof(simp only: tstep_red, cases "steps (s, l, r) B stp") + fix a b c + assume h5: "steps (s, l, r) B stp = (a, b, c)" "tstep (steps (s, l, r) B stp) B = (s', l', r')" + hence h6: "(steps (s + length A div 2, l, r) (A |+| B) stp) = + ((if a = 0 then 0 else a + length A div 2, b, c))" + using h2 h3 h4 + by(rule_tac ind, auto) + thus "tstep (steps (s + length A div 2, l, r) (A |+| B) stp) (A |+| B) = + (if s' = 0 then 0 else s'+ length A div 2, l', r')" + using h5 + proof(auto) + assume "tstep (0, b, c) B = (0, l', r')" thus "tstep (0, b, c) (A |+| B) = (0, l', r')" + by(simp add: tstep_0) + next + assume "tstep (0, b, c) B = (s', l', r')" "0 < s'" + thus "tstep (0, b, c) (A |+| B) = (s' + length A div 2, l', r')" + by(simp add: tstep_0) + next + assume "tstep (a, b, c) B = (0, l', r')" "0 < a" + thus "tstep (a + length A div 2, b, c) (A |+| B) = (0, l', r')" + using h2 h3 + by(drule_tac t_merge_snd_eq_step, auto) + next + assume "tstep (a, b, c) B = (s', l', r')" "0 < a" "0 < s'" + thus "tstep (a + length A div 2, b, c) (A |+| B) = (s' + length A div 2, l', r')" + using h2 h3 + by(drule_tac t_merge_snd_eq_step, auto) + qed + qed +qed + +lemma t_merge_snd_halt_eq: + "\steps (Suc 0, tp) B stp = (0, tp'); t_correct A; t_correct B\ + \ \stp. steps (Suc (length A div 2), tp) (A |+| B) stp = (0, tp')" +apply(case_tac tp, cases tp', simp) +apply(drule_tac s = "Suc 0" in t_merge_snd_eq_steps, auto) +done + +lemma t_inj: "\steps (Suc 0, tp) A stpa = (0, tp1); steps (Suc 0, tp) A stpb = (0, tp2)\ + \ tp1 = tp2" +proof - + assume h1: "steps (Suc 0, tp) A stpa = (0, tp1)" + and h2: "steps (Suc 0, tp) A stpb = (0, tp2)" + thus "?thesis" + proof(cases "stpa < stpb") + case True thus "?thesis" + using h1 h2 + apply(drule_tac less_imp_Suc_add, auto) + apply(simp del: add_Suc_right add_Suc add: add_Suc_right[THEN sym] steps_add steps_0) + done + next + case False thus "?thesis" + using h1 h2 + apply(drule_tac leI) + apply(case_tac "stpb = stpa", auto) + apply(subgoal_tac "stpb < stpa") + apply(drule_tac less_imp_Suc_add, auto) + apply(simp del: add_Suc_right add_Suc add: add_Suc_right[THEN sym] steps_add steps_0) + done + qed +qed + +type_synonym t_assert = "tape \ bool" + +definition t_imply :: "t_assert \ t_assert \ bool" ("_ \-> _" [0, 0] 100) + where + "t_imply a1 a2 = (\ tp. a1 tp \ a2 tp)" + + +locale turing_merge = + fixes A :: "tprog" and B :: "tprog" and P1 :: "t_assert" + and P2 :: "t_assert" + and P3 :: "t_assert" + and P4 :: "t_assert" + and Q1:: "t_assert" + and Q2 :: "t_assert" + assumes + A_wf : "t_correct A" + and B_wf : "t_correct B" + and A_halt : "P1 tp \ \ stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \ Q1 tp'" + and B_halt : "P2 tp \ \ stp. let (s, tp') = steps (Suc 0, tp) B stp in s = 0 \ Q2 tp'" + and A_uhalt : "P3 tp \ \ (\ stp. isS0 (steps (Suc 0, tp) A stp))" + and B_uhalt: "P4 tp \ \ (\ stp. isS0 (steps (Suc 0, tp) B stp))" +begin + + +text {* + The following lemma tries to derive the Hoare logic rule for sequentially combined TMs. + It deals with the situtation when both @{text "A"} and @{text "B"} are terminated. +*} + +lemma t_merge_halt: + assumes aimpb: "Q1 \-> P2" + shows "P1 \-> \ tp. (\ stp tp'. steps (Suc 0, tp) (A |+| B) stp = (0, tp') \ Q2 tp')" +proof(simp add: t_imply_def, auto) + fix a b + assume h: "P1 (a, b)" + hence "\ stp. let (s, tp') = steps (Suc 0, a, b) A stp in s = 0 \ Q1 tp'" + using A_halt by simp + from this obtain stp1 where "let (s, tp') = steps (Suc 0, a, b) A stp1 in s = 0 \ Q1 tp'" .. + thus "\stp aa ba. steps (Suc 0, a, b) (A |+| B) stp = (0, aa, ba) \ Q2 (aa, ba)" + proof(case_tac "steps (Suc 0, a, b) A stp1", simp, erule_tac conjE) + fix aa ba c + assume g1: "Q1 (ba, c)" + and g2: "steps (Suc 0, a, b) A stp1 = (0, ba, c)" + hence "P2 (ba, c)" + using aimpb apply(simp add: t_imply_def) + done + hence "\ stp. let (s, tp') = steps (Suc 0, ba, c) B stp in s = 0 \ Q2 tp'" + using B_halt by simp + from this obtain stp2 where "let (s, tp') = steps (Suc 0, ba, c) B stp2 in s = 0 \ Q2 tp'" .. + thus "?thesis" + proof(case_tac "steps (Suc 0, ba, c) B stp2", simp, erule_tac conjE) + fix aa bb ca + assume g3: " Q2 (bb, ca)" "steps (Suc 0, ba, c) B stp2 = (0, bb, ca)" + have "\ stp. steps (Suc 0, a, b) (A |+| B) stp = (Suc (length A div 2), ba , c)" + using g2 A_wf B_wf + by(rule_tac t_merge_pre_halt_same, auto) + moreover have "\ stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)" + using g3 A_wf B_wf + apply(rule_tac t_merge_snd_halt_eq, auto) + done + ultimately show "\stp aa ba. steps (Suc 0, a, b) (A |+| B) stp = (0, aa, ba) \ Q2 (aa, ba)" + apply(erule_tac exE, erule_tac exE) + apply(rule_tac x = "stp + stpa" in exI, simp add: steps_add) + using g3 by simp + qed + qed +qed + +lemma t_merge_uhalt_tmp: + assumes B_uh: "\stp. \ isS0 (steps (Suc 0, b, c) B stp)" + and merge_ah: "steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" + shows "\ stp. \ isS0 (steps (Suc 0, tp) (A |+| B) stp)" + using B_uh merge_ah +apply(rule_tac allI) +apply(case_tac "stp > stpa") +apply(erule_tac x = "stp - stpa" in allE) +apply(case_tac "(steps (Suc 0, b, c) B (stp - stpa))", simp) +proof - + fix stp a ba ca + assume h1: "\ isS0 (a, ba, ca)" "stpa < stp" + and h2: "steps (Suc 0, b, c) B (stp - stpa) = (a, ba, ca)" + have "steps (Suc 0 + length A div 2, b, c) (A |+| B) (stp - stpa) = + (if a = 0 then 0 else a + length A div 2, ba, ca)" + using A_wf B_wf h2 + by(rule_tac t_merge_snd_eq_steps, auto) + moreover have "a > 0" using h1 by(simp add: isS0_def) + moreover have "\ stpb. stp = stpa + stpb" + using h1 by(rule_tac x = "stp - stpa" in exI, simp) + ultimately show "\ isS0 (steps (Suc 0, tp) (A |+| B) stp)" + using merge_ah + by(auto simp: steps_add isS0_def) +next + fix stp + assume h: "steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" "\ stpa < stp" + hence "\ stpb. stpa = stp + stpb" apply(rule_tac x = "stpa - stp" in exI, auto) done + thus "\ isS0 (steps (Suc 0, tp) (A |+| B) stp)" + using h + apply(auto) + apply(cases "steps (Suc 0, tp) (A |+| B) stp", simp add: steps_add isS0_def steps_0) + done +qed + +text {* + The following lemma deals with the situation when TM @{text "B"} can not terminate. + *} + +lemma t_merge_uhalt: + assumes aimpb: "Q1 \-> P4" + shows "P1 \-> \ tp. \ (\ stp. isS0 (steps (Suc 0, tp) (A |+| B) stp))" +proof(simp only: t_imply_def, rule_tac allI, rule_tac impI) + fix tp + assume init_asst: "P1 tp" + show "\ (\stp. isS0 (steps (Suc 0, tp) (A |+| B) stp))" + proof - + have "\ stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \ Q1 tp'" + using A_halt[of tp] init_asst + by(simp) + from this obtain stpx where "let (s, tp') = steps (Suc 0, tp) A stpx in s = 0 \ Q1 tp'" .. + thus "?thesis" + proof(cases "steps (Suc 0, tp) A stpx", simp, erule_tac conjE) + fix a b c + assume "Q1 (b, c)" + and h3: "steps (Suc 0, tp) A stpx = (0, b, c)" + hence h2: "P4 (b, c)" using aimpb + by(simp add: t_imply_def) + have "\ stp. steps (Suc 0, tp) (A |+| B) stp = (Suc (length A div 2), b, c)" + using h3 A_wf B_wf + apply(rule_tac stp = stpx in t_merge_pre_halt_same, auto) + done + from this obtain stpa where h4:"steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" .. + have " \ (\ stp. isS0 (steps (Suc 0, b, c) B stp))" + using B_uhalt [of "(b, c)"] h2 apply simp + done + from this and h4 show "\stp. \ isS0 (steps (Suc 0, tp) (A |+| B) stp)" + by(rule_tac t_merge_uhalt_tmp, auto) + qed + qed +qed +end + +end + diff -r cbb4ac6c8081 -r 1ce04eb1c8ad utm/uncomputable.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/utm/uncomputable.thy Sat Sep 29 12:38:12 2012 +0000 @@ -0,0 +1,1693 @@ +(* Title: Turing machine's definition and its charater + Author: XuJian + Maintainer: Xujian +*) + +header {* Undeciablity of the {\em Halting problem} *} + +theory uncomputable +imports Main turing_basic +begin + +text {* + The {\em Copying} TM, which duplicates its input. +*} +definition tcopy :: "tprog" +where +"tcopy \ [(W0, 0), (R, 2), (R, 3), (R, 2), + (W1, 3), (L, 4), (L, 4), (L, 5), (R, 11), (R, 6), + (R, 7), (W0, 6), (R, 7), (R, 8), (W1, 9), (R, 8), + (L, 10), (L, 9), (L, 10), (L, 5), (R, 12), (R, 12), + (W1, 13), (L, 14), (R, 12), (R, 12), (L, 15), (W0, 14), + (R, 0), (L, 15)]" + +text {* + @{text "wipeLastBs tp"} removes all blanks at the end of tape @{text "tp"}. +*} +fun wipeLastBs :: "block list \ block list" + where + "wipeLastBs bl = rev (dropWhile (\a. a = Bk) (rev bl))" + +fun isBk :: "block \ bool" + where + "isBk b = (b = Bk)" + +text {* + The following functions are used to expression invariants of {\em Copying} TM. +*} +fun tcopy_F0 :: "nat \ tape \ bool" + where + "tcopy_F0 x tp = (let (ln, rn) = tp in + list_all isBk ln & rn = replicate x Oc + @ [Bk] @ replicate x Oc)" + +fun tcopy_F1 :: "nat \ tape \ bool" + where + "tcopy_F1 x (ln, rn) = (ln = [] & rn = replicate x Oc)" + +fun tcopy_F2 :: "nat \ tape \ bool" + where + "tcopy_F2 0 tp = False" | + "tcopy_F2 (Suc x) (ln, rn) = (length ln > 0 & + ln @ rn = replicate (Suc x) Oc)" + +fun tcopy_F3 :: "nat \ tape \ bool" + where + "tcopy_F3 0 tp = False" | + "tcopy_F3 (Suc x) (ln, rn) = + (ln = Bk # replicate (Suc x) Oc & length rn <= 1)" + +fun tcopy_F4 :: "nat \ tape \ bool" + where + "tcopy_F4 0 tp = False" | + "tcopy_F4 (Suc x) (ln, rn) = + ((ln = replicate x Oc & rn = [Oc, Bk, Oc]) + | (ln = replicate (Suc x) Oc & rn = [Bk, Oc])) " + +fun tcopy_F5 :: "nat \ tape \ bool" + where + "tcopy_F5 0 tp = False" | + "tcopy_F5 (Suc x) (ln, rn) = + (if rn = [] then False + else if hd rn = Bk then (ln = [] & + rn = Bk # (Oc # replicate (Suc x) Bk + @ replicate (Suc x) Oc)) + else if hd rn = Oc then + (\n. ln = replicate (x - n) Oc + & rn = Oc # (Oc # replicate n Bk @ replicate n Oc) + & n > 0 & n <= x) + else False)" + + +fun tcopy_F6 :: "nat \ tape \ bool" + where + "tcopy_F6 0 tp = False" | + "tcopy_F6 (Suc x) (ln, rn) = + (\n. ln = replicate (Suc x -n) Oc + & tl rn = replicate n Bk @ replicate n Oc + & n > 0 & n <= x)" + +fun tcopy_F7 :: "nat \ tape \ bool" + where + "tcopy_F7 0 tp = False" | + "tcopy_F7 (Suc x) (ln, rn) = + (let lrn = (rev ln) @ rn in + (\n. lrn = replicate ((Suc x) - n) Oc @ + replicate (Suc n) Bk @ replicate n Oc + & n > 0 & n <= x & + length rn >= n & length rn <= 2 * n ))" + +fun tcopy_F8 :: "nat \ tape \ bool" + where + "tcopy_F8 0 tp = False" | + "tcopy_F8 (Suc x) (ln, rn) = + (let lrn = (rev ln) @ rn in + (\n. lrn = replicate ((Suc x) - n) Oc @ + replicate (Suc n) Bk @ replicate n Oc + & n > 0 & n <= x & length rn < n)) " + +fun tcopy_F9 :: "nat \ tape \ bool" + where + "tcopy_F9 0 tp = False" | + "tcopy_F9 (Suc x) (ln, rn) = + (let lrn = (rev ln) @ rn in + (\n. lrn = replicate (Suc (Suc x) - n) Oc + @ replicate n Bk @ replicate n Oc + & n > Suc 0 & n <= Suc x & length rn > 0 + & length rn <= Suc n))" + +fun tcopy_F10 :: "nat \ tape \ bool" + where + "tcopy_F10 0 tp = False" | + "tcopy_F10 (Suc x) (ln, rn) = + (let lrn = (rev ln) @ rn in + (\n. lrn = replicate (Suc (Suc x) - n) Oc + @ replicate n Bk @ replicate n Oc & n > Suc 0 + & n <= Suc x & length rn > Suc n & + length rn <= 2*n + 1 ))" + +fun tcopy_F11 :: "nat \ tape \ bool" + where + "tcopy_F11 0 tp = False" | + "tcopy_F11 (Suc x) (ln, rn) = + (ln = [Bk] & rn = Oc # replicate (Suc x) Bk + @ replicate (Suc x) Oc)" + +fun tcopy_F12 :: "nat \ tape \ bool" + where + "tcopy_F12 0 tp = False" | + "tcopy_F12 (Suc x) (ln, rn) = + (let lrn = ((rev ln) @ rn) in + (\n. n > 0 & n <= Suc (Suc x) + & lrn = Bk # replicate n Oc @ replicate (Suc (Suc x) - n) Bk + @ replicate (Suc x) Oc + & length ln = Suc n))" + +fun tcopy_F13 :: "nat \ tape \ bool" + where + "tcopy_F13 0 tp = False" | + "tcopy_F13 (Suc x) (ln, rn) = + (let lrn = ((rev ln) @ rn) in + (\n. n > Suc 0 & n <= Suc (Suc x) + & lrn = Bk # replicate n Oc @ replicate (Suc (Suc x) - n) Bk + @ replicate (Suc x) Oc + & length ln = n))" + +fun tcopy_F14 :: "nat \ tape \ bool" + where + "tcopy_F14 0 tp = False" | + "tcopy_F14 (Suc x) (ln, rn) = + (ln = replicate (Suc x) Oc @ [Bk] & + tl rn = replicate (Suc x) Oc)" + +fun tcopy_F15 :: "nat \ tape \ bool" + where + "tcopy_F15 0 tp = False" | + "tcopy_F15 (Suc x) (ln, rn) = + (let lrn = ((rev ln) @ rn) in + lrn = Bk # replicate (Suc x) Oc @ [Bk] @ + replicate (Suc x) Oc & length ln <= (Suc x))" + +text {* + The following @{text "inv_tcopy"} is the invariant of the {\em Copying} TM. +*} +fun inv_tcopy :: "nat \ t_conf \ bool" + where + "inv_tcopy x c = (let (state, tp) = c in + if state = 0 then tcopy_F0 x tp + else if state = 1 then tcopy_F1 x tp + else if state = 2 then tcopy_F2 x tp + else if state = 3 then tcopy_F3 x tp + else if state = 4 then tcopy_F4 x tp + else if state = 5 then tcopy_F5 x tp + else if state = 6 then tcopy_F6 x tp + else if state = 7 then tcopy_F7 x tp + else if state = 8 then tcopy_F8 x tp + else if state = 9 then tcopy_F9 x tp + else if state = 10 then tcopy_F10 x tp + else if state = 11 then tcopy_F11 x tp + else if state = 12 then tcopy_F12 x tp + else if state = 13 then tcopy_F13 x tp + else if state = 14 then tcopy_F14 x tp + else if state = 15 then tcopy_F15 x tp + else False)" +declare tcopy_F0.simps [simp del] + tcopy_F1.simps [simp del] + tcopy_F2.simps [simp del] + tcopy_F3.simps [simp del] + tcopy_F4.simps [simp del] + tcopy_F5.simps [simp del] + tcopy_F6.simps [simp del] + tcopy_F7.simps [simp del] + tcopy_F8.simps [simp del] + tcopy_F9.simps [simp del] + tcopy_F10.simps [simp del] + tcopy_F11.simps [simp del] + tcopy_F12.simps [simp del] + tcopy_F13.simps [simp del] + tcopy_F14.simps [simp del] + tcopy_F15.simps [simp del] + +lemma list_replicate_Bk[dest]: "list_all isBk list \ + list = replicate (length list) Bk" +apply(induct list) +apply(simp+) +done + +lemma [simp]: "dropWhile (\ a. a = b) (replicate x b @ ys) = + dropWhile (\ a. a = b) ys" +apply(induct x) +apply(simp) +apply(simp) +done + +lemma [elim]: "\tstep (0, a, b) tcopy = (s, l, r); s \ 0\ \ RR" +apply(simp add: tstep.simps tcopy_def fetch.simps) +done + +lemma [elim]: "\tstep (Suc 0, a, b) tcopy = (s, l, r); s \ 0; s \ 2\ + \ RR" +apply(simp add: tstep.simps tcopy_def fetch.simps) +apply(simp split: block.splits list.splits) +done + +lemma [elim]: "\tstep (2, a, b) tcopy = (s, l, r); s \ 2; s \ 3\ + \ RR" +apply(simp add: tstep.simps tcopy_def fetch.simps) +apply(simp split: block.splits list.splits) +done + +lemma [elim]: "\tstep (3, a, b) tcopy = (s, l, r); s \ 3; s \ 4\ + \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (4, a, b) tcopy = (s, l, r); s \ 4; s \ 5\ + \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (5, a, b) tcopy = (s, l, r); s \ 6; s \ 11\ + \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (6, a, b) tcopy = (s, l, r); s \ 6; s \ 7\ + \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (7, a, b) tcopy = (s, l, r); s \ 7; s \ 8\ + \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (8, a, b) tcopy = (s, l, r); s \ 8; s \ 9\ + \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (9, a, b) tcopy = (s, l, r); s \ 9; s \ 10\ + \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (10, a, b) tcopy = (s, l, r); s \ 10; s \ 5\ + \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (11, a, b) tcopy = (s, l, r); s \ 12\ \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (12, a, b) tcopy = (s, l, r); s \ 13; s \ 14\ + \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (13, a, b) tcopy = (s, l, r); s \ 12\ \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (14, a, b) tcopy = (s, l, r); s \ 14; s \ 15\ + \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma [elim]: "\tstep (15, a, b) tcopy = (s, l, r); s \ 0; s \ 15\ + \ RR" +by(simp add: tstep.simps tcopy_def fetch.simps + split: block.splits list.splits) + +lemma min_Suc4: "min (Suc (Suc x)) x = x" +by auto + +lemma takeWhile2replicate: + "\n. takeWhile (\a. a = b) list = replicate n b" +apply(induct list) +apply(rule_tac x = 0 in exI, simp) +apply(auto) +apply(rule_tac x = "Suc n" in exI, simp) +done + +lemma rev_replicate_same: "rev (replicate x b) = replicate x b" +by(simp) + +lemma rev_equal: "a = b \ rev a = rev b" +by simp + +lemma rev_equal_rev: "rev a = rev b \ a = b" +by simp + +lemma rep_suc_rev[simp]:"replicate n b @ [b] = replicate (Suc n) b" +apply(rule rev_equal_rev) +apply(simp only: rev_append rev_replicate_same) +apply(auto) +done + +lemma replicate_Cons_simp: "b # replicate n b @ xs = + replicate n b @ b # xs" +apply(simp) +done + + +lemma [elim]: "\tstep (14, b, c) tcopy = (15, ab, ba); + tcopy_F14 x (b, c)\ \ tcopy_F15 x (ab, ba)" +apply(case_tac x) +apply(auto simp: tstep.simps tcopy_def + tcopy_F14.simps tcopy_F15.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +done + +lemma dropWhile_drophd: "\ p a \ + (dropWhile p xs @ (a # as)) = (dropWhile p (xs @ [a]) @ as)" +apply(induct xs) +apply(auto) +done + +lemma dropWhile_append3: "\\ p a; + listall ((dropWhile p xs) @ [a]) isBk\ \ + listall (dropWhile p (xs @ [a])) isBk" +apply(drule_tac p = p and xs = xs and a = a in dropWhile_drophd, simp) +done + +lemma takeWhile_append3: "\\p a; (takeWhile p xs) = b\ + \ takeWhile p (xs @ (a # as)) = b" +apply(drule_tac P = p and xs = xs and x = a and l = as in + takeWhile_tail) +apply(simp) +done + +lemma listall_append: "list_all p (xs @ ys) = + (list_all p xs \ list_all p ys)" +apply(induct xs) +apply(simp+) +done + +lemma [elim]: "\tstep (15, b, c) tcopy = (15, ab, ba); + tcopy_F15 x (b, c)\ \ tcopy_F15 x (ab, ba)" +apply(case_tac x) +apply(auto simp: tstep.simps tcopy_F15.simps + tcopy_def fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +apply(case_tac b, simp+) +done + +lemma [elim]: "\tstep (14, b, c) tcopy = (14, ab, ba); + tcopy_F14 x (b, c)\ \ tcopy_F14 x (ab, ba)" +apply(case_tac x) +apply(auto simp: tcopy_F14.simps tcopy_def tstep.simps + tcopy_F14.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +done + +lemma [intro]: "list_all isBk (replicate x Bk)" +apply(induct x, simp+) +done + +lemma [elim]: "list_all isBk (dropWhile (\a. a = Oc) b) \ + list_all isBk (dropWhile (\a. a = Oc) (tl b))" +apply(case_tac b, auto split: if_splits) +apply(drule list_replicate_Bk) +apply(case_tac "length list", auto) +done + +lemma [elim]: "list_all (\ a. a = Oc) list \ + list = replicate (length list) Oc" +apply(induct list) +apply(simp+) +done + +lemma append_length: "\as @ bs = cs @ ds; length bs = length ds\ + \ as = cs & bs = ds" +apply(auto) +done + +lemma Suc_elim: "Suc (Suc m) - n = Suc na \ Suc m - n = na" +apply(simp) +done + +lemma [elim]: "\0 < n; n \ Suc (Suc na); + rev b @ Oc # list = + Bk # replicate n Oc @ replicate (Suc (Suc na) - n) Bk @ + Oc # replicate na Oc; + length b = Suc n; b \ []\ + \ list_all isBk (dropWhile (\a. a = Oc) (tl b))" +apply(case_tac "rev b", auto) +done + +lemma b_cons_same: "b#bs = replicate x a @ as \ a \ b \ x = 0" +apply(case_tac x, simp+) +done + +lemma tcopy_tmp[elim]: + "\0 < n; n \ Suc (Suc na); + rev b @ Oc # list = + Bk # replicate n Oc @ replicate (Suc (Suc na) - n) Bk + @ Oc # replicate na Oc; length b = Suc n; b \ []\ + \ list = replicate na Oc" +apply(case_tac "rev b", simp+) +apply(auto) +apply(frule b_cons_same, auto) +done + +lemma [elim]: "\tstep (12, b, c) tcopy = (14, ab, ba); + tcopy_F12 x (b, c)\ \ tcopy_F14 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F12.simps tcopy_F14.simps + tcopy_def tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +apply(frule tcopy_tmp, simp+) +apply(case_tac n, simp+) +apply(case_tac nata, simp+) +done + +lemma replicate_app_Cons: "replicate a b @ b # replicate c b + = replicate (Suc (a + c)) b" +apply(simp) +apply(simp add: replicate_app_Cons_same) +apply(simp only: replicate_add[THEN sym]) +done + +lemma replicate_same_exE_pref: "\x. bs @ (b # cs) = replicate x y + \ (\n. bs = replicate n y)" +apply(induct bs) +apply(rule_tac x = 0 in exI, simp) +apply(drule impI) +apply(erule impE) +apply(erule exE, simp+) +apply(case_tac x, auto) +apply(case_tac x, auto) +apply(rule_tac x = "Suc n" in exI, simp+) +done + +lemma replicate_same_exE_inf: "\x. bs @ (b # cs) = replicate x y \ b = y" +apply(induct bs, auto) +apply(case_tac x, auto) +apply(drule impI) +apply(erule impE) +apply(case_tac x, simp+) +done + +lemma replicate_same_exE_suf: + "\x. bs @ (b # cs) = replicate x y \ \n. cs = replicate n y" +apply(induct bs, auto) +apply(case_tac x, simp+) +apply(drule impI, erule impE) +apply(case_tac x, simp+) +done + +lemma replicate_same_exE: "\x. bs @ (b # cs) = replicate x y + \ (\n. bs = replicate n y) & (b = y) & (\m. cs = replicate m y)" +apply(rule conjI) +apply(drule replicate_same_exE_pref, simp) +apply(rule conjI) +apply(drule replicate_same_exE_inf, simp) +apply(drule replicate_same_exE_suf, simp) +done + +lemma replicate_same: "bs @ (b # cs) = replicate x y + \ (\n. bs = replicate n y) & (b = y) & (\m. cs = replicate m y)" +apply(rule_tac replicate_same_exE) +apply(rule_tac x = x in exI) +apply(assumption) +done + +lemma [elim]: "\ 0 < n; n \ Suc (Suc na); + (rev ab @ Bk # list) = Bk # replicate n Oc + @ replicate (Suc (Suc na) - n) Bk @ Oc # replicate na Oc; ab \ []\ + \ n \ Suc na" +apply(rule contrapos_pp, simp+) +apply(case_tac "rev ab", simp+) +apply(auto) +apply(simp only: replicate_app_Cons) +apply(drule replicate_same) +apply(auto) +done + + +lemma [elim]: "\0 < n; n \ Suc (Suc na); + rev ab @ Bk # list = Bk # replicate n Oc @ + replicate (Suc (Suc na) - n) Bk @ Oc # replicate na Oc; + length ab = Suc n; ab \ []\ + \ rev ab @ Oc # list = Bk # Oc # replicate n Oc @ + replicate (Suc na - n) Bk @ Oc # replicate na Oc" +apply(case_tac "rev ab", simp+) +apply(auto) +apply(simp only: replicate_Cons_simp) +apply(simp) +apply(case_tac "Suc (Suc na) - n", simp+) +done + +lemma [elim]: "\tstep (12, b, c) tcopy = (13, ab, ba); + tcopy_F12 x (b, c)\ \ tcopy_F13 x (ab, ba)" +apply(case_tac x) +apply(simp_all add:tcopy_F12.simps tcopy_F13.simps + tcopy_def tstep.simps fetch.simps new_tape.simps) +apply(simp split: if_splits list.splits block.splits) +apply(auto) +done + + +lemma [elim]: "\tstep (11, b, c) tcopy = (12, ab, ba); + tcopy_F11 x (b, c)\ \ tcopy_F12 x (ab, ba)" +apply(case_tac x) +apply(simp_all add:tcopy_F12.simps tcopy_F11.simps + tcopy_def tstep.simps fetch.simps new_tape.simps) +apply(auto) +done + +lemma equal_length: "a = b \ length a = length b" +by(simp) + +lemma [elim]: "\tstep (13, b, c) tcopy = (12, ab, ba); + tcopy_F13 x (b, c)\ \ tcopy_F12 x (ab, ba)" +apply(case_tac x) +apply(simp_all add:tcopy_F12.simps tcopy_F13.simps + tcopy_def tstep.simps fetch.simps new_tape.simps) +apply(simp split: if_splits list.splits block.splits) +apply(auto) +apply(drule equal_length, simp) +done + +lemma [elim]: "\tstep (5, b, c) tcopy = (11, ab, ba); + tcopy_F5 x (b, c)\ \ tcopy_F11 x (ab, ba)" +apply(case_tac x) +apply(simp_all add:tcopy_F11.simps tcopy_F5.simps tcopy_def + tstep.simps fetch.simps new_tape.simps) +apply(simp split: if_splits list.splits block.splits) +done + +lemma less_equal: "\length xs <= b; \ Suc (length xs) <= b\ \ + length xs = b" +apply(simp) +done + +lemma length_cons_same: "\xs @ b # ys = as @ bs; + length ys = length bs\ \ xs @ [b] = as & ys = bs" +apply(drule rev_equal) +apply(simp) +apply(auto) +apply(drule rev_equal, simp) +done + +lemma replicate_set_equal: "\ xs @ [a] = replicate n b; a \ b\ \ RR" +apply(drule rev_equal, simp) +apply(case_tac n, simp+) +done + +lemma [elim]: "\tstep (10, b, c) tcopy = (10, ab, ba); + tcopy_F10 x (b, c)\ \ tcopy_F10 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F10.simps tcopy_def tstep.simps fetch.simps + new_tape.simps + split: if_splits list.splits block.splits) +apply(rule_tac x = n in exI, auto) +apply(case_tac b, simp+) +apply(rule contrapos_pp, simp+) +apply(frule less_equal, simp+) +apply(drule length_cons_same, auto) +apply(drule replicate_set_equal, simp+) +done + +lemma less_equal2: "\ (n::nat) \ m \ \x. n = x + m & x > 0" +apply(rule_tac x = "n - m" in exI) +apply(auto) +done + +lemma replicate_tail_length[dest]: + "\rev b @ Bk # list = xs @ replicate n Bk @ replicate n Oc\ + \ length list >= n" +apply(rule contrapos_pp, simp+) +apply(drule less_equal2, auto) +apply(drule rev_equal) +apply(simp add: replicate_add) +apply(auto) +apply(case_tac x, simp+) +done + + +lemma [elim]: "\tstep (9, b, c) tcopy = (10, ab, ba); + tcopy_F9 x (b, c)\ \ tcopy_F10 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F10.simps tcopy_F9.simps tcopy_def + tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +apply(rule_tac x = n in exI, auto) +apply(case_tac b, simp+) +done + +lemma [elim]: "\tstep (9, b, c) tcopy = (9, ab, ba); + tcopy_F9 x (b, c)\ \ tcopy_F9 x (ab, ba)" +apply(case_tac x) +apply(simp_all add: tcopy_F9.simps tcopy_def + tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +apply(rule_tac x = n in exI, auto) +apply(case_tac b, simp+) +apply(rule contrapos_pp, simp+) +apply(drule less_equal, simp+) +apply(drule rev_equal, auto) +apply(case_tac "length list", simp+) +done + +lemma app_cons_app_simp: "xs @ a # bs @ ys = (xs @ [a]) @ bs @ ys" +apply(simp) +done + +lemma [elim]: "\tstep (8, b, c) tcopy = (9, ab, ba); + tcopy_F8 x (b, c)\ \ tcopy_F9 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F8.simps tcopy_F9.simps tcopy_def + tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +apply(rule_tac x = "Suc n" in exI, auto) +apply(rule_tac x = "n" in exI, auto) +apply(simp only: app_cons_app_simp) +apply(frule replicate_tail_length, simp) +done + +lemma [elim]: "\tstep (8, b, c) tcopy = (8, ab, ba); + tcopy_F8 x (b, c)\ \ tcopy_F8 x (ab, ba)" +apply(case_tac x) +apply(simp_all add:tcopy_F8.simps tcopy_def tstep.simps + fetch.simps new_tape.simps) +apply(simp split: if_splits list.splits block.splits) +apply(rule_tac x = "n" in exI, auto) +done + +lemma ex_less_more: "\(x::nat) >= m ; x <= n\ \ + \y. x = m + y & y <= n - m" +by(rule_tac x = "x -m" in exI, auto) + +lemma replicate_split: "x <= n \ + (\y. replicate n b = replicate (y + x) b)" +apply(rule_tac x = "n - x" in exI) +apply(simp) +done + +lemma app_app_app_app_simp: "as @ bs @ cs @ ds = + (as @ bs) @ (cs @ ds)" +by simp + +lemma lengthtailsame_append_elim: + "\as @ bs = cs @ ds; length bs = length ds\ \ bs = ds" +apply(simp) +done + +lemma rep_suc: "replicate (Suc n) x = replicate n x @ [x]" +by(induct n, auto) + +lemma length_append_diff_cons: + "\b @ x # ba = xs @ replicate m y @ replicate n x; x \ y; + Suc (length ba) \ m + n\ + \ length ba < n" +apply(induct n arbitrary: ba, simp) +apply(drule_tac b = y in replicate_split, + simp add: replicate_add, erule exE, simp del: replicate.simps) +proof - + fix ba ya + assume h1: + "b @ x # ba = xs @ y # replicate ya y @ replicate (length ba) y" + and h2: "x \ y" + thus "False" + using append_eq_append_conv[of "b @ [x]" + "xs @ y # replicate ya y" "ba" "replicate (length ba) y"] + apply(auto) + apply(case_tac ya, simp, + simp add: rep_suc del: rep_suc_rev replicate.simps) + done +next + fix n ba + assume ind: "\ba. \b @ x # ba = xs @ replicate m y @ replicate n x; + x \ y; Suc (length ba) \ m + n\ + \ length ba < n" + and h1: "b @ x # ba = xs @ replicate m y @ replicate (Suc n) x" + and h2: "x \ y" and h3: "Suc (length ba) \ m + Suc n" + show "length ba < Suc n" + proof(cases "length ba") + case 0 thus "?thesis" by simp + next + fix nat + assume "length ba = Suc nat" + hence "\ ys a. ba = ys @ [a]" + apply(rule_tac x = "butlast ba" in exI) + apply(rule_tac x = "last ba" in exI) + using append_butlast_last_id[of ba] + apply(case_tac ba, auto) + done + from this obtain ys where "\ a. ba = ys @ [a]" .. + from this obtain a where "ba = ys @ [a]" .. + thus "?thesis" + using ind[of ys] h1 h2 h3 + apply(simp del: rep_suc_rev replicate.simps add: rep_suc) + done + qed +qed + +lemma [elim]: + "\b @ Oc # ba = xs @ Bk # replicate n Bk @ replicate n Oc; + Suc (length ba) \ 2 * n\ + \ length ba < n" + apply(rule_tac length_append_diff_cons[of b Oc ba xs "Suc n" Bk n]) + apply(simp, simp, simp) + done + +lemma [elim]: "\tstep (7, b, c) tcopy = (8, ab, ba); + tcopy_F7 x (b, c)\ \ tcopy_F8 x (ab, ba)" +apply(case_tac x) +apply(simp_all add:tcopy_F8.simps tcopy_F7.simps + tcopy_def tstep.simps fetch.simps new_tape.simps) +apply(simp split: if_splits list.splits block.splits) +apply(rule_tac x = "n" in exI, auto) +done + +lemma [elim]: "\tstep (7, b, c) tcopy = (7, ab, ba); + tcopy_F7 x (b, c)\ \ tcopy_F7 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F7.simps tcopy_def tstep.simps + fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +apply(rule_tac x = "n" in exI, auto) +apply(simp only: app_cons_app_simp) +apply(frule replicate_tail_length, simp) +done + +lemma Suc_more: "n <= m \ Suc m - n = Suc (m - n)" +by simp + +lemma [elim]: "\tstep (6, b, c) tcopy = (7, ab, ba); + tcopy_F6 x (b, c)\ \ tcopy_F7 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F7.simps tcopy_F6.simps + tcopy_def tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +done + +lemma [elim]: "\tstep (6, b, c) tcopy = (6, ab, ba); + tcopy_F6 x (b, c)\ \ tcopy_F6 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F6.simps tcopy_def tstep.simps + new_tape.simps fetch.simps + split: if_splits list.splits block.splits) +done + +lemma [elim]: "\tstep (5, b, c) tcopy = (6, ab, ba); + tcopy_F5 x (b, c)\ \ tcopy_F6 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F5.simps tcopy_F6.simps tcopy_def + tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +apply(rule_tac x = n in exI, simp) +apply(rule_tac x = n in exI, simp) +apply(drule Suc_more, simp) +done + +lemma ex_less_more2: "\(n::nat) < x ; x <= 2 * n\ \ + \y. (x = n + y & y <= n)" +apply(rule_tac x = "x - n" in exI) +apply(auto) +done + +lemma app_app_app_simp: "xs @ ys @ za = (xs @ ys) @ za" +apply(simp) +done + +lemma [elim]: "rev xs = replicate n b \ xs = replicate n b" +using rev_replicate[of n b] +thm rev_equal +by(drule_tac rev_equal, simp) + +lemma app_cons_tail_same[dest]: + "\rev b @ Oc # list = + replicate (Suc (Suc na) - n) Oc @ replicate n Bk @ replicate n Oc; + Suc 0 < n; n \ Suc na; n < length list; length list \ 2 * n; b \ []\ + \ list = replicate n Bk @ replicate n Oc + & b = replicate (Suc na - n) Oc" +using length_append_diff_cons[of "rev b" Oc list + "replicate (Suc (Suc na) - n) Oc" n Bk n] +apply(case_tac "length list = 2*n", simp) +using append_eq_append_conv[of "rev b @ [Oc]" "replicate + (Suc (Suc na) - n) Oc" "list" "replicate n Bk @ replicate n Oc"] +apply(case_tac n, simp, simp add: Suc_more rep_suc + del: rep_suc_rev replicate.simps, auto) +done + +lemma hd_replicate_false1: "\replicate x Oc \ []; + hd (replicate x Oc) = Bk\ \ RR" +apply(case_tac x, auto) +done + +lemma hd_replicate_false2: "\replicate x Oc \ []; + hd (replicate x Oc) \ Oc\ \ RR" +apply(case_tac x, auto) +done + +lemma Suc_more_less: "\n \ Suc m; n >= m\ \ n = m | n = Suc m" +apply(auto) +done + +lemma replicate_not_Nil: "replicate x a \ [] \ x > 0" +apply(case_tac x, simp+) +done + +lemma [elim]: "\tstep (10, b, c) tcopy = (5, ab, ba); + tcopy_F10 x (b, c)\ \ tcopy_F5 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F5.simps tcopy_F10.simps tcopy_def + tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +apply(frule app_cons_tail_same, simp+) +apply(rule_tac x = n in exI, auto) +done + +lemma [elim]: "\tstep (4, b, c) tcopy = (5, ab, ba); + tcopy_F4 x (b, c)\ \ tcopy_F5 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F5.simps tcopy_F4.simps tcopy_def + tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +done + +lemma [elim]: "\tstep (3, b, c) tcopy = (4, ab, ba); + tcopy_F3 x (b, c)\ \ tcopy_F4 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F3.simps tcopy_F4.simps + tcopy_def tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +done + +lemma [elim]: "\tstep (4, b, c) tcopy = (4, ab, ba); + tcopy_F4 x (b, c)\ \ tcopy_F4 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F3.simps tcopy_F4.simps + tcopy_def tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +done + +lemma [elim]: "\tstep (3, b, c) tcopy = (3, ab, ba); + tcopy_F3 x (b, c)\ \ tcopy_F3 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F3.simps tcopy_F4.simps + tcopy_def tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +done + +lemma replicate_cons_back: "y # replicate x y = replicate (Suc x) y" +apply(simp) +done + +lemma replicate_cons_same: "bs @ (b # cs) = y # replicate x y \ + (\n. bs = replicate n y) & (b = y) & (\m. cs = replicate m y)" +apply(simp only: replicate_cons_back) +apply(drule_tac replicate_same) +apply(simp) +done + +lemma [elim]: "\tstep (2, b, c) tcopy = (3, ab, ba); + tcopy_F2 x (b, c)\ \ tcopy_F3 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F3.simps tcopy_F2.simps + tcopy_def tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +apply(drule replicate_cons_same, auto)+ +done + +lemma [elim]: "\tstep (2, b, c) tcopy = (2, ab, ba); + tcopy_F2 x (b, c)\ \ tcopy_F2 x (ab, ba)" +apply(case_tac x) +apply(auto simp:tcopy_F3.simps tcopy_F2.simps + tcopy_def tstep.simps fetch.simps new_tape.simps + split: if_splits list.splits block.splits) +apply(frule replicate_cons_same, auto) +apply(simp add: replicate_app_Cons_same) +done + +lemma [elim]: "\tstep (Suc 0, b, c) tcopy = (2, ab, ba); + tcopy_F1 x (b, c)\ \ tcopy_F2 x (ab, ba)" +apply(case_tac x) +apply(simp_all add:tcopy_F2.simps tcopy_F1.simps + tcopy_def tstep.simps fetch.simps new_tape.simps) +apply(auto) +done + +lemma [elim]: "\tstep (Suc 0, b, c) tcopy = (0, ab, ba); + tcopy_F1 x (b, c)\ \ tcopy_F0 x (ab, ba)" +apply(case_tac x) +apply(simp_all add:tcopy_F0.simps tcopy_F1.simps + tcopy_def tstep.simps fetch.simps new_tape.simps) +done + +lemma ex_less: "Suc x <= y \ \z. y = x + z & z > 0" +apply(rule_tac x = "y - x" in exI, auto) +done + +lemma [elim]: "\xs @ Bk # ba = + Bk # Oc # replicate n Oc @ Bk # Oc # replicate n Oc; + length xs \ Suc n; xs \ []\ \ RR" +apply(case_tac xs, auto) +apply(case_tac list, auto) +apply(drule ex_less, auto) +apply(simp add: replicate_add) +apply(auto) +apply(case_tac z, simp+) +done + +lemma [elim]: "\tstep (15, b, c) tcopy = (0, ab, ba); + tcopy_F15 x (b, c)\ \ tcopy_F0 x (ab, ba)" +apply(case_tac x) +apply(auto simp: tcopy_F15.simps tcopy_F0.simps + tcopy_def tstep.simps new_tape.simps fetch.simps + split: if_splits list.splits block.splits) +done + +lemma [elim]: "\tstep (0, b, c) tcopy = (0, ab, ba); + tcopy_F0 x (b, c)\ \ tcopy_F0 x (ab, ba)" +apply(case_tac x) +apply(simp_all add: tcopy_F0.simps tcopy_def + tstep.simps new_tape.simps fetch.simps) +done + +declare tstep.simps[simp del] + +text {* + Finally establishes the invariant of Copying TM, which is used to dervie + the parital correctness of Copying TM. +*} +lemma inv_tcopy_step:"inv_tcopy x c \ inv_tcopy x (tstep c tcopy)" +apply(induct c) +apply(auto split: if_splits block.splits list.splits taction.splits) +apply(auto simp: tstep.simps tcopy_def fetch.simps new_tape.simps + split: if_splits list.splits block.splits taction.splits) +done + +declare inv_tcopy.simps[simp del] + +text {* + Invariant under mult-step execution. + *} +lemma inv_tcopy_steps: + "inv_tcopy x (steps (Suc 0, [], replicate x Oc) tcopy stp) " +apply(induct stp) +apply(simp add: tstep.simps tcopy_def steps.simps + tcopy_F1.simps inv_tcopy.simps) +apply(drule_tac inv_tcopy_step, simp add: tstep_red) +done + + +text {* + The followng lemmas gives the parital correctness of Copying TM. +*} +theorem inv_tcopy_rs: + "steps (Suc 0, [], replicate x Oc) tcopy stp = (0, l, r) + \ \ n. l = replicate n Bk \ + r = replicate x Oc @ Bk # replicate x Oc" +apply(insert inv_tcopy_steps[of x stp]) +apply(auto simp: inv_tcopy.simps tcopy_F0.simps isBk.simps) +done + + + + +(*----------halt problem of tcopy----------------------------------------*) + +section {* + The following definitions are used to construct the measure function used to show + the termnation of Copying TM. +*} + +definition lex_pair :: "((nat \ nat) \ nat \ nat) set" + where + "lex_pair \ less_than <*lex*> less_than" + +definition lex_triple :: + "((nat \ (nat \ nat)) \ (nat \ (nat \ nat))) set" + where +"lex_triple \ less_than <*lex*> lex_pair" + +definition lex_square :: + "((nat \ nat \ nat \ nat) \ (nat \ nat \ nat \ nat)) set" + where +"lex_square \ less_than <*lex*> lex_triple" + +lemma wf_lex_triple: "wf lex_triple" + by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def) + +lemma wf_lex_square: "wf lex_square" + by (auto intro:wf_lex_prod + simp:lex_triple_def lex_square_def lex_pair_def) + +text {* + A measurement functions used to show the termination of copying machine: +*} +fun tcopy_phase :: "t_conf \ nat" + where + "tcopy_phase c = (let (state, tp) = c in + if state > 0 & state <= 4 then 5 + else if state >=5 & state <= 10 then 4 + else if state = 11 then 3 + else if state = 12 | state = 13 then 2 + else if state = 14 | state = 15 then 1 + else 0)" + +fun tcopy_phase4_stage :: "tape \ nat" + where + "tcopy_phase4_stage (ln, rn) = + (let lrn = (rev ln) @ rn + in length (takeWhile (\a. a = Oc) lrn))" + +fun tcopy_stage :: "t_conf \ nat" + where + "tcopy_stage c = (let (state, ln, rn) = c in + if tcopy_phase c = 5 then 0 + else if tcopy_phase c = 4 then + tcopy_phase4_stage (ln, rn) + else if tcopy_phase c = 3 then 0 + else if tcopy_phase c = 2 then length rn + else if tcopy_phase c = 1 then 0 + else 0)" + +fun tcopy_phase4_state :: "t_conf \ nat" + where + "tcopy_phase4_state c = (let (state, ln, rn) = c in + if state = 6 & hd rn = Oc then 0 + else if state = 5 then 1 + else 12 - state)" + +fun tcopy_state :: "t_conf \ nat" + where + "tcopy_state c = (let (state, ln, rn) = c in + if tcopy_phase c = 5 then 4 - state + else if tcopy_phase c = 4 then + tcopy_phase4_state c + else if tcopy_phase c = 3 then 0 + else if tcopy_phase c = 2 then 13 - state + else if tcopy_phase c = 1 then 15 - state + else 0)" + +fun tcopy_step2 :: "t_conf \ nat" + where + "tcopy_step2 (s, l, r) = length r" + +fun tcopy_step3 :: "t_conf \ nat" + where + "tcopy_step3 (s, l, r) = (if r = [] | r = [Bk] then Suc 0 else 0)" + +fun tcopy_step4 :: "t_conf \ nat" + where + "tcopy_step4 (s, l, r) = length l" + +fun tcopy_step7 :: "t_conf \ nat" + where + "tcopy_step7 (s, l, r) = length r" + +fun tcopy_step8 :: "t_conf \ nat" + where + "tcopy_step8 (s, l, r) = length r" + +fun tcopy_step9 :: "t_conf \ nat" + where + "tcopy_step9 (s, l, r) = length l" + +fun tcopy_step10 :: "t_conf \ nat" + where + "tcopy_step10 (s, l, r) = length l" + +fun tcopy_step14 :: "t_conf \ nat" + where + "tcopy_step14 (s, l, r) = (case hd r of + Oc \ 1 | + Bk \ 0)" + +fun tcopy_step15 :: "t_conf \ nat" + where + "tcopy_step15 (s, l, r) = length l" + +fun tcopy_step :: "t_conf \ nat" + where + "tcopy_step c = (let (state, ln, rn) = c in + if state = 0 | state = 1 | state = 11 | + state = 5 | state = 6 | state = 12 | state = 13 then 0 + else if state = 2 then tcopy_step2 c + else if state = 3 then tcopy_step3 c + else if state = 4 then tcopy_step4 c + else if state = 7 then tcopy_step7 c + else if state = 8 then tcopy_step8 c + else if state = 9 then tcopy_step9 c + else if state = 10 then tcopy_step10 c + else if state = 14 then tcopy_step14 c + else if state = 15 then tcopy_step15 c + else 0)" + +text {* + The measure function used to show the termination of Copying TM. +*} +fun tcopy_measure :: "t_conf \ (nat * nat * nat * nat)" + where + "tcopy_measure c = + (tcopy_phase c, tcopy_stage c, tcopy_state c, tcopy_step c)" + +definition tcopy_LE :: "((nat \ block list \ block list) \ + (nat \ block list \ block list)) set" + where + "tcopy_LE \ (inv_image lex_square tcopy_measure)" + +lemma wf_tcopy_le: "wf tcopy_LE" +by(auto intro:wf_inv_image wf_lex_square simp:tcopy_LE_def) + + +declare steps.simps[simp del] + +declare tcopy_phase.simps[simp del] tcopy_stage.simps[simp del] + tcopy_state.simps[simp del] tcopy_step.simps[simp del] + inv_tcopy.simps[simp del] +declare tcopy_F0.simps [simp] + tcopy_F1.simps [simp] + tcopy_F2.simps [simp] + tcopy_F3.simps [simp] + tcopy_F4.simps [simp] + tcopy_F5.simps [simp] + tcopy_F6.simps [simp] + tcopy_F7.simps [simp] + tcopy_F8.simps [simp] + tcopy_F9.simps [simp] + tcopy_F10.simps [simp] + tcopy_F11.simps [simp] + tcopy_F12.simps [simp] + tcopy_F13.simps [simp] + tcopy_F14.simps [simp] + tcopy_F15.simps [simp] + fetch.simps[simp] + new_tape.simps[simp] +lemma [elim]: "tcopy_F1 x (b, c) \ + (tstep (Suc 0, b, c) tcopy, Suc 0, b, c) \ tcopy_LE" +apply(simp add: tcopy_F1.simps tstep.simps tcopy_def tcopy_LE_def + lex_square_def lex_triple_def lex_pair_def tcopy_phase.simps + tcopy_stage.simps tcopy_state.simps tcopy_step.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +done + +lemma [elim]: "tcopy_F2 x (b, c) \ + (tstep (2, b, c) tcopy, 2, b, c) \ tcopy_LE" +apply(simp add:tstep.simps tcopy_def tcopy_LE_def lex_square_def + lex_triple_def lex_pair_def tcopy_phase.simps tcopy_stage.simps + tcopy_state.simps tcopy_step.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +done + +lemma [elim]: "tcopy_F3 x (b, c) \ + (tstep (3, b, c) tcopy, 3, b, c) \ tcopy_LE" +apply(simp add: tstep.simps tcopy_def tcopy_LE_def lex_square_def + lex_triple_def lex_pair_def tcopy_phase.simps tcopy_stage.simps + tcopy_state.simps tcopy_step.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +apply(case_tac x, simp+) +done + +lemma [elim]: "tcopy_F4 x (b, c) \ + (tstep (4, b, c) tcopy, 4, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tcopy_F4.simps tstep.simps tcopy_def tcopy_LE_def + lex_square_def lex_triple_def lex_pair_def tcopy_phase.simps + tcopy_stage.simps tcopy_state.simps tcopy_step.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +done + +lemma[simp]: "takeWhile (\a. a = b) (replicate x b @ ys) = + replicate x b @ (takeWhile (\a. a = b) ys)" +apply(induct x) +apply(simp+) +done + +lemma [elim]: "tcopy_F5 x (b, c) \ + (tstep (5, b, c) tcopy, 5, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tstep.simps tcopy_def tcopy_LE_def + lex_square_def lex_triple_def lex_pair_def tcopy_phase.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +apply(simp_all add: tcopy_phase.simps + tcopy_stage.simps tcopy_state.simps) +done + +lemma [elim]: "\replicate n x = []; n > 0\ \ RR" +apply(case_tac n, simp+) +done + +lemma [elim]: "tcopy_F6 x (b, c) \ + (tstep (6, b, c) tcopy, 6, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tstep.simps tcopy_def tcopy_LE_def + lex_square_def lex_triple_def lex_pair_def + tcopy_phase.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +apply(simp_all add: tcopy_phase.simps tcopy_stage.simps + tcopy_state.simps tcopy_step.simps) +done + +lemma [elim]: "tcopy_F7 x (b, c) \ + (tstep (7, b, c) tcopy, 7, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tstep.simps tcopy_def tcopy_LE_def lex_square_def + lex_triple_def lex_pair_def tcopy_phase.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +apply(simp_all add: tcopy_phase.simps tcopy_stage.simps + tcopy_state.simps tcopy_step.simps) +done + +lemma [elim]: "tcopy_F8 x (b, c) \ + (tstep (8, b, c) tcopy, 8, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tstep.simps tcopy_def tcopy_LE_def lex_square_def + lex_triple_def lex_pair_def tcopy_phase.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +apply(simp_all add: tcopy_phase.simps tcopy_stage.simps + tcopy_state.simps tcopy_step.simps) +apply(simp only: app_cons_app_simp, frule replicate_tail_length, simp) +done + +lemma app_app_app_equal: "xs @ ys @ zs = (xs @ ys) @ zs" +by simp + +lemma append_cons_assoc: "as @ b # bs = (as @ [b]) @ bs" +apply(rule rev_equal_rev) +apply(simp) +done + +lemma rev_tl_hd_merge: "bs \ [] \ + rev (tl bs) @ hd bs # as = rev bs @ as" +apply(rule rev_equal_rev) +apply(simp) +done + +lemma [elim]: "tcopy_F9 x (b, c) \ + (tstep (9, b, c) tcopy, 9, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tstep.simps tcopy_def tcopy_LE_def lex_square_def + lex_triple_def lex_pair_def tcopy_phase.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +apply(simp_all add: tcopy_phase.simps tcopy_stage.simps + tcopy_state.simps tcopy_step.simps) +apply(drule_tac bs = b and as = "Bk # list" in rev_tl_hd_merge) +apply(simp) +apply(drule_tac bs = b and as = "Oc # list" in rev_tl_hd_merge) +apply(simp) +done + +lemma [elim]: "tcopy_F10 x (b, c) \ + (tstep (10, b, c) tcopy, 10, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tstep.simps tcopy_def tcopy_LE_def lex_square_def + lex_triple_def lex_pair_def tcopy_phase.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +apply(simp_all add: tcopy_phase.simps tcopy_stage.simps + tcopy_state.simps tcopy_step.simps) +apply(drule_tac bs = b and as = "Bk # list" in rev_tl_hd_merge) +apply(simp) +apply(drule_tac bs = b and as = "Oc # list" in rev_tl_hd_merge) +apply(simp) +done + +lemma [elim]: "tcopy_F11 x (b, c) \ + (tstep (11, b, c) tcopy, 11, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tstep.simps tcopy_def tcopy_LE_def + lex_square_def lex_triple_def lex_pair_def + tcopy_phase.simps) +done + +lemma [elim]: "tcopy_F12 x (b, c) \ + (tstep (12, b, c) tcopy, 12, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tstep.simps tcopy_def tcopy_LE_def lex_square_def + lex_triple_def lex_pair_def tcopy_phase.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +apply(simp_all add: tcopy_phase.simps tcopy_stage.simps + tcopy_state.simps tcopy_step.simps) +done + +lemma [elim]: "tcopy_F13 x (b, c) \ + (tstep (13, b, c) tcopy, 13, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tstep.simps tcopy_def tcopy_LE_def lex_square_def + lex_triple_def lex_pair_def tcopy_phase.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +apply(simp_all add: tcopy_phase.simps tcopy_stage.simps + tcopy_state.simps tcopy_step.simps) +apply(drule equal_length, simp)+ +done + +lemma [elim]: "tcopy_F14 x (b, c) \ + (tstep (14, b, c) tcopy, 14, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tstep.simps tcopy_def tcopy_LE_def lex_square_def + lex_triple_def lex_pair_def tcopy_phase.simps) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +apply(simp_all add: tcopy_phase.simps tcopy_stage.simps + tcopy_state.simps tcopy_step.simps) +done + +lemma [elim]: "tcopy_F15 x (b, c) \ + (tstep (15, b, c) tcopy, 15, b, c) \ tcopy_LE" +apply(case_tac x, simp) +apply(simp add: tstep.simps tcopy_def tcopy_LE_def lex_square_def + lex_triple_def lex_pair_def tcopy_phase.simps ) +apply(simp split: if_splits list.splits block.splits taction.splits) +apply(auto) +apply(simp_all add: tcopy_phase.simps tcopy_stage.simps + tcopy_state.simps tcopy_step.simps) +done + +lemma tcopy_wf_step:"\a > 0; inv_tcopy x (a, b, c)\ \ + (tstep (a, b, c) tcopy, (a, b, c)) \ tcopy_LE" +apply(simp add:inv_tcopy.simps split: if_splits, auto) +apply(auto simp: tstep.simps tcopy_def tcopy_LE_def lex_square_def + lex_triple_def lex_pair_def tcopy_phase.simps + tcopy_stage.simps tcopy_state.simps tcopy_step.simps + split: if_splits list.splits block.splits taction.splits) +done + +lemma tcopy_wf: +"\n. let nc = steps (Suc 0, [], replicate x Oc) tcopy n in + let Sucnc = steps (Suc 0, [], replicate x Oc) tcopy (Suc n) in + \ isS0 nc \ ((Sucnc, nc) \ tcopy_LE)" +proof(rule allI, case_tac + "steps (Suc 0, [], replicate x Oc) tcopy n", auto simp: tstep_red) + fix n a b c + assume h: "\ isS0 (a, b, c)" + "steps (Suc 0, [], replicate x Oc) tcopy n = (a, b, c)" + hence "inv_tcopy x (a, b, c)" + using inv_tcopy_steps[of x n] by(simp) + thus "(tstep (a, b, c) tcopy, a, b, c) \ tcopy_LE" + using h + by(rule_tac tcopy_wf_step, auto simp: isS0_def) +qed + +text {* + The termination of Copying TM: +*} +lemma tcopy_halt: + "\n. isS0 (steps (Suc 0, [], replicate x Oc) tcopy n)" +apply(insert halt_lemma + [of tcopy_LE isS0 "steps (Suc 0, [], replicate x Oc) tcopy"]) +apply(insert tcopy_wf [of x]) +apply(simp only: Let_def) +apply(insert wf_tcopy_le) +apply(simp) +done + +text {* + The total correntess of Copying TM: +*} +theorem tcopy_halt_rs: "\stp m. + steps (Suc 0, [], replicate x Oc) tcopy stp = + (0, replicate m Bk, replicate x Oc @ Bk # replicate x Oc)" +using tcopy_halt[of x] +proof(erule_tac exE) + fix n + assume h: "isS0 (steps (Suc 0, [], replicate x Oc) tcopy n)" + have "inv_tcopy x (steps (Suc 0, [], replicate x Oc) tcopy n)" + using inv_tcopy_steps[of x n] by simp + thus "?thesis" + using h + apply(cases "(steps (Suc 0, [], replicate x Oc) tcopy n)", + auto simp: isS0_def inv_tcopy.simps) + apply(rule_tac x = n in exI, auto) + done +qed + +section {* + The {\em Dithering} Turing Machine +*} + +text {* + The {\em Dithering} TM, when the input is @{text "1"}, it will loop forever, otherwise, it will + terminate. +*} +definition dither :: "tprog" + where + "dither \ [(W0, 1), (R, 2), (L, 1), (L, 0)] " + +lemma dither_halt_rs: + "\ stp. steps (Suc 0, Bk\<^bsup>m\<^esup>, [Oc, Oc]) dither stp = + (0, Bk\<^bsup>m\<^esup>, [Oc, Oc])" +apply(rule_tac x = "Suc (Suc (Suc 0))" in exI) +apply(simp add: dither_def steps.simps + tstep.simps fetch.simps new_tape.simps) +done + +lemma dither_unhalt_state: + "(steps (Suc 0, Bk\<^bsup>m\<^esup>, [Oc]) dither stp = + (Suc 0, Bk\<^bsup>m\<^esup>, [Oc])) \ + (steps (Suc 0, Bk\<^bsup>m\<^esup>, [Oc]) dither stp = (2, Oc # Bk\<^bsup>m\<^esup>, []))" + apply(induct stp, simp add: steps.simps) + apply(simp add: tstep_red, auto) + apply(auto simp: tstep.simps fetch.simps dither_def new_tape.simps) + done + +lemma dither_unhalt_rs: + "\ (\ stp. isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, [Oc]) dither stp))" +proof(auto) + fix stp + assume h1: "isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, [Oc]) dither stp)" + have "\ isS0 ((steps (Suc 0, Bk\<^bsup>m\<^esup>, [Oc]) dither stp))" + using dither_unhalt_state[of m stp] + by(auto simp: isS0_def) + from h1 and this show False by (auto) +qed + +section {* + The final diagnal arguments to show the undecidability of Halting problem. +*} + +text {* + @{text "haltP tp x"} means TM @{text "tp"} terminates on input @{text "x"} + and the final configuration is standard. +*} +definition haltP :: "tprog \ nat \ bool" + where + "haltP t x = (\n a b c. steps (Suc 0, [], Oc\<^bsup>x\<^esup>) t n = (0, Bk\<^bsup>a\<^esup>, Oc\<^bsup>b\<^esup> @ Bk\<^bsup>c\<^esup>))" + +lemma [simp]: "length (A |+| B) = length A + length B" +by(auto simp: t_add.simps tshift.simps) + +lemma [intro]: "\iseven (x::nat); iseven y\ \ iseven (x + y)" +apply(auto simp: iseven_def) +apply(rule_tac x = "x + xa" in exI, simp) +done + +lemma t_correct_add[intro]: + "\t_correct A; t_correct B\ \ t_correct (A |+| B)" +apply(auto simp: t_correct.simps tshift.simps t_add.simps + change_termi_state.simps list_all_iff) +apply(erule_tac x = "(a, b)" in ballE, auto) +apply(case_tac "ba = 0", auto) +done + +lemma [intro]: "t_correct tcopy" +apply(simp add: t_correct.simps tcopy_def iseven_def) +apply(rule_tac x = 15 in exI, simp) +done + +lemma [intro]: "t_correct dither" +apply(simp add: t_correct.simps dither_def iseven_def) +apply(rule_tac x = 2 in exI, simp) +done + +text {* + The following locale specifies that TM @{text "H"} can be used to solve + the {\em Halting Problem} and @{text "False"} is going to be derived + under this locale. Therefore, the undecidability of {\em Halting Problem} + is established. +*} +locale uncomputable = + -- {* The coding function of TM, interestingly, the detailed definition of this + funciton @{text "code"} does not affect the final result. *} + fixes code :: "tprog \ nat" + -- {* + The TM @{text "H"} is the one which is assummed being able to solve the Halting problem. + *} + and H :: "tprog" + assumes h_wf[intro]: "t_correct H" + -- {* + The following two assumptions specifies that @{text "H"} does solve the Halting problem. + *} + and h_case: + "\ M n. \(haltP M n)\ \ + \ na nb. (steps (Suc 0, Bk\<^bsup>x\<^esup>, Oc\<^bsup>code M\<^esup> @ Bk # Oc\<^bsup>n\<^esup>) H na = (0, Bk\<^bsup>nb\<^esup>, [Oc]))" + and nh_case: + "\ M n. \(\ haltP M n)\ \ + \ na nb. (steps (Suc 0, Bk\<^bsup>x\<^esup>, Oc\<^bsup>code M\<^esup> @ Bk # Oc\<^bsup>n\<^esup>) H na = (0, Bk\<^bsup>nb\<^esup>, [Oc, Oc]))" +begin + +term t_correct +declare haltP_def[simp del] +definition tcontra :: "tprog \ tprog" + where + "tcontra h \ ((tcopy |+| h) |+| dither)" + +lemma [simp]: "a\<^bsup>0\<^esup> = []" + by(simp add: exponent_def) +lemma haltP_weaking: + "haltP (tcontra H) (code (tcontra H)) \ + \stp. isS0 (steps (Suc 0, [], Oc\<^bsup>code (tcontra H)\<^esup>) + ((tcopy |+| H) |+| dither) stp)" + apply(simp add: haltP_def, auto) + apply(rule_tac x = n in exI, simp add: isS0_def tcontra_def) + done + +lemma h_uh: "haltP (tcontra H) (code (tcontra H)) + \ \ haltP (tcontra H) (code (tcontra H))" +proof - + let ?cn = "code (tcontra H)" + let ?P1 = "\ tp. let (l, r) = tp in (l = [] \ + (r::block list) = Oc\<^bsup>(?cn)\<^esup>)" + let ?Q1 = "\ (l, r).(\ nb. l = Bk\<^bsup>nb\<^esup> \ + r = Oc\<^bsup>(?cn)\<^esup> @ Bk # Oc\<^bsup>(?cn)\<^esup>)" + let ?P2 = ?Q1 + let ?Q2 = "\ (l, r). (\ nd. l = Bk\<^bsup>nd \<^esup>\ r = [Oc])" + let ?P3 = "\ tp. False" + assume h: "haltP (tcontra H) (code (tcontra H))" + hence h1: "\ x. \ na nb. steps (Suc 0, Bk\<^bsup>x\<^esup>, Oc\<^bsup>code (tcontra H)\<^esup> @ Bk # + Oc\<^bsup>code (tcontra H)\<^esup>) H na = (0, Bk\<^bsup>nb\<^esup>, [Oc])" + by(drule_tac x = x in h_case, simp) + have "?P1 \-> \ tp. (\ stp tp'. steps (Suc 0, tp) (tcopy |+| H) stp = (0, tp') \ ?Q2 tp')" + proof(rule_tac turing_merge.t_merge_halt[of tcopy H "?P1" "?P2" "?P3" + "?P3" "?Q1" "?Q2"], auto simp: turing_merge_def) + show "\stp. case steps (Suc 0, [], Oc\<^bsup>?cn\<^esup>) tcopy stp of (s, tp') \ + s = 0 \ (case tp' of (l, r) \ (\nb. l = Bk\<^bsup>nb\<^esup>) \ r = Oc\<^bsup>?cn\<^esup> @ Bk # Oc\<^bsup>?cn\<^esup>)" + using tcopy_halt_rs[of "?cn"] + apply(auto) + apply(rule_tac x = stp in exI, auto simp: exponent_def) + done + next + fix nb + show "\stp. case steps (Suc 0, Bk\<^bsup>nb\<^esup>, Oc\<^bsup>code (tcontra H)\<^esup> @ Bk # Oc\<^bsup>code (tcontra H)\<^esup>) H stp of + (s, tp') \ s = 0 \ (case tp' of (l, r) \ (\nd. l = Bk\<^bsup>nd\<^esup>) \ r = [Oc])" + using h1[of nb] + apply(auto) + apply(rule_tac x = na in exI, auto) + done + next + show "\(l, r). ((\nb. l = Bk\<^bsup>nb\<^esup>) \ r = Oc\<^bsup>code (tcontra H)\<^esup> @ Bk # Oc\<^bsup>code (tcontra H)\<^esup>) \-> + \(l, r). ((\nb. l = Bk\<^bsup>nb\<^esup>) \ r = Oc\<^bsup>code (tcontra H)\<^esup> @ Bk # Oc\<^bsup>code (tcontra H)\<^esup>)" + apply(simp add: t_imply_def) + done + qed + hence "\stp tp'. steps (Suc 0, [], Oc\<^bsup>?cn\<^esup>) (tcopy |+| H) stp = (0, tp') \ + (case tp' of (l, r) \ \nd. l = Bk\<^bsup>nd\<^esup> \ r = [Oc])" + apply(simp add: t_imply_def) + done + hence "?P1 \-> \ tp. \ (\ stp. isS0 (steps (Suc 0, tp) ((tcopy |+| H) |+| dither) stp))" + proof(rule_tac turing_merge.t_merge_uhalt[of "tcopy |+| H" dither "?P1" "?P3" "?P3" + "?Q2" "?Q2" "?Q2"], simp add: turing_merge_def, auto) + fix stp nd + assume "steps (Suc 0, [], Oc\<^bsup>code (tcontra H)\<^esup>) (tcopy |+| H) stp = (0, Bk\<^bsup>nd\<^esup>, [Oc])" + thus "\stp. case steps (Suc 0, [], Oc\<^bsup>code (tcontra H)\<^esup>) (tcopy |+| H) stp of (s, tp') + \ s = 0 \ (case tp' of (l, r) \ (\nd. l = Bk\<^bsup>nd\<^esup>) \ r = [Oc])" + apply(rule_tac x = stp in exI, auto) + done + next + fix stp nd nda stpa + assume "isS0 (steps (Suc 0, Bk\<^bsup>nda\<^esup>, [Oc]) dither stpa)" + thus "False" + using dither_unhalt_rs[of nda] + apply auto + done + next + fix stp nd + show "\(l, r). ((\nd. l = Bk\<^bsup>nd\<^esup>) \ r = [Oc]) \-> + \(l, r). ((\nd. l = Bk\<^bsup>nd\<^esup>) \ r = [Oc])" + by (simp add: t_imply_def) + qed + thus "\ haltP (tcontra H) (code (tcontra H))" + apply(simp add: t_imply_def haltP_def tcontra_def, auto) + apply(erule_tac x = n in allE, simp add: isS0_def) + done +qed + +lemma uh_h: + assumes uh: "\ haltP (tcontra H) (code (tcontra H))" + shows "haltP (tcontra H) (code (tcontra H))" +proof - + let ?cn = "code (tcontra H)" + have h1: "\ x. \ na nb. steps (Suc 0, Bk\<^bsup>x\<^esup>, Oc\<^bsup>?cn\<^esup> @ Bk # Oc\<^bsup>?cn\<^esup>) + H na = (0, Bk\<^bsup>nb\<^esup>, [Oc, Oc])" + using uh + by(drule_tac x = x in nh_case, simp) + let ?P1 = "\ tp. let (l, r) = tp in (l = [] \ + (r::block list) = Oc\<^bsup>(?cn)\<^esup>)" + let ?Q1 = "\ (l, r).(\ na. l = Bk\<^bsup>na\<^esup> \ r = [Oc, Oc])" + let ?P2 = ?Q1 + let ?Q2 = ?Q1 + let ?P3 = "\ tp. False" + have "?P1 \-> \ tp. (\ stp tp'. steps (Suc 0, tp) ((tcopy |+| H ) |+| dither) + stp = (0, tp') \ ?Q2 tp')" + proof(rule_tac turing_merge.t_merge_halt[of "tcopy |+| H" dither ?P1 ?P2 ?P3 ?P3 + ?Q1 ?Q2], auto simp: turing_merge_def) + show "\stp. case steps (Suc 0, [], Oc\<^bsup>code (tcontra H)\<^esup>) (tcopy |+| H) stp of (s, tp') \ + + s = 0 \ (case tp' of (l, r) \ (\na. l = Bk\<^bsup>na\<^esup>) \ r = [Oc, Oc])" + proof - + let ?Q1 = "\ (l, r).(\ nb. l = Bk\<^bsup>nb\<^esup> \ r = Oc\<^bsup>(?cn)\<^esup> @ Bk # Oc\<^bsup>(?cn)\<^esup>)" + let ?P2 = "?Q1" + let ?Q2 = "\ (l, r).(\ na. l = Bk\<^bsup>na\<^esup> \ r = [Oc, Oc])" + have "?P1 \-> \ tp. (\ stp tp'. steps (Suc 0, tp) (tcopy |+| H ) + stp = (0, tp') \ ?Q2 tp')" + proof(rule_tac turing_merge.t_merge_halt[of tcopy H ?P1 ?P2 ?P3 ?P3 + ?Q1 ?Q2], auto simp: turing_merge_def) + show "\stp. case steps (Suc 0, [], Oc\<^bsup>code (tcontra H)\<^esup>) tcopy stp of (s, tp') \ s = 0 + \ (case tp' of (l, r) \ (\nb. l = Bk\<^bsup>nb\<^esup>) \ r = Oc\<^bsup>code (tcontra H)\<^esup> @ Bk # Oc\<^bsup>code (tcontra H)\<^esup>)" + using tcopy_halt_rs[of "?cn"] + apply(auto) + apply(rule_tac x = stp in exI, simp add: exponent_def) + done + next + fix nb + show "\stp. case steps (Suc 0, Bk\<^bsup>nb\<^esup>, Oc\<^bsup>code (tcontra H)\<^esup> @ Bk # Oc\<^bsup>code (tcontra H)\<^esup>) H stp of + (s, tp') \ s = 0 \ (case tp' of (l, r) \ (\na. l = Bk\<^bsup>na\<^esup>) \ r = [Oc, Oc])" + using h1[of nb] + apply(auto) + apply(rule_tac x = na in exI, auto) + done + next + show "\(l, r). ((\nb. l = Bk\<^bsup>nb\<^esup>) \ r = Oc\<^bsup>code (tcontra H)\<^esup> @ Bk # Oc\<^bsup>code (tcontra H)\<^esup>) \-> + \(l, r). ((\nb. l = Bk\<^bsup>nb\<^esup>) \ r = Oc\<^bsup>code (tcontra H)\<^esup> @ Bk # Oc\<^bsup>code (tcontra H)\<^esup>)" + by(simp add: t_imply_def) + qed + hence "(\ stp tp'. steps (Suc 0, [], Oc\<^bsup>?cn\<^esup>) (tcopy |+| H ) stp = (0, tp') \ ?Q2 tp')" + apply(simp add: t_imply_def) + done + thus "?thesis" + apply(auto) + apply(rule_tac x = stp in exI, auto) + done + qed + next + fix na + show "\stp. case steps (Suc 0, Bk\<^bsup>na\<^esup>, [Oc, Oc]) dither stp of (s, tp') + \ s = 0 \ (case tp' of (l, r) \ (\na. l = Bk\<^bsup>na\<^esup>) \ r = [Oc, Oc])" + using dither_halt_rs[of na] + apply(auto) + apply(rule_tac x = stp in exI, auto) + done + next + show "\(l, r). ((\na. l = Bk\<^bsup>na\<^esup>) \ r = [Oc, Oc]) \-> + (\(l, r). (\na. l = Bk\<^bsup>na\<^esup>) \ r = [Oc, Oc])" + by (simp add: t_imply_def) + qed + hence "\ stp tp'. steps (Suc 0, [], Oc\<^bsup>?cn\<^esup>) ((tcopy |+| H ) |+| dither) + stp = (0, tp') \ ?Q2 tp'" + apply(simp add: t_imply_def) + done + thus "haltP (tcontra H) (code (tcontra H))" + apply(auto simp: haltP_def tcontra_def) + apply(rule_tac x = stp in exI, + rule_tac x = na in exI, + rule_tac x = "Suc (Suc 0)" in exI, + rule_tac x = "0" in exI, simp add: exp_ind_def) + done +qed + +text {* + @{text "False"} is finally derived. +*} + +lemma "False" +using uh_h h_uh +by auto +end + +end +