updated
authorChristian Urban <christian dot urban at kcl dot ac dot uk>
Wed, 23 Jan 2013 20:18:40 +0100
changeset 70 2363eb91d9fd
parent 69 32ec97f94a07
child 71 8c7f10b3da7b
updated
thys/UF.thy
thys/rec_def.thy
thys/recursive.thy
thys/uncomputable.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/UF.thy	Wed Jan 23 20:18:40 2013 +0100
@@ -0,0 +1,4910 @@
+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 {* Univeral Function *}
+
+subsection {* 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 \<equiv>  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 \<Rightarrow> 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 \<Rightarrow> 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, \<dots>, id n k]"}, 
+  the effect of which is to take out the first @{text "Suc k"} 
+  arguments out of the @{text "n"} input arguments.
+  *}
+
+fun get_fstn_args :: "nat \<Rightarrow>  nat \<Rightarrow> 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 \<Rightarrow> 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 \<Rightarrow> nat list \<Rightarrow> 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 (\<lambda> 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 [\<lambda> (r, xs). size r, (\<lambda> (r, xs). last xs)])" 
+    by auto
+next
+  fix n f gs xs x
+  assume "(x::recf) \<in> set gs" 
+  thus "((x, xs), Cn n f gs, xs) \<in> 
+    measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]"
+    by(induct gs, auto)
+next
+  fix n f gs xs x
+  assume "x = map (\<lambda>a. rec_exec a xs) gs"
+    "\<And>x. x \<in> set gs \<Longrightarrow> rec_exec_dom (x, xs)" 
+  thus "((f, x), Cn n f gs, xs) \<in> 
+    measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]"
+    by(auto)
+next
+  fix n f g xs
+  show "((f, butlast xs), Pr n f g, xs) \<in>
+    measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]"
+    by auto
+next
+  fix n f g xs
+  assume "last xs \<noteq> (0::nat)" thus 
+    "((Pr n f g, butlast xs @ [last xs - 1]), Pr n f g, xs) 
+    \<in> measures [\<lambda>(r, xs). size r, \<lambda>(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) \<in> measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]"
+    by auto
+next
+  fix n f xs x
+  show "((f, xs @ [x]), Mn n f, xs) \<in> 
+    measures [\<lambda>(r, xs). size r, \<lambda>(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: "\<And> 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: "\<And> 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: "\<And> 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: "\<And> 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: "\<And> 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: "\<And> 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: 
+  "\<And> 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: "\<And> 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: "\<And> x y. rec_exec rec_conj [x, y] = (if x = 0 \<or> 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: "\<And> x y. rec_exec rec_disj [x, y] = (if x = 0 \<and> 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 \<Rightarrow> nat \<Rightarrow> bool"
+  where
+prime_z[intro]:  "primerec z (Suc 0)" |
+prime_s[intro]:  "primerec s (Suc 0)" |
+prime_id[intro!]: "\<lbrakk>n < m\<rbrakk> \<Longrightarrow> primerec (id m n) m" |
+prime_cn[intro!]: "\<lbrakk>primerec f k; length gs = k; 
+  \<forall> i < length gs. primerec (gs ! i) m; m = n\<rbrakk> 
+  \<Longrightarrow> primerec (Cn n f gs) m" |
+prime_pr[intro!]: "\<lbrakk>primerec f n; 
+  primerec g (Suc (Suc n)); m = Suc n\<rbrakk> 
+  \<Longrightarrow> 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 \<Rightarrow> nat) \<Rightarrow> nat list \<Rightarrow> 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 (\<lambda> (f, xs). last xs))" by auto
+next
+  fix g xs
+  assume "last (xs::nat list) \<noteq> 0"
+  thus "((g, butlast xs @ [last xs - 1]), g, xs)  
+                   \<in> measure (\<lambda>(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]
+
+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)
+
+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: "\<lbrakk>length xs = m; n \<le> m\<rbrakk> \<Longrightarrow> 
+  map (\<lambda> 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 \<Longrightarrow> 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))
+    \<Longrightarrow> 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))
+     \<Longrightarrow> 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, \<dots>, xn, k) = 
+           f(x1, x2, \<dots>, xn, 0) * 
+           f(x1, x2, \<dots>, xn, 1) *
+               \<dots> 
+           f(x1, x2, \<dots>, xn, k)"}
+*}
+fun rec_accum :: "recf \<Rightarrow> 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 \<Rightarrow> nat) \<Rightarrow> nat list \<Rightarrow> 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 (\<lambda> (f, xs). last xs))"
+    by auto
+next
+  fix f xs
+  assume "last xs \<noteq> (0::nat)"
+  thus "((f, butlast xs @ [last xs - 1]), f, xs) \<in> 
+            measure (\<lambda>(f, xs). last xs)"
+    by auto
+qed
+
+lemma rec_accum_Suc_simp_rewrite: 
+  "primerec f (Suc (length xs))
+    \<Longrightarrow> 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))
+     \<Longrightarrow> 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, \<dots>, xn)"} 
+  computes the charactrization function of the following FOL formula:
+  @{text "(\<forall> x \<le> t(x1, x2, \<dots>, xn). (f(x1, x2, \<dots>, xn, x) > 0))"}
+*}
+fun rec_all :: "recf \<Rightarrow> recf \<Rightarrow> 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)) \<Longrightarrow>
+     (rec_exec (rec_accum rf) (xs @ [x]) = 0) = 
+      (\<exists> t \<le> 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: 
+  "\<lbrakk>primerec rf (Suc (length xs));
+    primerec rt (length xs)\<rbrakk>
+  \<Longrightarrow> rec_exec (rec_all rt rf) xs = (if (\<forall> x \<le> (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, \<dots>, xn)"} 
+  computes the charactrization function of the following FOL formula:
+  @{text "(\<exists> x \<le> t(x1, x2, \<dots>, xn). (f(x1, x2, \<dots>, xn, x) > 0))"}
+*}
+fun rec_ex :: "recf \<Rightarrow> recf \<Rightarrow> 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))
+          \<Longrightarrow> (rec_exec (rec_sigma rf) (xs @ [x]) = 0) = 
+                          (\<forall> t \<le> 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:"
+  \<lbrakk>primerec rf (Suc (length xs));
+   primerec rt (length xs)\<rbrakk>
+\<Longrightarrow> (rec_exec (rec_ex rt rf) xs =
+    (if (\<exists> x \<le> (rec_exec rt xs). 0 <rec_exec rf (xs @ [x])) then 1
+     else 0))"
+apply(auto simp: rec_ex.simps rec_exec.simps map_append get_fstn_args_take 
+            split: if_splits)
+apply(drule_tac x = "rec_exec rt xs" in rec_sigma_ex, simp)
+apply(drule_tac x = "rec_exec rt xs" in rec_sigma_ex, simp)
+done
+
+text {*
+  Defintiion of @{text "Min[R]"} on page 77 of Boolos's book.
+*}
+
+fun Minr :: "(nat list \<Rightarrow> bool) \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> nat"
+  where "Minr Rr xs w = (let setx = {y | y. (y \<le> w) \<and> 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 \<le> w \<or> Minr Rr xs w = Suc w"
+apply(auto simp: Minr.simps)
+apply(subgoal_tac "Min {x. x \<le> w \<and> Rr (xs @ [x])} \<le> x")
+apply(erule_tac order_trans, simp)
+apply(rule_tac Min_le, auto)
+done
+
+lemma [simp]: "{x. x \<le> Suc w \<and> Rr (xs @ [x])}
+    = (if Rr (xs @ [Suc w]) then insert (Suc w) 
+                              {x. x \<le> w \<and> Rr (xs @ [x])}
+      else {x. x \<le> w \<and> Rr (xs @ [x])})"
+by(auto, case_tac "x = Suc w", auto)
+
+lemma [simp]: "Minr Rr xs w \<le> w \<Longrightarrow> Minr Rr xs (Suc w) = Minr Rr xs w"
+apply(simp add: Minr.simps, auto)
+apply(case_tac "\<forall>x\<le>w. \<not> Rr (xs @ [x])", auto)
+done
+
+lemma [simp]: "\<forall>x\<le>w. \<not> Rr (xs @ [x]) \<Longrightarrow>  
+                           {x. x \<le> w \<and> Rr (xs @ [x])} = {} "
+by auto
+
+lemma [simp]: "\<lbrakk>Minr Rr xs w = Suc w; Rr (xs @ [Suc w])\<rbrakk> \<Longrightarrow> 
+                                       Minr Rr xs (Suc w) = Suc w"
+apply(simp add: Minr.simps)
+apply(case_tac "\<forall>x\<le>w. \<not> Rr (xs @ [x])", auto)
+done
+ 
+lemma [simp]: "\<lbrakk>Minr Rr xs w = Suc w; \<not> Rr (xs @ [Suc w])\<rbrakk> \<Longrightarrow> 
+                                   Minr Rr xs (Suc w) = Suc (Suc w)"
+apply(simp add: Minr.simps)
+apply(case_tac "\<forall>x\<le>w. \<not> Rr (xs @ [x])", auto)
+apply(subgoal_tac "Min {x. x \<le> w \<and> Rr (xs @ [x])} \<in> 
+                                {x. x \<le> w \<and> 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 \<le> 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 \<Rightarrow> 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 \<Longrightarrow> 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 \<Longrightarrow>
+                       primerec rt n \<and> 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]: "\<lbrakk>i < n; n \<le> m\<rbrakk> \<Longrightarrow> 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]: "\<lbrakk>primerec rf n; n \<ge> Suc (Suc 0)\<rbrakk>   \<Longrightarrow> 
+                        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: 
+  "\<lbrakk>primerec rt n; primerec rf (Suc n); n > 0\<rbrakk> \<Longrightarrow> 
+                                 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]) \<Longrightarrow> 
+                   Min {x. x = (0::nat) \<and> 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]: "\<lbrakk>\<not> Min {uu. uu \<le> w \<and> 0 < rec_exec rf (xs @ [uu])} \<le> w;
+       x \<le> w; 0 < rec_exec rf (xs @ [x])\<rbrakk>
+      \<Longrightarrow>  False"
+apply(subgoal_tac "finite {uu. uu \<le> w \<and> 0 < rec_exec rf (xs @ [uu])}")
+apply(subgoal_tac "{uu. uu \<le> w \<and> 0 < rec_exec rf (xs @ [uu])} \<noteq> {}")
+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 (\<lambda>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]))) \<and>
+        primerec ?rt (length (xs @ [0]))"
+    apply(auto simp: prrf nth_append)+
+    done
+  show "Sigma (rec_exec (rec_all ?rt ?rf)) (xs @ [0])
+       = Minr (\<lambda>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 (\<lambda>args. 0 < rec_exec rf args) xs w"
+  have prrf: "primerec ?rf (Suc (length (xs @ [Suc w]))) \<and>
+        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 (\<lambda>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: "
+  \<lbrakk>primerec rf (Suc (length xs))\<rbrakk> 
+     \<Longrightarrow> rec_exec (rec_Minr rf) (xs @ [w]) = 
+            Minr (\<lambda> 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 (\<lambda> 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: 
+  "\<And>x y. rec_exec rec_le [x, y] = (if (x \<le> 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 \<Rightarrow> bool) \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "Maxr Rr xs w = (let setx = {y. y \<le> w \<and> 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 \<Rightarrow> 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 \<Longrightarrow> (take n ys @ [ys ! n, ys ! n]) =  
+                                                  ys @ [ys ! n]"
+apply(simp)
+apply(subgoal_tac "\<exists> 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 Max.insert)
+apply(rule_tac Max_eqI, auto)
+done
+
+lemma [simp]: "min (Suc n) n = n" by simp
+
+lemma Sigma_0: "\<forall> i \<le> n. (f (xs @ [i]) = 0) \<Longrightarrow> 
+                              Sigma f (xs @ [n]) = 0"
+apply(induct n, simp add: Sigma.simps)
+apply(simp add: Sigma_Suc_simp_rewrite)
+done
+  
+lemma [elim]: "\<forall>k<Suc w. f (xs @ [k]) = Suc 0
+        \<Longrightarrow> 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: "\<lbrakk>\<forall> k < ma. f (xs @ [k]) = 1;
+        \<forall> k \<ge> ma. f (xs @ [k]) = 0; ma \<le> w\<rbrakk>
+    \<Longrightarrow> 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 (\<lambda>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: "\<forall>x\<le>w. rec_exec rf (xs @ [x]) = 0"
+    have "primerec ?rf (Suc (length (xs @ [w, i]))) \<and> 
+          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 \<le> w" "0 < rec_exec rf (xs @ [x])"
+    hence "\<exists> ma. Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])} = ma"
+      by auto
+    from this obtain ma where k1: 
+      "Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])} = ma" ..
+    hence k2: "ma \<le> w \<and> 0 < rec_exec rf (xs @ [ma])"
+      using h
+      apply(subgoal_tac
+        "Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])} \<in>  {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])}")
+      apply(erule_tac CollectE, simp)
+      apply(rule_tac Max_in, auto)
+      done
+    hence k3: "\<forall> k < ma. (rec_exec ?notrq (xs @ [w, k]) = 1)"
+      apply(auto simp: nth_append)
+      apply(subgoal_tac "primerec ?rf (Suc (length (xs @ [w, k]))) \<and> 
+        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: "\<forall> k \<ge> ma. (rec_exec ?notrq (xs @ [w, k]) = 0)"
+      apply(auto)
+      apply(subgoal_tac "primerec ?rf (Suc (length (xs @ [w, k]))) \<and> 
+        primerec ?rt (length (xs @ [w, k]))")
+      apply(auto simp: rec_exec.simps all_lemma get_fstn_args_take nth_append)
+      apply(subgoal_tac "x \<le> Max {y. y \<le> w \<and> 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 \<le> w \<and> 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 (\<lambda> 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 (\<lambda>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 (\<lambda>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 \<Rightarrow> nat"
+  where
+  "quo [x, y] = (let Rr = 
+                         (\<lambda> zs. ((zs ! (Suc 0) * zs ! (Suc (Suc 0))
+                                 \<le> zs ! 0) \<and> zs ! Suc 0 \<noteq> (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 \<Longrightarrow> 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 \<le> x"  "y > 0"
+  hence "(y * ya) div y \<le> x div y"
+    by(insert div_le_mono[of "y * ya" x y], simp)
+  from this and h show "ya \<le> x div y" by simp
+next
+  fix xa
+  show "y * (x div y) \<le> 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: 
+  "\<And> x y. rec_exec rec_noteq [x, y] = 
+               (if x \<noteq> 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 (\<lambda> 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 (\<lambda> args. if rec_exec ?rR args = 0 then False
+                           else True) [x, y] x" 
+    by simp
+  have g2: "Maxr (\<lambda> 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: "\<And> 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 \<Rightarrow> nat"
+type_synonym rtype = "nat list \<Rightarrow> bool"
+
+text {*
+  The specifation of the mutli-way branching statement on
+  page 79 of Boolos's book.
+  *}
+fun Embranch :: "(ftype * rtype) list \<Rightarrow> nat list \<Rightarrow> 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 \<Rightarrow> nat \<Rightarrow> 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 \<Rightarrow> 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: 
+  "\<lbrakk>\<forall> j < length rcs. rec_exec (rcs ! j) xs = 0;
+    length rgs = length rcs;  
+  rcs \<noteq> []; 
+  list_all (\<lambda> rf. primerec rf (length xs)) (rgs @ rcs)\<rbrakk>  \<Longrightarrow> 
+  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: 
+    "\<And>rgs. \<lbrakk>\<forall>j<length rcs. rec_exec (rcs ! j) xs = 0; 
+             length rgs = length rcs; rcs \<noteq> []; 
+            list_all (\<lambda>rf. primerec rf (length xs)) (rgs @ rcs)\<rbrakk> \<Longrightarrow> 
+                      rec_exec (rec_embranch (zip rgs rcs)) xs = 0"
+  and h:  "\<forall>j<length (a # rcs). rec_exec ((a # rcs) ! j) xs = 0"
+  "length rgs = length (a # rcs)" 
+    "a # rcs \<noteq> []" 
+    "list_all (\<lambda>rf. primerec rf (length xs)) (rgs @ a # rcs)"
+    "rgs = aa # list"
+  have g: "rcs \<noteq> [] \<Longrightarrow> 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 \<noteq> []"
+    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: "\<lbrakk>rec_exec aa xs = 0; zip rgs list \<noteq> []; 
+       list_all (\<lambda> rf. primerec rf (length xs)) ([a, aa] @ rgs @ list)\<rbrakk>
+       \<Longrightarrow> 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: "\<lbrakk>length xs = k; length ys = k; zip xs ys = []\<rbrakk> \<Longrightarrow> xs = [] \<and> ys = []"
+apply(case_tac xs, simp, simp)
+apply(case_tac ys, simp, simp)
+done
+
+lemma zip_null_gr: "\<lbrakk>length xs = k; length ys = k; zip xs ys \<noteq> []\<rbrakk> \<Longrightarrow> 0 < k"
+apply(case_tac xs, simp, simp)
+done
+
+lemma Embranch_0:  
+  "\<lbrakk>length rgs = k; length rcs = k; k > 0; 
+  \<forall> j < k. rec_exec (rcs ! j) xs = 0\<rbrakk> \<Longrightarrow>
+  Embranch (zip (map rec_exec rgs) (map (\<lambda>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: 
+    "\<And>rcs k. \<lbrakk>length rgs = k; length rcs = k; 0 < k; \<forall>j<k. rec_exec (rcs ! j) xs = 0\<rbrakk> 
+    \<Longrightarrow> Embranch (zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) rcs)) xs = 0"
+  and h: "Suc (length rgs) = k" "length rcs = k"
+    "\<forall>j<k. rec_exec (rcs ! j) xs = 0"
+  from h show  
+    "Embranch (zip (rec_exec a # map rec_exec rgs) 
+           (map (\<lambda>r 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: 
+  "(\<exists> i < n. (rec_exec (rcs ! i) xs = 1 \<and> (\<forall> j < n. j \<noteq> i \<longrightarrow> 
+                                      rec_exec (rcs ! j) xs = 0)))"
+  and prime_all: "list_all (\<lambda> rf. primerec rf (length xs)) (rgs @ rcs)"
+  shows "rec_exec (rec_embranch (zip rgs rcs)) xs =
+                  Embranch (zip (map rec_exec rgs) 
+                     (map (\<lambda> 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: 
+    "\<And>rcs n. \<lbrakk>length rgs = n; length rcs = n; 0 < n;
+    \<exists>i<n. rec_exec (rcs ! i) xs = 1 \<and> (\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec (rcs ! j) xs = 0);
+    list_all (\<lambda>rf. primerec rf (length xs)) (rgs @ rcs)\<rbrakk>
+    \<Longrightarrow> rec_exec (rec_embranch (zip rgs rcs)) xs =
+    Embranch (zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) rcs)) xs"
+  and h: "length (a # rgs) = n" "length (rcs::recf list) = n" "0 < n"
+         " \<exists>i<n. rec_exec (rcs ! i) xs = 1 \<and> 
+         (\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec (rcs ! j) xs = 0)" 
+    "list_all (\<lambda>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 (\<lambda>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 = [] \<and> list = []", simp add: Embranch.simps rec_exec.simps rec_embranch.simps)
+    apply(rule_tac  zip_null_iff, simp, simp, simp)
+  proof -
+    fix aa list
+    assume g:
+      "Suc (length rgs) = n" "Suc (length list) = n" 
+      "\<exists>i<n. rec_exec ((aa # list) ! i) xs = Suc 0 \<and> 
+          (\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec ((aa # list) ! j) xs = 0)"
+      "primerec a (length xs) \<and> 
+      list_all (\<lambda>rf. primerec rf (length xs)) rgs \<and>
+      primerec aa (length xs) \<and> 
+      list_all (\<lambda>rf. primerec rf (length xs)) list"
+      "rec_exec aa xs = 0" "rcs = aa # list" "zip rgs list \<noteq> []"
+    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, \<lambda>args. 0 < rec_exec aa args) # 
+           zip (map rec_exec rgs) (map (\<lambda>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"
+      "\<exists>i<n. rec_exec ((aa # list) ! i) xs = Suc 0 \<and> 
+      (\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec ((aa # list) ! j) xs = 0)"
+      "primerec a (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) rgs \<and>
+      primerec aa (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) list"
+    "rcs = aa # list" "rec_exec aa xs \<noteq> 0" "zip rgs list = []"
+    thus "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = 
+        Embranch ((rec_exec a, \<lambda>args. 0 < rec_exec aa args) # 
+       zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) list)) xs"
+      apply(subgoal_tac "rgs = [] \<and> 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"
+      "\<exists>i<n. rec_exec ((aa # list) ! i) xs = Suc 0 \<and>  
+           (\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec ((aa # list) ! j) xs = 0)"
+      "primerec a (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) rgs
+      \<and> primerec aa (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) list"
+      "rcs = aa # list" "rec_exec aa xs \<noteq> 0" "zip rgs list \<noteq> []"
+    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 " \<forall>j<length list. rec_exec (list ! j) xs = 0"
+          using g
+          apply(auto)
+          apply(case_tac i, simp)
+          apply(erule_tac x = "Suc j" in allE, simp)
+          apply(simp)
+          apply(erule_tac x = 0 in allE, simp)
+          done
+      next
+        show "length rgs = length list"
+          using g
+          apply(case_tac n, simp, simp)
+          done
+      next
+        show "list \<noteq> []"
+          using g
+          apply(case_tac list, simp, simp)
+          done
+      next
+        show "list_all (\<lambda>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 (\<lambda>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, \<lambda>args. 0 < rec_exec aa args) #
+           zip (map rec_exec rgs) (map (\<lambda>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 \<Rightarrow> bool"
+  where
+  "Prime x = (1 < x \<and> (\<forall> u < x. (\<forall> v < x. u * v \<noteq> x)))"
+
+declare Prime.simps [simp del]
+
+lemma primerec_all1: 
+  "primerec (rec_all rt rf) n \<Longrightarrow> primerec rt n"
+  by (simp add: primerec_all)
+
+lemma primerec_all2: "primerec (rec_all rt rf) n \<Longrightarrow> 
+  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]))]"
+
+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 (\<forall>w\<le>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 (\<forall>k\<le>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 \<longrightarrow>  (rec_exec (rec_all ?rt2 ?rf2) [x] = 0 \<longrightarrow> 
+    \<not> Prime x) \<and>
+     (0 < rec_exec (rec_all ?rt2 ?rf2) [x] \<longrightarrow> Prime x)) \<and>
+    (\<not> Suc 0 < x \<longrightarrow> \<not> Prime x \<and> (rec_exec (rec_all ?rt2 ?rf2) [x] = 0
+    \<longrightarrow> \<not> Prime x))"
+    apply(auto simp:rec_exec.simps)
+    apply(simp add: exec_tmp rec_exec.simps)
+  proof -
+    assume "\<forall>k\<le>x - Suc 0. (0::nat) < (if \<forall>w\<le>x - Suc 0. 
+           0 < (if k * w \<noteq> 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 "\<not> 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 \<le> 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 \<le> 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 \<Rightarrow> 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 (\<lambda> 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 \<Rightarrow> nat"
+  where
+  "Np x = Min {y. y \<le> Suc (x!) \<and> x < y \<and> 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: 
+"\<lbrakk>\<not> Prime x; x > Suc 0\<rbrakk> \<Longrightarrow> (\<exists> u > Suc 0. (\<exists> v > Suc 0. u * v = x))"
+ by(auto simp: Prime.simps)
+
+lemma divsor_prime_ex: "\<lbrakk>\<not> Prime x; x > Suc 0\<rbrakk> \<Longrightarrow> 
+  \<exists> p. Prime p \<and> p dvd x"
+apply(induct x rule: wf_induct[where r = "measure (\<lambda> 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: "\<lbrakk>0 < q; q \<le> n\<rbrakk> \<Longrightarrow> q dvd n!"
+apply(induct n, simp)
+apply(case_tac "q \<le> 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: "\<lbrakk>Suc 0 < q; q dvd n!; q \<le> n\<rbrakk> \<Longrightarrow> \<not> q dvd Suc (n!)"
+proof(auto simp: dvd_def)
+  fix k ka
+  assume h1: "Suc 0 < q" "q \<le> 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 "\<exists>d. d > 0 \<and>  ka = d + k"  
+    by(rule_tac x = "ka - k" in exI, simp)
+  from this obtain d where "d > 0 \<and> ka = d + k" ..
+  from h2 and this and h1 show "False"
+    by(simp add: add_mult_distrib2)
+qed
+    
+lemma prime_ex: "\<exists> p. n < p \<and> p \<le> Suc (n!) \<and> Prime p"
+proof(cases "Prime (n! + 1)")
+  case True thus "?thesis" 
+    by(rule_tac x = "Suc (n!)" in exI, simp)
+next
+  assume h: "\<not> Prime (n! + 1)"  
+  hence "\<exists> p. Prime p \<and> p dvd (n! + 1)"
+    by(erule_tac divsor_prime_ex, auto)
+  from this obtain q where k: "Prime q \<and> 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: "\<not> 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 "\<not> 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!]: "\<lbrakk>i < Suc (Suc 0); 
+  primerec (ys ! 0) n; primerec (ys ! 1) n\<rbrakk> \<Longrightarrow> 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 = "\<lambda> zs. zs ! 0 < zs ! 1 \<and> Prime (zs ! 1)"
+  have g1: "rec_exec (rec_Minr ?rr) ([x] @ [Suc (x!)]) = 
+    Minr (\<lambda> 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 (\<lambda> 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 \<longrightarrow> (x < uu \<longrightarrow> uu \<le> Suc (x!)) \<and> x < uu) \<and> Prime uu}
+    = {y. y \<le> Suc (x!) \<and> x < y \<and> Prime y}", auto)
+    done
+  from g1 and g2 show "rec_exec (rec_Minr ?rr) ([x, Suc (x!)]) = Np x"
+    by simp
+qed
+
+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 \<Rightarrow> 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
+
+fun loR :: "nat list \<Rightarrow> 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 \<Rightarrow> nat \<Rightarrow> nat"
+  where 
+  "lo x y  = (if x > 1 \<and> y > 1 \<and> {u. loR [x, y, u]} \<noteq> {} then Max {u. loR [x, y, u]}
+                                                         else 0)"
+
+declare lo.simps[simp del]
+
+lemma [elim]: "primerec rf n \<Longrightarrow> n > 0"
+apply(induct rule: primerec.induct, auto)
+done
+
+lemma primerec_sigma[intro!]:  
+  "\<lbrakk>n > Suc 0; primerec rf n\<rbrakk> \<Longrightarrow> 
+  primerec (rec_sigma rf) n"
+apply(simp add: rec_sigma.simps)
+apply(auto, auto simp: nth_append)
+done
+
+lemma [intro!]:  "\<lbrakk>primerec rf n; n > 0\<rbrakk> \<Longrightarrow> 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!]: 
+  "\<lbrakk>i < Suc (Suc (Suc (0::nat))); primerec (ys ! 0) n;
+  primerec (ys ! 1) n;  
+  primerec (ys ! 2) n\<rbrakk> \<Longrightarrow> 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:
+  "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow>  
+        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 (\<lambda> 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 (\<lambda> 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 \<and> loR [0, y, ya]} = 0"
+apply(rule_tac Max_eqI, auto simp: loR.simps)
+done
+
+lemma [simp]: "Suc 0 < y \<Longrightarrow> Suc (Suc 0) < y * y"
+apply(induct y, simp)
+apply(case_tac y, simp, simp)
+done
+
+lemma less_mult: "\<lbrakk>x > 0; y > Suc 0\<rbrakk> \<Longrightarrow> x < y * x"
+apply(case_tac y, simp, simp)
+done
+
+lemma x_less_exp: "\<lbrakk>y > Suc 0\<rbrakk> \<Longrightarrow> 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 \<noteq> (0::nat) \<Longrightarrow> x \<le> x * y"  
+  by(induct y, simp, simp)
+
+lemma uplimit_loR:  "\<lbrakk>Suc 0 < x; Suc 0 < y; loR [x, y, xa]\<rbrakk> \<Longrightarrow> 
+  xa \<le> x"
+apply(simp add: loR.simps)
+apply(rule_tac classical, auto)
+apply(subgoal_tac "xa < y^xa")
+apply(subgoal_tac "y^xa \<le> y^xa * q", simp)
+apply(rule_tac le_mult, case_tac q, simp, simp)
+apply(rule_tac x_less_exp, simp)
+done
+
+lemma [simp]: "\<lbrakk>xa \<le> x; loR [x, y, xa]; Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow>
+  {u. loR [x, y, u]} = {ya. ya \<le> x \<and> loR [x, y, ya]}"
+apply(rule_tac Collect_cong, auto)
+apply(erule_tac uplimit_loR, simp, simp)
+done
+
+lemma Maxr_lo: "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow>
+  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': "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> 
+  rec_exec rec_lo [x, y] = lo x y"
+by(simp add: Maxr_lo  rec_lo_Maxr_lor)
+
+lemma lo_lemma'': "\<lbrakk>\<not> Suc 0 < x\<rbrakk> \<Longrightarrow> 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''': "\<lbrakk>\<not> Suc 0 < y\<rbrakk> \<Longrightarrow> 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 \<and> Suc 0 < y")
+apply(auto simp: lo_lemma' lo_lemma'' lo_lemma''')
+done
+
+fun lgR :: "nat list \<Rightarrow> bool"
+  where
+  "lgR [x, y, u] = (y^u \<le> 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 \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "lg x y = (if x > 1 \<and> y > 1 \<and> {u. lgR [x, y, u]} \<noteq> {} 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: "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> 
+                      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 ((\<lambda> 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 ((\<lambda> 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]: "\<lbrakk>Suc 0 < y; lgR [x, y, xa]\<rbrakk> \<Longrightarrow> xa \<le> x"
+apply(simp add: lgR.simps)
+apply(subgoal_tac "y^xa > xa", simp)
+apply(erule x_less_exp)
+done
+
+lemma [simp]: "\<lbrakk>Suc 0 < x; Suc 0 < y; lgR [x, y, xa]\<rbrakk> \<Longrightarrow>
+           {u. lgR [x, y, u]} =  {ya. ya \<le> x \<and> lgR [x, y, ya]}"
+apply(rule_tac Collect_cong, auto)
+done
+
+lemma maxr_lg: "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> 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': "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> rec_exec rec_lg [x, y] = lg x y"
+apply(simp add: maxr_lg lg_maxr)
+done
+
+lemma lg_lemma'': "\<not> Suc 0 < x \<Longrightarrow> 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''': "\<not> Suc 0 < y \<Longrightarrow> 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 \<and> 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 \<Rightarrow> nat \<Rightarrow> 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)
+
+
+subsection {* The construction of 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 \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "listsum2 xs 0 = 0"
+| "listsum2 xs (Suc n) = listsum2 xs n + xs ! n"
+
+fun rec_listsum2 :: "nat \<Rightarrow> nat \<Rightarrow> 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: "\<lbrakk>length xs = vl; n \<le> vl\<rbrakk> \<Longrightarrow> 
+      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 \<Rightarrow> nat \<Rightarrow> 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 \<Rightarrow> nat \<Rightarrow> 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: "\<lbrakk>length xs = vl; n \<le> vl\<rbrakk> \<Longrightarrow> 
+  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 \<Rightarrow> nat"
+  where
+  "strt xs = (let ys = map Suc xs in 
+              strt' ys (length ys))"
+
+fun rec_map :: "recf \<Rightarrow> nat \<Rightarrow> recf list"
+  where
+  "rec_map rf vl = map (\<lambda> i. Cn vl rf [id vl (i)]) [0..<vl]"
+
+text {*
+  @{text "rec_strt"} is the recursive function used to implement @{text "strt"}.
+  *}
+fun rec_strt :: "nat \<Rightarrow> recf"
+  where
+  "rec_strt vl = Cn vl (rec_strt' vl vl) (rec_map s vl)"
+
+lemma map_s_lemma: "length xs = vl \<Longrightarrow> 
+  map ((\<lambda>a. rec_exec a xs) \<circ> (\<lambda>i. Cn vl s [recf.id vl i]))
+  [0..<vl]
+        = map Suc xs"
+apply(induct vl arbitrary: xs, simp, auto simp: rec_exec.simps)
+apply(subgoal_tac "\<exists> ys y. xs = ys @ [y]", auto)
+proof -
+  fix ys y
+  assume ind: "\<And>xs. length xs = length (ys::nat list) \<Longrightarrow>
+      map ((\<lambda>a. rec_exec a xs) \<circ> (\<lambda>i. Cn (length ys) s 
+        [recf.id (length ys) (i)])) [0..<length ys] = map Suc xs"
+  show
+    "map ((\<lambda>a. rec_exec a (ys @ [y])) \<circ> (\<lambda>i. Cn (Suc (length ys)) s 
+  [recf.id (Suc (length ys)) (i)])) [0..<length ys] = map Suc ys"
+  proof -
+    have "map ((\<lambda>a. rec_exec a ys) \<circ> (\<lambda>i. Cn (length ys) s
+        [recf.id (length ys) (i)])) [0..<length ys] = map Suc ys"
+      apply(rule_tac ind, simp)
+      done
+    moreover have
+      "map ((\<lambda>a. rec_exec a (ys @ [y])) \<circ> (\<lambda>i. Cn (Suc (length ys)) s
+           [recf.id (Suc (length ys)) (i)])) [0..<length ys]
+         = map ((\<lambda>a. rec_exec a ys) \<circ> (\<lambda>i. Cn (length ys) s 
+                 [recf.id (length ys) (i)])) [0..<length ys]"
+      apply(rule_tac map_ext, auto simp: rec_exec.simps nth_append)
+      done
+    ultimately show "?thesis"
+      by simp
+  qed
+next
+  fix vl xs
+  assume "length xs = Suc vl"
+  thus "\<exists>ys y. xs = ys @ [y]"
+    apply(rule_tac x = "butlast xs" in exI, rule_tac x = "last xs" in exI)
+    apply(subgoal_tac "xs \<noteq> []", auto)
+    done
+qed
+
+text {*
+  The correctness of @{text "rec_strt"}.
+  *}
+lemma strt_lemma: "length xs = vl \<Longrightarrow> 
+  rec_exec (rec_strt vl) xs = strt xs"
+apply(simp add: strt.simps rec_exec.simps strt'_lemma)
+apply(subgoal_tac "(map ((\<lambda>a. rec_exec a xs) \<circ> (\<lambda>i. Cn vl s [recf.id vl (i)])) [0..<vl])
+                  = map Suc xs", auto)
+apply(rule map_s_lemma, simp)
+done
+
+text {*
+  The @{text "scan"} function on page 90 of B book.
+  *}
+fun scan :: "nat \<Rightarrow> 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 \<Rightarrow> nat"
+  where
+  "newleft0 [p, r] = p"
+
+definition rec_newleft0 :: "recf"
+  where
+  "rec_newleft0 = id 2 0"
+
+fun newrgt0 :: "nat list \<Rightarrow> 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 \<Rightarrow> nat"
+  where
+  "newleft1 [p, r] = p"
+
+definition rec_newleft1 :: "recf"
+  where
+  "rec_newleft1 = id 2 0"
+
+fun newrgt1 :: "nat list \<Rightarrow> 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 \<Rightarrow> 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 \<Rightarrow> 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 \<Rightarrow> 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 \<Rightarrow> 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 \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "newleft p r a = (if a = 0 \<or> 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: 
+  "\<lbrakk>i < Suc (Suc (Suc (Suc 0))); i = 0 \<Longrightarrow>  P i;
+    i = 1 \<Longrightarrow> P i; i =2 \<Longrightarrow> P i; 
+    i =3 \<Longrightarrow> P i\<rbrakk> \<Longrightarrow> 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 (\<lambda>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 \<or> 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 k2: "Embranch (zip (map rec_exec ?rgs) (map (\<lambda>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 \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 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: 
+  "\<lbrakk>i < Suc (Suc (Suc (Suc (Suc 0)))); i = 0 \<Longrightarrow> P 0;
+  i = 1 \<Longrightarrow> P 1; i = 2 \<Longrightarrow> P 2; i = 3 \<Longrightarrow> P 3; i = 4 \<Longrightarrow> P 4\<rbrakk> \<Longrightarrow> 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, \<lambda> zs. zs ! 1]"
+  let ?r0 = "\<lambda> zs. zs ! 2 = 0"
+  let ?r1 = "\<lambda> zs. zs ! 2 = 1"
+  let ?r2 = "\<lambda> zs. zs ! 2 = 2"
+  let ?r3 = "\<lambda> zs. zs ! 2 = 3"
+  let ?r4 = "\<lambda> zs. zs ! 2 > 3"
+  let ?gs = "map (\<lambda> g. (\<lambda> 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 (\<lambda>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 (\<lambda>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 \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "actn m q r = (if q \<noteq> 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)
+
+fun newstat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "newstat m q r = (if q \<noteq> 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 \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 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 \<Rightarrow> nat"
+  where
+  "left c = lo c (Pi 0)"
+
+fun stat :: "nat \<Rightarrow> nat"
+  where
+  "stat c = lo c (Pi 1)"
+
+fun rght :: "nat \<Rightarrow> nat"
+  where
+  "rght c = lo c (Pi 2)"
+
+thm Prime.simps
+
+fun inpt :: "nat \<Rightarrow> nat list \<Rightarrow> nat"
+  where
+  "inpt m xs = trpl 0 1 (strt xs)"
+
+fun newconf :: "nat \<Rightarrow> nat \<Rightarrow> 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 \<Rightarrow> 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 (\<lambda> i. id vl (i)) [1..<vl])]"
+
+lemma left_lemma: "rec_exec rec_left [c] = left c"
+by(simp add: rec_exec.simps rec_left_def left.simps lo_lemma)
+      
+lemma right_lemma: "rec_exec rec_right [c] = rght c"
+by(simp add: rec_exec.simps rec_right_def rght.simps lo_lemma)
+
+lemma stat_lemma: "rec_exec rec_stat [c] = stat c"
+by(simp add: rec_exec.simps rec_stat_def stat.simps lo_lemma)
+ 
+declare rec_strt.simps[simp del] strt.simps[simp del]
+
+lemma map_cons_eq: 
+  "(map ((\<lambda>a. rec_exec a (m # xs)) \<circ> 
+    (\<lambda>i. recf.id (Suc (length xs)) (i))) 
+          [Suc 0..<Suc (length xs)])
+        = map (\<lambda> i. xs ! (i - 1)) [Suc 0..<Suc (length xs)]"
+apply(rule map_ext, auto)
+apply(auto simp: rec_exec.simps nth_append nth_Cons split: nat.split)
+done
+
+lemma list_map_eq: 
+  "vl = length (xs::nat list) \<Longrightarrow> map (\<lambda> i. xs ! (i - 1))
+                                          [Suc 0..<Suc vl] = xs"
+apply(induct vl arbitrary: xs, simp)
+apply(subgoal_tac "\<exists> ys y. xs = ys @ [y]", auto)
+proof -
+  fix ys y
+  assume ind: 
+    "\<And>xs. length (ys::nat list) = length (xs::nat list) \<Longrightarrow>
+            map (\<lambda>i. xs ! (i - Suc 0)) [Suc 0..<length xs] @
+                                [xs ! (length xs - Suc 0)] = xs"
+  and h: "Suc 0 \<le> length (ys::nat list)"
+  have "map (\<lambda>i. ys ! (i - Suc 0)) [Suc 0..<length ys] @ 
+                                   [ys ! (length ys - Suc 0)] = ys"
+    apply(rule_tac ind, simp)
+    done
+  moreover have 
+    "map (\<lambda>i. (ys @ [y]) ! (i - Suc 0)) [Suc 0..<length ys]
+      = map (\<lambda>i. ys ! (i - Suc 0)) [Suc 0..<length ys]"
+    apply(rule map_ext)
+    using h
+    apply(auto simp: nth_append)
+    done
+  ultimately show "map (\<lambda>i. (ys @ [y]) ! (i - Suc 0)) 
+        [Suc 0..<length ys] @ [(ys @ [y]) ! (length ys - Suc 0)] = ys"
+    apply(simp del: map_eq_conv add: nth_append, auto)
+    using h
+    apply(simp)
+    done
+next
+  fix vl xs
+  assume "Suc vl = length (xs::nat list)"
+  thus "\<exists>ys y. xs = ys @ [y]"
+    apply(rule_tac x = "butlast xs" in exI, 
+          rule_tac x = "last xs" in exI)
+    apply(case_tac "xs \<noteq> []", auto)
+    done
+qed
+
+lemma [elim]: 
+  "Suc 0 \<le> length xs \<Longrightarrow> 
+     (map ((\<lambda>a. rec_exec a (m # xs)) \<circ> 
+         (\<lambda>i. recf.id (Suc (length xs)) (i))) 
+             [Suc 0..<length xs] @ [(m # xs) ! length xs]) = xs"
+using map_cons_eq[of m xs]
+apply(simp del: map_eq_conv add: rec_exec.simps)
+using list_map_eq[of "length xs" xs]
+apply(simp)
+done
+
+    
+lemma inpt_lemma:
+  "\<lbrakk>Suc (length xs) = vl\<rbrakk> \<Longrightarrow> 
+            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 ((\<lambda>a. rec_exec a (m # xs)) \<circ> 
+          (\<lambda>i. recf.id (Suc (length xs)) (i))) 
+            [Suc 0..<length xs] @ [(m # xs) ! length xs]) = xs", simp)
+apply(auto, case_tac xs, auto)
+done
+
+definition rec_newconf:: "recf"
+  where
+  "rec_newconf = 
+    Cn 2 rec_trpl 
+        [Cn 2 rec_newleft [Cn 2 rec_left [id 2 1], 
+                           Cn 2 rec_right [id 2 1], 
+                           Cn 2 rec_actn [id 2 0, 
+                                          Cn 2 rec_stat [id 2 1], 
+                           Cn 2 rec_right [id 2 1]]],
+          Cn 2 rec_newstat [id 2 0, 
+                            Cn 2 rec_stat [id 2 1], 
+                            Cn 2 rec_right [id 2 1]],
+           Cn 2 rec_newrght [Cn 2 rec_left [id 2 1], 
+                             Cn 2 rec_right [id 2 1], 
+                             Cn 2 rec_actn [id 2 0, 
+                                   Cn 2 rec_stat [id 2 1], 
+                             Cn 2 rec_right [id 2 1]]]]"
+
+lemma newconf_lemma: "rec_exec rec_newconf [m ,c] = newconf m c"
+by(auto simp: rec_newconf_def rec_exec.simps 
+              trpl_lemma newleft_lemma left_lemma
+              right_lemma stat_lemma newrght_lemma actn_lemma 
+               newstat_lemma stat_lemma newconf.simps)
+
+declare newconf_lemma[simp]
+
+text {*
+  @{text "conf m r k"} computes the TM configuration after @{text "k"} steps of execution
+  of TM coded as @{text "m"} starting from the initial configuration where the left number equals @{text "0"}, 
+  right number equals @{text "r"}. 
+  *}
+fun conf :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 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 \<Rightarrow> bool"
+  where
+  "NSTD c = (stat c \<noteq> 0 \<or> left c \<noteq> 0 \<or> 
+             rght c \<noteq> 2^(lg (rght c + 1) 2) - 1 \<or> 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 \<or>
+                   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) \<Longrightarrow> 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 \<Longrightarrow> (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 \<Rightarrow> 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{* 
+  @{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 \<Rightarrow> nat  \<Rightarrow> nat \<Rightarrow> 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]
+
+lemma primerec_not0: "primerec f n \<Longrightarrow> n > 0"
+by(induct f n rule: primerec.induct, auto)
+
+lemma [elim]: "primerec f 0 \<Longrightarrow> RR"
+apply(drule_tac primerec_not0, simp)
+done
+
+lemma [simp]: "length xs = Suc n \<Longrightarrow> length (butlast xs) = n"
+apply(subgoal_tac "\<exists> 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) 
+           \<Longrightarrow> 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:
+    "\<And>x rs. \<lbrakk>x \<in> set gs; primerec x (length xs)\<rbrakk> \<Longrightarrow>
+                rec_calc_rel x xs rs = (rec_exec x xs = rs)"
+    and ind2: 
+    "\<And>x rs. \<lbrakk>x = map (\<lambda>a. rec_exec a xs) gs; 
+             primerec f (length gs)\<rbrakk> \<Longrightarrow> 
+            rec_calc_rel f (map (\<lambda>a. rec_exec a xs) gs) rs = 
+           (rec_exec f (map (\<lambda>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:"\<forall>k<length gs. rec_calc_rel (gs ! k) xs (ys ! k)"
+      and g2: "length ys = length gs"
+      and g3: "rec_calc_rel f ys rs"
+    have "rec_calc_rel f (map (\<lambda>a. rec_exec a xs) gs) rs =
+                  (rec_exec f (map (\<lambda>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 (\<lambda>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 (\<lambda>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 (\<lambda>a. rec_exec a xs) gs))"
+      apply(rule_tac rs = "(map (\<lambda>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)" 
+             "\<forall>i<length gs. primerec (gs ! i) (length xs)"
+      thus "rec_calc_rel (gs ! k) xs (rec_exec (gs ! k) xs)"
+        using ind1[of "gs!k" "(rec_exec (gs ! k) xs)"]
+        by(simp)
+    next
+      assume "primerec f (length gs)" 
+             "\<forall>i<length gs. primerec (gs ! i) (length xs)"
+      thus "rec_calc_rel f (map (\<lambda>a. rec_exec a xs) gs) 
+        (rec_exec f (map (\<lambda>a. rec_exec a xs) gs))"
+        using ind2[of "(map (\<lambda>a. rec_exec a xs) gs)" 
+                   "(rec_exec f (map (\<lambda>a. rec_exec a xs) gs))"]
+        by simp
+    qed
+  qed
+next
+  fix n f g xs rs
+  assume ind1: 
+    "\<And>rs. \<lbrakk>last xs = 0; primerec f (length xs - Suc 0)\<rbrakk> 
+    \<Longrightarrow> rec_calc_rel f (butlast xs) rs = 
+                     (rec_exec f (butlast xs) = rs)"
+  and ind2 : 
+    "\<And>rs. \<lbrakk>0 < last xs; 
+           primerec (Pr n f g) (Suc (length xs - Suc 0))\<rbrakk> \<Longrightarrow>
+           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: 
+    "\<And>rs. \<lbrakk>0 < last xs; primerec g (Suc (Suc (length xs - Suc 0)))\<rbrakk>
+       \<Longrightarrow> 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!]: "\<lbrakk>0 < vl; n \<le> vl\<rbrakk> \<Longrightarrow> 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]: "\<lbrakk>0 < vl; n \<le> vl\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> 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 \<Longrightarrow> primerec ((map (\<lambda>i. recf.id (Suc vl) (i)) 
+        [Suc 0..<vl] @ [recf.id (Suc vl) (vl)]) ! i) (Suc vl)"
+apply(induct i, auto simp: nth_append)
+done
+
+lemma [intro]: "primerec rec_newleft0 ((Suc (Suc 0)))"
+apply(simp add: rec_newleft_def rec_embranch.simps 
+                Let_def arity.simps rec_newleft0_def
+                rec_newleft1_def rec_newleft2_def rec_newleft3_def)
+apply(tactic {* resolve_tac [@{thm prime_cn}, 
+    @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+
+done
+
+lemma [intro]: "primerec rec_newleft1 ((Suc (Suc 0)))"
+apply(simp add: rec_newleft_def rec_embranch.simps 
+                Let_def arity.simps rec_newleft0_def
+                rec_newleft1_def rec_newleft2_def rec_newleft3_def)
+apply(tactic {* resolve_tac [@{thm prime_cn}, 
+    @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+
+done
+
+lemma [intro]: "primerec rec_newleft2 ((Suc (Suc 0)))"
+apply(simp add: rec_newleft_def rec_embranch.simps 
+                Let_def arity.simps rec_newleft0_def
+                rec_newleft1_def rec_newleft2_def rec_newleft3_def)
+apply(tactic {* resolve_tac [@{thm prime_cn}, 
+    @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+
+done
+
+lemma [intro]: "primerec rec_newleft3 ((Suc (Suc 0)))"
+apply(simp add: rec_newleft_def rec_embranch.simps 
+                Let_def arity.simps rec_newleft0_def
+                rec_newleft1_def rec_newleft2_def rec_newleft3_def)
+apply(tactic {* resolve_tac [@{thm prime_cn}, 
+    @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+
+done
+
+lemma [intro]: "primerec rec_newleft (Suc (Suc (Suc 0)))"
+apply(simp add: rec_newleft_def rec_embranch.simps 
+                Let_def arity.simps)
+apply(rule_tac prime_cn, auto+)
+done
+
+lemma [intro]: "primerec rec_left (Suc 0)"
+apply(simp add: rec_left_def rec_lo_def rec_entry_def Let_def)
+apply(tactic {* resolve_tac [@{thm prime_cn}, 
+    @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+
+done
+
+lemma [intro]: "primerec rec_actn (Suc (Suc (Suc 0)))"
+apply(simp add: rec_left_def rec_lo_def rec_entry_def
+                Let_def rec_actn_def)
+apply(tactic {* resolve_tac [@{thm prime_cn}, 
+    @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+
+done
+
+lemma [intro]: "primerec rec_stat (Suc 0)"
+apply(simp add: rec_left_def rec_lo_def rec_entry_def Let_def 
+                rec_actn_def rec_stat_def)
+apply(tactic {* resolve_tac [@{thm prime_cn}, 
+    @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+
+done
+
+lemma [intro]: "primerec rec_newstat (Suc (Suc (Suc 0)))"
+apply(simp add: rec_left_def rec_lo_def rec_entry_def 
+                Let_def rec_actn_def rec_stat_def rec_newstat_def)
+apply(tactic {* resolve_tac [@{thm prime_cn}, 
+    @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+
+done
+
+lemma [intro]: "primerec rec_newrght (Suc (Suc (Suc 0)))"
+apply(simp add: rec_newrght_def rec_embranch.simps
+                Let_def arity.simps rec_newrgt0_def 
+                rec_newrgt1_def rec_newrgt2_def rec_newrgt3_def)
+apply(tactic {* resolve_tac [@{thm prime_cn}, 
+    @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+
+done
+
+lemma [intro]: "primerec rec_newconf (Suc (Suc 0))"
+apply(simp add: rec_newconf_def)
+apply(tactic {* resolve_tac [@{thm prime_cn}, 
+    @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+
+done
+
+lemma [intro]: "0 < vl \<Longrightarrow> 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 \<and> 
+  (\<forall> t'< t. 
+      (\<exists> y. rec_calc_rel rec_nonstop [m, r, t'] y \<and>
+            y \<noteq> 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 \<and> 
+           (\<forall> t'< t. (\<exists> y. rec_exec rec_nonstop [m, r, t'] = y
+                    \<and> y \<noteq> 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 \<Rightarrow> 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 \<Longrightarrow> (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]: 
+  "\<lbrakk>ys \<noteq> [];  k < length ys\<rbrakk> \<Longrightarrow>
+  (get_fstn_args (length ys) (length ys) ! k) = 
+                                  id (length ys) (k)"
+by(erule_tac get_fstn_args_nth)
+
+lemma calc_rel_get_pren: 
+  "\<lbrakk>ys \<noteq> [];  k < length ys\<rbrakk> \<Longrightarrow> 
+  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]:
+  "\<lbrakk>xs \<noteq> []; k < Suc (length xs)\<rbrakk> \<Longrightarrow> 
+  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 \<Longrightarrow>
+  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 \<or>  k = Suc 0 \<or> 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: 
+  "\<forall> t. \<not> rec_calc_rel rec_halt [m, r] t \<Longrightarrow> 
+                \<forall> rs. \<not> 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:
+    "\<forall>t. \<not> 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: "\<forall>k<Suc (Suc (Suc 0)). rec_calc_rel ([recf.id (Suc (Suc 0)) 0, 
+          recf.id (Suc (Suc 0)) (Suc 0), rec_halt] ! k) [m, r] (rsc ! k)"
+  have "rec_calc_rel (rec_halt ) [m, r]
+                              (rsc ! (Suc (Suc 0)))"
+    using g
+    apply(erule_tac x = "(Suc (Suc 0))" in allE)
+    apply(simp add:nth_append)
+    done
+  thus "False"
+    using h
+    apply(erule_tac x = "ysb ! (Suc (Suc 0))" in allE, simp)
+    done
+qed
+
+
+subsection {* Coding function of TMs *}
+
+text {*
+  The purpose of this section is to get the coding function of Turing Machine, which is 
+  going to be named @{text "code"}.
+  *}
+
+fun bl2nat :: "cell list \<Rightarrow> nat \<Rightarrow> 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 :: "cell list \<Rightarrow> nat"
+  where
+  "bl2wc xs = bl2nat xs 0"
+
+fun trpl_code :: "config \<Rightarrow> 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 :: "action \<Rightarrow> 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 \<Rightarrow> action"
+  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 :: "cell \<Rightarrow> nat"
+  where
+  "block_map Bk = 0"
+| "block_map Oc = 1"
+
+fun godel_code' :: "nat list \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "godel_code' [] n = 1"
+| "godel_code' (x#xs) n = (Pi n)^x * godel_code' xs (Suc n) "
+
+fun godel_code :: "nat list \<Rightarrow> nat"
+  where
+  "godel_code xs = (let lh = length xs in 
+                   2^lh * (godel_code' xs (Suc 0)))"
+
+fun modify_tprog :: "instr list \<Rightarrow> 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 :: "instr list \<Rightarrow> nat"
+  where 
+  "code tp = (let nl = modify_tprog tp in 
+              godel_code nl)"
+
+subsection {* Relating interperter functions to the execution of TMs *}
+
+lemma [simp]: "bl2wc [] = 0" by(simp add: bl2wc.simps bl2nat.simps)
+term trpl
+
+lemma [simp]: "\<lbrakk>fetch tp 0 b = (nact, ns)\<rbrakk> \<Longrightarrow> action_map nact = 4"
+apply(simp add: fetch.simps)
+done
+
+lemma Pi_gr_1[simp]: "Pi n > Suc 0"
+proof(induct n, auto simp: Pi.simps Np.simps)
+  fix n
+  let ?setx = "{y. y \<le> Suc (Pi n!) \<and> Pi n < y \<and> Prime y}"
+  have "finite ?setx" by auto
+  moreover have "?setx \<noteq> {}"
+    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]: 
+  "\<lbrakk>i < length nl; \<not> Suc 0 < godel_code nl\<rbrakk> \<Longrightarrow> nl ! i = 0"
+using godel_code_great[of nl] godel_code_eq_1[of nl]
+apply(simp)
+done
+
+term set_of
+lemma prime_coprime: "\<lbrakk>Prime x; Prime y; x\<noteq>y\<rbrakk> \<Longrightarrow> 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: "\<forall>u<d * ka. \<forall>v<d * ka. u * v \<noteq> d * ka" 
+    and case_k: "\<forall>u<d * k. \<forall>v<d * k. u * v \<noteq> d * k"
+    and h: "(0::nat) < d" "d \<noteq> Suc 0" "Suc 0 < d * ka" 
+           "ka \<noteq> k" "Suc 0 < d * k"
+  from h have "k > Suc 0 \<or> 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 \<and> 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 \<and> 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 \<le> Suc (Pi i!) \<and> Pi i < y \<and> Prime y}"
+  have "finite ?setx" by simp
+  moreover have "?setx \<noteq> {}"
+    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 \<Longrightarrow> Pi i < Pi j"
+proof(induct j, simp)
+  fix j
+  assume ind: "i < j \<Longrightarrow> 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" "\<not> 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 \<noteq> j \<Longrightarrow> Pi i \<noteq> 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 \<le> Suc (Pi n!) \<and> Pi n < y \<and> Prime y}"
+  show "Prime (Min ?setx)"
+  proof -
+    have "finite ?setx" by simp
+    moreover have "?setx \<noteq> {}" 
+      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 \<noteq> j \<Longrightarrow> 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 \<noteq> j \<Longrightarrow> coprime ((Pi i)^m) ((Pi j)^n)"
+by(rule_tac coprime_exp2_nat, erule_tac Pi_coprime)
+
+lemma coprime_dvd_mult_nat2: "\<lbrakk>coprime (k::nat) n; k dvd n * m\<rbrakk> \<Longrightarrow> 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 \<noteq> [] \<Longrightarrow> godel_code' xs (Suc j) = 
+  godel_code' (butlast xs) (Suc j) * Pi (length xs + j)^(last xs)"
+apply(subgoal_tac "\<exists> 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 \<noteq> 0"
+apply(induct xs, auto simp: godel_code'.simps)
+done
+
+lemma godel_code_append_cons: 
+  "length xs = i \<Longrightarrow> 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: 
+    "\<And>xs i y ys. \<lbrakk>x = i; length xs = i\<rbrakk> \<Longrightarrow> 
+       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 \<le> i \<Longrightarrow> 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: 
+    "\<And>ps. \<lbrakk>x = length ps; length ps \<le> i\<rbrakk> \<Longrightarrow>
+                  coprime (Pi (Suc i)) (godel_code' ps (Suc 0))"
+  and h: "Suc x = length ps"
+          "length (ps::nat list) \<le> 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 \<Longrightarrow> coprime (Pi i) (godel_code' ps j)"
+proof(induct "length ps" arbitrary: ps, simp add: godel_code'.simps)
+  fix x ps
+  assume ind: 
+    "\<And>ps. \<lbrakk>x = length ps; i < j\<rbrakk> \<Longrightarrow> 
+                    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: "\<not> ia < godel_code' nl (Suc 0)"
+  from g1 have "Pi (Suc i)^ia \<le> 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 \<Longrightarrow>  nl ! i  \<in> {u. Pi (Suc i) ^ u dvd
+                                     godel_code' nl (Suc 0)}"
+proof -
+ assume h: "i<length nl"
+  hence "godel_code' (take i nl@(nl!i)#drop (Suc i) nl) (Suc 0)
+           = godel_code' (take i nl) (Suc 0) *  Pi (Suc i)^(nl!i) *
+                               godel_code' (drop (Suc i) nl) (i + 2)"
+    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"]
+    apply(simp)
+    done
+  ultimately  show 
+    "nl ! i \<in> {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}"
+    by(simp)
+qed
+     
+lemma godel_code'_get_nth:
+  "i < length nl \<Longrightarrow> 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 \<in> {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\<le>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 \<le>  Pi (Suc i) ^ nl ! i "
+      apply(rule_tac dvd_imp_le, auto)
+      done
+    thus "y \<le> nl ! i"
+      apply(rule_tac power_le_imp_le_exp, auto)
+      done
+  qed
+next
+  assume h: "i<length nl"
+
+  thus "nl ! i \<in> {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 \<Longrightarrow> 
+           Max {u. Pi (Suc i) ^ u dvd godel_code nl} = nl ! i"
+by(simp add: godel_code.simps godel_code'_get_nth)
+
+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: "\<lbrakk>a > Suc 0; a ^ y dvd a ^ l\<rbrakk> \<Longrightarrow> y \<le> l"
+apply(case_tac "y \<le> l", simp, simp)
+apply(subgoal_tac "\<exists> d. y = l + d", auto simp: power_add)
+apply(rule_tac x = "y - l" in exI, simp)
+done
+
+
+lemma [elim]: "Pi n = 0 \<Longrightarrow> RR"
+  using Pi_not_0[of n] by simp
+
+lemma [elim]: "Pi n = Suc 0 \<Longrightarrow> RR"
+  using Pi_gr_1[of n] by simp
+
+lemma finite_power_dvd:
+  "\<lbrakk>(a::nat) > Suc 0; y \<noteq> 0\<rbrakk> \<Longrightarrow> 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: "\<lbrakk>m \<noteq> n; m \<noteq> k; k \<noteq> n\<rbrakk> \<Longrightarrow> 
+  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 \<noteq> n" "m \<noteq> k" "k \<noteq> 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 \<in> ?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 \<le> (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 \<in> ?setx" by simp
+  qed
+qed  
+
+lemma conf_decode2: 
+  "\<lbrakk>m \<noteq> n; m \<noteq> k; n \<noteq> k; 
+  \<not> Suc 0 < Pi m ^ l * Pi n ^ st * Pi k ^ r\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> 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 \<Longrightarrow> 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: 
+  "\<lbrakk>st \<le> length tp div 2; st > 0\<rbrakk> \<Longrightarrow>
+  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: 
+    "\<And>tp. \<lbrakk>st \<le> length tp div 2; 0 < st\<rbrakk> \<Longrightarrow> 
+     modify_tprog tp ! (4 * (st - Suc 0)) =
+               action_map (fst ((tp::instr list) ! (2 * (st - Suc 0))))"
+  and h: "Suc st \<le> length (tp::instr list) 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 \<noteq> 0"
+    hence "\<exists> 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'::instr list) ! (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: 
+  "\<lbrakk>st \<le> length tp div 2; st > 0\<rbrakk> \<Longrightarrow> 
+       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: 
+    "\<And>tp. \<lbrakk>st \<le> length tp div 2; 0 < st\<rbrakk> \<Longrightarrow>  
+       modify_tprog tp ! Suc (Suc (4 * (st - Suc 0))) = 
+          action_map (fst (tp ! Suc (2 * (st - Suc 0))))"
+  and h: "Suc st \<le> length (tp::instr list) 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 \<noteq> 0"
+    hence "\<exists> 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:
+  "\<lbrakk>st \<le> length tp div 2; st > 0; b = 1 \<or> b = 0\<rbrakk> \<Longrightarrow> 
+      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: 
+  "\<lbrakk>block_map b = scan r; fetch tp st b = (nact, ns);
+   st \<le> length tp div 2\<rbrakk> \<Longrightarrow> 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 \<le> 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)
+   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(case_tac st, simp_all add: fetch.simps nth_of.simps)
+    apply(case_tac b, auto simp: block_map.simps nth_of.simps fetch.simps 
+                    split: if_splits)
+    apply(case_tac "r mod 2", simp, simp)
+    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) \<Longrightarrow> ns = 0"
+by(simp add: fetch.simps)
+
+lemma Five_Suc: "5 = Suc 4" by simp
+
+lemma modify_tprog_fetch_state:
+  "\<lbrakk>st \<le> length tp div 2; st > 0; b = 1 \<or> b = 0\<rbrakk> \<Longrightarrow> 
+     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: 
+    "\<And>tp. \<lbrakk>st \<le> length tp div 2; 0 < st; b = 1 \<or> b = 0\<rbrakk> \<Longrightarrow> 
+    modify_tprog tp ! Suc (4 * (st - Suc 0) + 2 * b) =
+                             snd (tp ! (2 * (st - Suc 0) + b))"
+  and h:
+    "Suc st \<le> length (tp::instr list) div 2" 
+    "0 < Suc st" 
+    "b = 1 \<or> 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 \<noteq> 0"
+    hence "\<exists> 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:
+  "\<lbrakk>block_map b = scan r; 
+  fetch tp st b = (nact, ns);
+  st \<le> length tp div 2\<rbrakk> \<Longrightarrow> 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 \<le> 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)
+   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(case_tac st, simp)
+    apply(case_tac b, auto simp: block_map.simps nth_of.simps
+                                 fetch.simps
+                                 split: if_splits)
+    apply(subgoal_tac "(2 * (Suc nat - r mod 2) + r mod 2) = 
+                       (2 * nat + r mod 2)", simp)
+    by (metis diff_Suc_Suc minus_nat.diff_0)
+  ultimately show "Entry (godel_code (modify_tprog tp)) (?i)
+           = ns" 
+    by simp
+qed
+
+
+lemma [intro!]: 
+  "\<lbrakk>a = a'; b = b'; c = c'\<rbrakk> \<Longrightarrow> 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: "\<And>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]: "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]:
+  "bl2wc (Oc # tl c) = Suc (bl2wc c) - bl2wc c mod 2 "
+apply(case_tac c, case_tac [2] a, simp)
+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 \<noteq> [] \<Longrightarrow> 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 \<noteq> [] \<Longrightarrow> 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]: "\<lbrakk>b \<noteq> []\<rbrakk> \<Longrightarrow> 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]
+
+text {*
+  The lemma relates the one step execution of TMs with the interpreter function @{text "rec_newconf"}.
+  *}
+lemma rec_t_eq_step: 
+  "(\<lambda> (s, l, r). s \<le> length tp div 2) c \<Longrightarrow>
+  trpl_code (step0 c tp) = 
+  rec_exec rec_newconf [code tp, trpl_code c]"
+  apply(cases c, simp)
+proof(case_tac "fetch tp a (read ca)",
+    simp add: newconf.simps trpl_code.simps step.simps)
+  fix a b ca aa ba
+  assume h: "(a::nat) \<le> length tp div 2" 
+    "fetch tp a (read ca) = (aa, ba)"
+  moreover hence "actn (code tp) a (bl2wc ca) = action_map aa"
+    apply(rule_tac b = "read ca" 
+          in fetch_action_eq, auto)
+    apply(case_tac "hd ca", auto)
+    apply(case_tac [!] ca, auto)
+    done
+  moreover from h have "(newstat (code tp) a (bl2wc ca)) = ba"
+    apply(rule_tac b = "read ca" 
+          in fetch_state_eq, auto split: list.splits)
+    apply(case_tac "hd ca", auto)
+    apply(case_tac [!] ca, auto)
+    done
+  ultimately show 
+    "trpl_code (ba, update aa (b, ca)) =
+          trpl (newleft (bl2wc b) (bl2wc ca) (actn (code tp) a (bl2wc ca))) 
+    (newstat (code tp) a (bl2wc ca)) (newrght (bl2wc b) (bl2wc ca) (actn (code tp) a (bl2wc ca)))"
+    apply(case_tac aa)
+    apply(auto simp: trpl_code.simps 
+         newleft.simps newrght.simps split: action.splits)
+    done
+qed
+
+lemma [simp]: "bl2nat (Oc # Oc\<up>x) 0 = (2 * 2 ^ x - Suc 0)"
+apply(induct x)
+apply(simp add: bl2nat.simps)
+apply(simp add: bl2nat.simps bl2nat_double exp_ind)
+done
+
+lemma [simp]: "bl2nat (Oc\<up>y) 0 = 2^y - Suc 0"
+apply(induct y, auto simp: bl2nat.simps bl2nat_double)
+apply(case_tac "(2::nat)^y", auto)
+done
+
+lemma [simp]: "bl2nat (Bk\<up>l) n = 0"
+apply(induct l, auto simp: bl2nat.simps bl2nat_double exp_ind)
+done
+
+lemma bl2nat_cons_bk: "bl2nat (ks @ [Bk]) 0 = bl2nat ks 0"
+apply(induct ks, auto simp: bl2nat.simps)
+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)
+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: 
+    "\<And>xs ys. x = length xs \<Longrightarrow> 
+             bl2nat (xs @ ys) 0 = bl2nat xs 0 + bl2nat ys (length xs)"
+  and h: "Suc x = length (xs::cell list)"
+  have "\<exists> 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 \<noteq> 0 \<Longrightarrow> 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: "\<lbrakk>a = b; c = d\<rbrakk> \<Longrightarrow> a - c = b - d"
+by auto
+
+lemma tape_of_nat_list_butlast_last:
+  "ys \<noteq> [] \<Longrightarrow> <ys @ [y]> = <ys> @ Bk # Oc\<up>Suc y"
+apply(induct ys, simp, simp)
+apply(case_tac "ys = []", simp add: tape_of_nl_abv 
+                                    tape_of_nat_list.simps)
+apply(simp add: tape_of_nl_cons)
+done
+
+lemma listsum2_append:
+  "\<lbrakk>n \<le> length xs\<rbrakk> \<Longrightarrow> listsum2 (xs @ ys) n = listsum2 xs n"
+apply(induct n)
+apply(auto simp: listsum2.simps nth_append)
+done
+
+lemma strt'_append:  
+  "\<lbrakk>n \<le> length xs\<rbrakk> \<Longrightarrow> 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: 
+    "\<And> xs ys. n \<le> length xs \<Longrightarrow> strt' xs n = strt' (xs @ ys) n"
+    and h: "Suc n \<le> 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: 
+  "\<lbrakk>length (ys::nat list) = k\<rbrakk>
+       \<Longrightarrow> length (<ys>) = listsum2 (map Suc ys) k + k - 1"
+apply(induct k arbitrary: ys, simp_all add: listsum2.simps)
+apply(subgoal_tac "\<exists> xs x. ys = xs @ [x]", auto)
+proof -
+  fix xs x
+  assume ind: "\<And>ys. length ys = length xs \<Longrightarrow> length (<ys>) 
+    = listsum2 (map Suc ys) (length xs) + 
+      length (xs::nat list) - Suc 0"
+  have "length (<xs>) 
+    = listsum2 (map Suc xs) (length xs) + length xs - Suc 0"
+    apply(rule_tac ind, simp)
+    done
+  thus "length (<xs @ [x]>) =
+    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 "\<exists>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 (steps0 (Suc 0, Bk\<up>l, <lm>) tp 0) = 
+    rec_exec rec_conf [code tp, bl2wc (<lm>), 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 state_in_range_step
+  : "\<lbrakk>a \<le> length A div 2; step0 (a, b, c) A = (st, l, r); tm_wf (A,0)\<rbrakk>
+  \<Longrightarrow> st \<le> length A div 2"
+apply(simp add: step.simps fetch.simps tm_wf.simps 
+  split: if_splits list.splits)
+apply(case_tac [!] a, auto simp: list_all_length 
+  fetch.simps nth_of.simps)
+apply(erule_tac x = "A ! (2*nat) " in ballE, auto)
+apply(case_tac "hd c", auto simp: fetch.simps nth_of.simps)
+apply(erule_tac x = "A !(2 * nat)" in ballE, auto)
+apply(erule_tac x = "A !Suc (2 * nat)" in ballE, auto)
+done
+
+lemma state_in_range: "\<lbrakk>steps0 (Suc 0, tp) A stp = (st, l, r); tm_wf (A, 0)\<rbrakk>
+  \<Longrightarrow> st \<le> length A div 2"
+proof(induct stp arbitrary: st l r)
+  case 0 thus "?case" by(auto simp: tm_wf.simps steps.simps)
+next
+  fix stp st l r
+  assume ind: "\<And>st l r. \<lbrakk>steps0 (Suc 0, tp) A stp = (st, l, r); tm_wf (A, 0)\<rbrakk> \<Longrightarrow> st \<le> length A div 2"
+  and h1: "steps0 (Suc 0, tp) A (Suc stp) = (st, l, r)"
+  and h2: "tm_wf (A,0::nat)"
+  from h1 h2 show "st \<le> length A div 2"
+  proof(simp add: step_red, cases "(steps0 (Suc 0, tp) A stp)", simp)
+    fix a b c 
+    assume h3: "step0 (a, b, c) A = (st, l, r)"
+    and h4: "steps0 (Suc 0, tp) A stp = (a, b, c)"
+    have "a \<le> length A div 2"
+      using h2 h4
+      by(rule_tac l = b and r = c in ind, auto)
+    thus "?thesis"
+      using h3 h2
+      apply(erule_tac state_in_range_step, simp_all)
+      done
+  qed
+qed
+
+lemma rec_t_eq_steps:
+  "tm_wf (tp,0) \<Longrightarrow>
+  trpl_code (steps0 (Suc 0, Bk\<up>l, <lm>) tp stp) = 
+  rec_exec rec_conf [code tp, bl2wc (<lm>), stp]"
+proof(induct stp)
+  case 0 thus "?case" by(simp)
+next
+  case (Suc n) thus "?case"
+  proof -
+    assume ind: 
+      "tm_wf (tp,0) \<Longrightarrow> trpl_code (steps0 (Suc 0, Bk\<up> l, <lm>) tp n) 
+      = rec_exec rec_conf [code tp, bl2wc (<lm>), n]"
+      and h: "tm_wf (tp, 0)"
+    show 
+      "trpl_code (steps0 (Suc 0, Bk\<up> l, <lm>) tp (Suc n)) =
+      rec_exec rec_conf [code tp, bl2wc (<lm>), Suc n]"
+    proof(case_tac "steps0 (Suc 0, Bk\<up> l, <lm>) tp  n", 
+        simp only: step_red conf_lemma conf.simps)
+      fix a b c
+      assume g: "steps0 (Suc 0, Bk\<up> l, <lm>) tp n = (a, b, c) "
+      hence "conf (code tp) (bl2wc (<lm>)) n= trpl_code (a, b, c)"
+        using ind h
+        apply(simp add: conf_lemma)
+        done
+      moreover hence 
+        "trpl_code (step0 (a, b, c) tp) = 
+        rec_exec rec_newconf [code tp, trpl_code (a, b, c)]"
+        apply(rule_tac rec_t_eq_step)
+        using h g
+        apply(simp add: state_in_range)
+        done
+      ultimately show 
+        "trpl_code (step0 (a, b, c) tp) =
+            newconf (code tp) (conf (code tp) (bl2wc (<lm>)) n)"
+        by(simp add: newconf_lemma)
+    qed
+  qed
+qed
+
+lemma [simp]: "bl2wc (Bk\<up> m) = 0"
+apply(induct m)
+apply(simp, simp)
+done
+
+lemma [simp]: "bl2wc (Oc\<up> rs@Bk\<up> n) = bl2wc (Oc\<up> rs)"
+apply(induct rs, simp, 
+  simp add: bl2wc.simps bl2nat.simps bl2nat_double)
+done
+
+lemma lg_power: "x > Suc 0 \<Longrightarrow> lg (x ^ rs) x = rs"
+proof(simp add: lg.simps, auto)
+  fix xa
+  assume h: "Suc 0 < x"
+  show "Max {ya. ya \<le> x ^ rs \<and> 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 "\<not> Suc 0 < x ^ rs" "Suc 0 < x" 
+  thus "rs = 0"
+    apply(case_tac "x ^ rs", simp, simp)
+    done
+next
+  assume "Suc 0 < x" "\<forall>xa. \<not> 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"}.
+  *}
+
+declare tm_wf.simps[simp del]
+
+lemma nonstop_t_eq: 
+  "\<lbrakk>steps0 (Suc 0, Bk\<up>l, <lm>) tp stp = (0, Bk\<up> m, Oc\<up> rs @ Bk\<up> n); 
+   tm_wf (tp, 0); 
+  rs > 0\<rbrakk> 
+  \<Longrightarrow> rec_exec rec_nonstop [code tp, bl2wc (<lm>), stp] = 0"
+proof(simp add: nonstop_lemma nonstop.simps nstd.simps)
+  assume h: "steps0 (Suc 0, Bk\<up>l, <lm>) tp stp = (0, Bk\<up> m, Oc\<up> rs @ Bk\<up> n)"
+  and tc_t: "tm_wf (tp, 0)" "rs > 0"
+  have g: "rec_exec rec_conf [code tp,  bl2wc (<lm>), stp] =
+                                        trpl_code (0, Bk\<up> m, Oc\<up> rs@Bk\<up> n)"
+    using rec_t_eq_steps[of tp l lm stp] tc_t h
+    by(simp)
+  thus "\<not> NSTD (conf (code tp) (bl2wc (<lm>)) stp)" 
+  proof(auto simp: NSTD.simps)
+    show "stat (conf (code tp) (bl2wc (<lm>)) stp) = 0"
+      using g
+      by(auto simp: conf_lemma trpl_code.simps)
+  next
+    show "left (conf (code tp) (bl2wc (<lm>)) stp) = 0"
+      using g
+      by(simp add: conf_lemma trpl_code.simps)
+  next
+    show "rght (conf (code tp) (bl2wc (<lm>)) stp) = 
+           2 ^ lg (Suc (rght (conf (code tp) (bl2wc (<lm>)) stp))) 2 - Suc 0"
+    using g h
+    proof(simp add: conf_lemma trpl_code.simps)
+      have "2 ^ lg (Suc (bl2wc (Oc\<up> rs))) 2 = Suc (bl2wc (Oc\<up> rs))"
+        apply(simp add: bl2wc.simps lg_power)
+        done
+      thus "bl2wc (Oc\<up> rs) = 2 ^ lg (Suc (bl2wc (Oc\<up> rs))) 2 - Suc 0"
+        apply(simp)
+        done
+    qed
+  next
+    show "0 < rght (conf (code tp) (bl2wc (<lm>)) 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 step_red[simp del]
+
+lemma halt_least_step: 
+  "\<lbrakk>steps0 (Suc 0, Bk\<up>l, <lm>) tp stp = 
+       (0, Bk\<up> m, Oc\<up>rs @ Bk\<up>n); 
+    tm_wf (tp, 0); 
+    0<rs\<rbrakk> \<Longrightarrow>
+    \<exists> stp. (nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and>
+       (\<forall> stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp'))"
+proof(induct stp, simp add: steps.simps, simp)
+  fix stp
+  assume ind: 
+    "steps0 (Suc 0, Bk\<up> l, <lm>) tp stp = (0, Bk\<up> m, Oc\<up> rs @ Bk\<up> n) \<Longrightarrow> 
+    \<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and> 
+          (\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp')"
+  and h: 
+    "steps0 (Suc 0, Bk\<up> l, <lm>) tp (Suc stp) = (0, Bk\<up> m, Oc\<up> rs @ Bk\<up> n)"
+    "tm_wf (tp, 0::nat)" 
+    "0 < rs"
+  from h show 
+    "\<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 
+    \<and> (\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp')"
+  proof(simp add: step_red, 
+      case_tac "steps0 (Suc 0, Bk\<up> l, <lm>) tp stp", simp, 
+       case_tac a, simp add: step_0)
+    assume "steps0 (Suc 0, Bk\<up> l, <lm>) tp stp = (0, Bk\<up> m, Oc\<up> rs @ Bk\<up> n)"
+    thus "\<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and> 
+      (\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp')"
+      apply(erule_tac ind)
+      done
+  next
+    fix a b c nat
+    assume "steps0 (Suc 0, Bk\<up> l, <lm>) tp stp = (a, b, c)"
+      "a = Suc nat"
+    thus "\<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and> 
+      (\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> 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:"steps0 (Suc 0, Bk\<up> l, <lm>) tp stp = (Suc nat, b, c)" 
+        "nonstop (code tp) (bl2wc (<lm>)) stp' = 0"
+      thus  "Suc stp \<le> stp'"
+      proof(case_tac "Suc stp \<le> stp'", simp, simp)
+        assume "\<not> Suc stp \<le> stp'"
+        hence "stp' \<le> stp" by simp
+        hence "\<not> is_final (steps0 (Suc 0, Bk\<up> l, <lm>) tp stp')"
+          using g
+          apply(case_tac "steps0 (Suc 0, Bk\<up> l, <lm>) tp stp'",auto, simp)
+          apply(subgoal_tac "\<exists> n. stp = stp' + n", auto simp: steps_add steps_0)
+          apply(case_tac a, simp_all add: steps.simps)
+          apply(rule_tac x = "stp - stp'"  in exI, simp)
+          done         
+        hence "nonstop (code tp) (bl2wc (<lm>)) stp' = 1"
+        proof(case_tac "steps0 (Suc 0, Bk\<up> l, <lm>) tp stp'",
+            simp add: nonstop.simps)
+          fix a b c
+          assume k: 
+            "0 < a" "steps0 (Suc 0, Bk\<up> l, <lm>) tp stp' = (a, b, c)"
+          thus " NSTD (conf (code tp) (bl2wc (<lm>)) 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 (<lm>)) 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 (<lm>)) stp')" by simp
+          qed
+        qed
+        thus "False" using g by simp
+      qed
+    qed
+  qed
+qed    
+
+lemma conf_trpl_ex: "\<exists> p q r. conf m (bl2wc (<lm>)) 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 (<lm>)" in exI)
+apply(simp)
+done
+  
+lemma nonstop_rgt_ex: 
+  "nonstop m (bl2wc (<lm>)) stpa = 0 \<Longrightarrow> \<exists> r. conf m (bl2wc (<lm>)) 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 \<Longrightarrow> 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 \<in> {u. x ^ u dvd x ^ r}"
+  thus "y \<le> r"
+    apply(case_tac "y\<le> r", simp)
+    apply(subgoal_tac "\<exists> d. y = r + d")
+    apply(auto simp: power_add)
+    apply(rule_tac x = "y - r" in exI, simp)
+    done
+next
+  show "r \<in> {u. x ^ u dvd x ^ r}" by simp
+qed  
+
+lemma lo_power: "x > Suc 0 \<Longrightarrow> 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  \<Longrightarrow>
+  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:
+  "\<lbrakk>nonstop m (bl2wc (<lm>)) stpa = 0\<rbrakk> \<Longrightarrow> 
+  conf m (bl2wc (<lm>)) stpa = conf m (bl2wc (<lm>)) (stpa + n)"
+apply(drule_tac nonstop_rgt_ex, auto simp: conf_keep)
+done
+
+lemma halt_state_keep: 
+  "\<lbrakk>nonstop m (bl2wc (<lm>)) stpa = 0; nonstop m (bl2wc (<lm>)) stpb = 0\<rbrakk> \<Longrightarrow>
+  conf m (bl2wc (<lm>)) stpa = conf m (bl2wc (<lm>)) 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
+
+text {*
+  The correntess of @{text "rec_F"} which relates the interpreter function @{text "rec_F"} with the
+  execution of of TMs.
+  *}
+
+lemma F_correct: 
+  "\<lbrakk>steps0 (Suc 0, Bk\<up>l, <lm>) tp stp = (0, Bk\<up>m, Oc\<up>rs@Bk\<up>n); 
+    tm_wf (tp,0); 0<rs\<rbrakk>
+   \<Longrightarrow> rec_calc_rel rec_F [code tp, (bl2wc (<lm>))] (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 (<lm>)) stpa = 0" 
+    "\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stpa \<le> stp'" 
+    "nonstop (code tp) (bl2wc (<lm>)) stp = 0" 
+    "trpl_code (0, Bk\<up> m, Oc\<up> rs @ Bk\<up> n) = conf (code tp) (bl2wc (<lm>)) stp"
+    "steps0 (Suc 0, Bk\<up> l, <lm>) tp stp = (0, Bk\<up> m, Oc\<up> rs @ Bk\<up> n)"
+  hence g1: "conf (code tp) (bl2wc (<lm>)) stpa = trpl_code (0, Bk\<up> m, Oc\<up> rs @ Bk\<up>n)"
+    using halt_state_keep[of "code tp" lm stpa stp]
+    by(simp)
+  moreover have g2:
+    "rec_calc_rel rec_halt [code tp, (bl2wc (<lm>))] stpa"
+    using h
+    apply(simp add: halt_lemma nonstop_lemma, auto)
+    done
+  show  
+    "rec_calc_rel rec_F [code tp, (bl2wc (<lm>))] (rs - Suc 0)"
+  proof -
+    have 
+      "rec_calc_rel rec_F [code tp, (bl2wc (<lm>))] 
+                         (valu (rght (conf (code tp) (bl2wc (<lm>)) stpa)))"
+      apply(rule F_lemma) using g2 h by auto
+    moreover have 
+      "valu (rght (conf (code tp) (bl2wc (<lm>)) 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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/rec_def.thy	Wed Jan 23 20:18:40 2013 +0100
@@ -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; \<dots> ;gm]"} 
+  computes @{text "f (g1(x1, x2, \<dots>, xn), g2(x1, x2, \<dots>, xn), \<dots> , 
+  gm(x1, x2, \<dots> , xn))"} for input argments @{text "x1, \<dots>, xn"}.
+  *}
+  Cn nat recf "recf list" | 
+-- {*
+  The primitive resursive operator, where @{text "Pr n f g"} computes:
+  @{text "Pr n f g (x1, x2, \<dots>, xn-1, 0) = f(x1, \<dots>, xn-1)"} 
+  and @{text "Pr n f g (x1, x2, \<dots>, xn-1, k') = g(x1, x2, \<dots>, xn-1, k, 
+                                                  Pr n f g (x1, \<dots>, xn-1, k))"}.
+  *}
+  Pr nat recf recf | 
+-- {*
+  The minimization operator, where @{text "Mn n f (x1, x2, \<dots> , xn)"} 
+  computes the first i such that @{text "f (x1, \<dots>, xn, i) = 0"} and for all
+  @{text "j"}, @{text "f (x1, x2, \<dots>, 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, \<dots>, xn] r"} means the computation of 
+  @{text "R"} over input arguments @{text "[x1, x2, \<dots>, xn"} terminates
+  and gives rise to a result @{text "r"}
+*}
+
+inductive rec_calc_rel :: "recf \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> bool"
+where
+  calc_z: "rec_calc_rel z [n] 0" |
+  calc_s: "rec_calc_rel s [n] (Suc n)" |
+  calc_id: "\<lbrakk>length args = i; j < i; args!j = r\<rbrakk> \<Longrightarrow> rec_calc_rel (id i j) args r" |
+  calc_cn: "\<lbrakk>length args = n;
+             \<forall> k < length gs. rec_calc_rel (gs ! k) args (rs ! k);
+             length rs = length gs; 
+             rec_calc_rel f rs r\<rbrakk> 
+            \<Longrightarrow> rec_calc_rel (Cn n f gs) args r" |
+  calc_pr_zero: 
+           "\<lbrakk>length args = n;
+             rec_calc_rel f args r0 \<rbrakk> 
+            \<Longrightarrow> rec_calc_rel (Pr n f g) (args @ [0]) r0" |
+  calc_pr_ind: "
+           \<lbrakk> length args = n;
+             rec_calc_rel (Pr n f g) (args @ [k]) rk; 
+             rec_calc_rel g (args @ [k] @ [rk]) rk'\<rbrakk>
+            \<Longrightarrow> rec_calc_rel (Pr n f g) (args @ [Suc k]) rk'"  |
+  calc_mn: "\<lbrakk>length args = n; 
+             rec_calc_rel f (args@[r]) 0; 
+             \<forall> i < r. (\<exists> ri. rec_calc_rel f (args@[i]) ri \<and> ri \<noteq> 0)\<rbrakk> 
+            \<Longrightarrow> 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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/recursive.thy	Wed Jan 23 20:18:40 2013 +0100
@@ -0,0 +1,5104 @@
+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 \<Rightarrow> 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 \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
+  where
+  "addition m n p = [Dec m 4, Inc n, Inc p, Goto 0, Dec p 7, 
+                       Inc m, Goto 4]"
+
+fun mv_box :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
+  where
+  "mv_box m n = [Dec m 3, Inc n, Goto 0]"
+
+fun abc_inst_shift :: "abc_inst \<Rightarrow> nat \<Rightarrow> 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 \<Rightarrow> nat \<Rightarrow> abc_inst list" 
+  where
+  "abc_shift xs n = map (\<lambda> x. abc_inst_shift x n) xs" 
+
+fun abc_append :: "abc_inst list \<Rightarrow> abc_inst list \<Rightarrow> 
+                           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 \<equiv> [Goto 1]"
+
+text {*
+  The compilation of @{text "s"}-operator.
+*}
+definition rec_ci_s :: "abc_inst list"
+  where
+  "rec_ci_s \<equiv> (addition 0 1 2 [+] [Inc 1])"
+
+
+text {*
+  The compilation of @{text "id i j"}-operator
+*}
+
+fun rec_ci_id :: "nat \<Rightarrow> nat \<Rightarrow> abc_inst list"
+  where
+  "rec_ci_id i j = addition j i (i + 1)"
+
+fun mv_boxes :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_inst list"
+  where
+  "mv_boxes ab bb 0 = []" |
+  "mv_boxes ab bb (Suc n) = mv_boxes ab bb n [+] mv_box (ab + n)
+  (bb + n)"
+
+fun empty_boxes :: "nat \<Rightarrow> 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 \<times> nat \<times> nat) list \<Rightarrow> nat \<Rightarrow> abc_inst list"
+  where
+  "cn_merge_gs [] p = []" |
+  "cn_merge_gs (g # gs) p = 
+      (let (gprog, gpara, gn) = g in 
+         gprog [+] mv_box 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 \<Rightarrow> abc_inst list \<times> nat \<times> 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 (\<lambda> g. rec_ci g) (f # gs) in
+       let (fprog, fpara, fn) = hd cied_gs in 
+       let pstr = 
+        Max (set (Suc n # fn # (map (\<lambda> (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 [+] 
+            mv_box fpara pstr [+] empty_boxes (length gs) [+] 
+             mv_box 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 
+           (mv_box n p [+] fprog [+] mv_box 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) \<in> set (f # gs)" 
+  thus "(x, Cn n f gs) \<in> measure size"
+    by(induct gs, auto)
+next
+  fix n f g
+  show "(f, Pr n f g) \<in> measure size" by auto
+next
+  fix n f g x xa y xb ya
+  show "(g, Pr n f g) \<in> measure size" by auto
+next
+  fix n f
+  show "(f, Mn n f) \<in> 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]
+        mv_box.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: 
+    "\<And>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: 
+  "\<lbrakk>rec_calc_rel z l x; rec_calc_rel z l y\<rbrakk> \<Longrightarrow> x = y"
+apply(auto elim: calc_z_reverse)
+done
+
+lemma  rec_calc_inj_case_s: 
+  "\<lbrakk>rec_calc_rel s l x; rec_calc_rel s l y\<rbrakk> \<Longrightarrow> x = y"
+apply(auto elim: calc_s_reverse)
+done
+
+lemma rec_calc_inj_case_id:
+  "\<lbrakk>rec_calc_rel (recf.id nat1 nat2) l x;
+    rec_calc_rel (recf.id nat1 nat2) l y\<rbrakk> \<Longrightarrow> x = y"
+apply(auto elim: calc_id_reverse)
+done
+
+lemma rec_calc_inj_case_mn:
+  assumes ind: "\<And> l x y. \<lbrakk>rec_calc_rel f l x; rec_calc_rel f l y\<rbrakk> 
+           \<Longrightarrow> 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" 
+    "\<forall>x<y. \<exists>v. rec_calc_rel f (l @ [x]) v \<and> 0 < v" "\<not> 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: 
+  "\<And>l x y. \<lbrakk>rec_calc_rel f l x; rec_calc_rel f l y\<rbrakk> \<Longrightarrow> x = y"
+  and g_ind:
+  "\<And>x xa y xb ya l xc yb. 
+  \<lbrakk>x = rec_ci f; (xa, y) = x; (xb, ya) = y; 
+  rec_calc_rel g l xc; rec_calc_rel g l yb\<rbrakk> \<Longrightarrow> 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: 
+    "\<And> l xc yb. \<lbrakk>rec_calc_rel g l xc; rec_calc_rel g l yb\<rbrakk>
+    \<Longrightarrow> 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:
+        "\<And>ry rya. \<lbrakk>rec_calc_rel (Pr (length l) f g) (l @ [y]) ry; 
+                   rec_calc_rel (Pr (length l) f g) (l @ [y]) rya\<rbrakk> \<Longrightarrow> 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:
+  "\<forall>k<Suc (length list). (a # xs) ! k = (aa # list) ! k
+       \<Longrightarrow> \<forall>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:  
+  "\<lbrakk>length xs = length ys; \<forall> k < length xs. xs ! k = ys ! k\<rbrakk> 
+  \<Longrightarrow> xs = ys"
+apply(induct xs arbitrary: ys, simp)
+apply(case_tac ys, simp, simp)
+proof -
+  fix a xs ys aa list
+  assume ind: 
+    "\<And>ys. \<lbrakk>length list = length ys; \<forall>k<length ys. xs ! k = ys ! k\<rbrakk>
+    \<Longrightarrow> xs = ys"
+    and h: "length xs = length list" 
+    "\<forall>k<Suc (length list). (a # xs) ! k = (aa # list) ! k"
+  from h show "a = aa \<and> 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: 
+  "\<And>x l xa y.
+  \<lbrakk>x = f \<or> x \<in> set gs; rec_calc_rel x l xa; rec_calc_rel x l y\<rbrakk>
+  \<Longrightarrow> 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:
+  "\<lbrakk>rec_calc_rel f l x; 
+    rec_calc_rel f l y\<rbrakk> \<Longrightarrow> 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: 
+  "\<lbrakk>rec_calc_rel (Pr n f g) (lm @ [Suc x]) rs\<rbrakk> 
+  \<Longrightarrow> \<exists> 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 \<le> y \<Longrightarrow> Suc (y - Suc x) = y - x"
+by arith
+
+lemma calc_pr_para_not_null: 
+  "rec_calc_rel (Pr n f g) lm rs \<Longrightarrow> lm \<noteq> []"
+apply(erule calc_pr_reverse, simp, simp)
+done
+
+lemma calc_pr_less_ex: 
+ "\<lbrakk>rec_calc_rel (Pr n f g) lm rs; x \<le> last lm\<rbrakk> \<Longrightarrow> 
+ \<exists>rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rs"
+apply(subgoal_tac "lm \<noteq> []")
+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 \<Longrightarrow> 
+             \<exists>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 \<Longrightarrow> 
+       (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 \<ge> length bp \<Longrightarrow> 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: 
+  "\<lbrakk>abc_steps_l (0, lm) bp stpa = (length bp, lm1); 
+    abc_steps_l (0, lm) bp stpb = (length bp, lm2)\<rbrakk> \<Longrightarrow> 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: 
+  "\<lbrakk>\<exists>stp. abc_steps_l (0, lm) bp stp = (bs, lm');
+    bs = length bp; bp \<noteq> []\<rbrakk> 
+  \<Longrightarrow> \<exists> stp. (\<lambda> (s, l). s < bs \<and> 
+              (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 \<noteq> []"
+  thus 
+    "\<exists>stp. (\<lambda>(s, l). s < bs \<and> 
+      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')
+       \<Longrightarrow> \<exists>stp. (\<lambda>(s, l). s < length bp  \<and> 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 \<noteq> []"
+    from h show 
+      "\<exists>stp. (\<lambda>(s, l). s < length bp \<and> 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
+
+declare Let_def[simp]
+
+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 Let_def)
+done
+
+lemma abc_halt_point_step[simp]: 
+  "\<lbrakk>a < length bp; abc_steps_l (a, b) bp (Suc 0) = (length bp, lm')\<rbrakk>
+  \<Longrightarrow> 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:
+  "\<lbrakk>bs < length bp;  abc_steps_l (a, b) bp (Suc 0) = (bs, l)\<rbrakk>
+  \<Longrightarrow> 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: 
+  "\<lbrakk>bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\<rbrakk>
+ \<Longrightarrow> 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: 
+    "\<And>bs l. \<lbrakk>bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\<rbrakk>
+    \<Longrightarrow> 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:
+  "\<lbrakk>\<exists> stp. abc_steps_l (0, lm) bp stp = (bs, lm');
+    bs = length bp; 
+    as = length ap\<rbrakk>
+    \<Longrightarrow> \<exists> 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: 
+  "\<lbrakk>\<exists> stp. abc_steps_l (0, am) bp stp = (bs, bm); ss = length ap\<rbrakk>
+   \<Longrightarrow>  \<exists> 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 " \<exists>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 "\<exists>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: 
+      "\<And>bs bm. \<lbrakk>abc_steps_l (0, am) bp stp = (bs, bm);
+                 ss = length ap\<rbrakk> \<Longrightarrow> 
+          \<exists>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 
+      "\<exists>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:
+  "\<lbrakk>ap \<noteq> []; 
+    abc_steps_l (0, am) ap astp = (a, b);
+    a < length ap\<rbrakk>
+     \<Longrightarrow> (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: 
+    "\<And>a b. \<lbrakk>abc_steps_l (0, am) ap astp = (a, b); 
+             a < length ap\<rbrakk> \<Longrightarrow> 
+                  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: 
+  "\<lbrakk>\<exists> astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap\<rbrakk>
+  \<Longrightarrow> \<exists> 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: 
+  "\<lbrakk>\<exists> astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap; 
+    \<exists> bstp. abc_steps_l (0, bm) bp bstp = (bs, bm'); bs = length bp;
+    cs = as + bs; bp \<noteq> []\<rbrakk>
+  \<Longrightarrow> \<exists> 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 exponent_add_iff: "a\<up>b @ a\<up>c@ xs = a\<up>(b+c) @ xs"
+apply(auto simp: replicate_add)
+done
+
+lemma exponent_cons_iff: "a # a\<up>c @ xs = a\<up>(Suc c) @ xs"
+apply(auto simp: replicate_add)
+done
+
+lemma  [simp]: "length lm = n \<Longrightarrow>  
+  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 \<Longrightarrow> 
+  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: 
+  "\<lbrakk>length lm = n\<rbrakk> \<Longrightarrow> 
+  \<exists> 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: "\<exists> 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: 
+    "\<exists> 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 
+    "\<exists> 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 \<Longrightarrow> 
+  (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 " \<exists>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 mv_box.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: 
+    "\<And>x. \<exists>stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp = 
+                                     (0, lm @ (x + y) # 0 # suf_lm)"
+  show "\<exists>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 \<and> Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
+      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 \<and> 0 < rs_pos \<Longrightarrow> 
+                           \<not> a_md - Suc 0 < rs_pos - Suc 0"
+apply(arith)
+done
+
+lemma [simp]: 
+  "Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
+           \<not> a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))"
+apply(arith)
+done
+
+lemma butlast_append_last: "lm \<noteq> [] \<Longrightarrow> lm = butlast lm @ [last lm]"
+apply(auto)
+done
+
+lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)
+           \<Longrightarrow> (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 \<Longrightarrow> 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)
+                  \<Longrightarrow> 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]:  
+  "\<lbrakk>rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm xs\<rbrakk>
+  \<Longrightarrow> length lm = rs_pos"
+apply(simp add: rec_ci.simps rec_ci_z_def)
+apply(erule_tac calc_z_reverse, simp)
+done
+
+lemma [intro]: 
+  "\<lbrakk>rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm xs\<rbrakk>
+  \<Longrightarrow> length lm = rs_pos"
+apply(simp add: rec_ci.simps rec_ci_s_def)
+apply(erule_tac calc_s_reverse, simp)
+done
+
+lemma [intro]: 
+  "\<lbrakk>rec_ci (recf.id nat1 nat2) = (aprog, rs_pos, a_md); 
+    rec_calc_rel (recf.id nat1 nat2) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
+apply(simp add: rec_ci.simps rec_ci_id.simps)
+apply(erule_tac calc_id_reverse, simp)
+done
+
+lemma [intro]: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+    rec_calc_rel (Cn n f gs) lm xs\<rbrakk> \<Longrightarrow> 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]:
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
+    rec_calc_rel (Pr n f g) lm xs\<rbrakk> \<Longrightarrow> 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]: 
+  "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md);
+    rec_calc_rel (Mn n f) lm xs\<rbrakk> \<Longrightarrow> 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: 
+  "\<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm xs\<rbrakk>
+  \<Longrightarrow> length lm = rs_pos"
+apply(case_tac f, auto)
+done
+
+lemma ci_pr_g_paras:
+  "\<lbrakk>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\<rbrakk> \<Longrightarrow> 
+    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: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
+    rec_ci g = (a, aa, ba)\<rbrakk> \<Longrightarrow> 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) \<Longrightarrow> rp < ad"
+  by(simp add: rec_ci.simps)
+
+lemma [intro]: "rec_ci s = (ap, rp, ad) \<Longrightarrow> rp < ad"
+  by(simp add: rec_ci.simps)
+
+lemma [intro]: "rec_ci (recf.id nat1 nat2) = (ap, rp, ad) \<Longrightarrow> rp < ad"
+  by(simp add: rec_ci.simps)
+
+lemma [intro]: "rec_ci (Cn n f gs) = (ap, rp, ad) \<Longrightarrow> 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) \<Longrightarrow> 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) \<Longrightarrow> 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) \<Longrightarrow> ad > rp"
+apply(case_tac f, auto)
+done
+
+lemma [elim]: "\<lbrakk>a [+] b = []; a \<noteq> [] \<or> b \<noteq> []\<rbrakk> \<Longrightarrow> RR"
+apply(auto simp: abc_append.simps abc_shift.simps)
+done
+
+lemma [intro]: "rec_ci z = ([], aa, ba) \<Longrightarrow> False"
+by(simp add: rec_ci.simps rec_ci_z_def)
+
+lemma [intro]: "rec_ci s = ([], aa, ba) \<Longrightarrow> False"
+by(auto simp: rec_ci.simps rec_ci_s_def addition.simps)
+
+lemma [intro]: "rec_ci (id m n) = ([], aa, ba) \<Longrightarrow> False"
+by(auto simp: rec_ci.simps rec_ci_id.simps addition.simps)
+
+lemma [intro]: "rec_ci (Cn n f gs) = ([], aa, ba) \<Longrightarrow> False"
+apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_append.simps)
+apply(simp add: abc_shift.simps mv_box.simps)
+done
+
+lemma [intro]: "rec_ci (Pr n f g) = ([], aa, ba) \<Longrightarrow> 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) \<Longrightarrow> False"
+apply(case_tac "rec_ci f", auto simp: rec_ci.simps)
+done
+
+lemma rec_ci_not_null:  "rec_ci g = (a, aa, ba) \<Longrightarrow> a \<noteq> []"
+by(case_tac g, auto)
+
+lemma calc_pr_g_def:
+ "\<lbrakk>rec_calc_rel (Pr rs_pos f g) (lm @ [Suc x]) rsa;
+   rec_calc_rel (Pr rs_pos f g) (lm @ [x]) rsxa\<rbrakk>
+ \<Longrightarrow> 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: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+    rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
+  \<Longrightarrow> a_md = Suc (max (n + 3) (max bc ba))"
+by(simp add: rec_ci.simps)
+
+lemma  ci_pr_f_paras: 
+  "\<lbrakk>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)\<rbrakk>  \<Longrightarrow> ac = rs_pos - Suc 0"
+apply(subgoal_tac "\<exists>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 \<noteq> []", simp)
+apply(erule_tac calc_pr_reverse, simp, simp)
+apply(erule calc_pr_zero_ex)
+done
+
+lemma ci_pr_md_ge_f:  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+        rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> Suc bc \<le> a_md"
+apply(case_tac "rec_ci g")
+apply(simp add: rec_ci.simps, auto)
+done
+
+lemma ci_pr_md_ge_g:  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+        rec_ci g = (ab, ac, bc)\<rbrakk> \<Longrightarrow> bc < a_md"
+apply(case_tac "rec_ci f")
+apply(simp add: rec_ci.simps, auto)
+done 
+
+lemma rec_calc_rel_def0: 
+  "\<lbrakk>rec_calc_rel (Pr n f g) lm rs; rec_calc_rel f (butlast lm) rsa\<rbrakk>
+  \<Longrightarrow> 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 (mv_box m n) = 3"
+by (auto simp: mv_box.simps)
+(*
+lemma
+  "\<lbrakk>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)\<rbrakk>
+\<Longrightarrow> \<exists>ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 3 + length ab \<and> bp = recursive.mv_box (n - Suc 0) n 3"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "recursive.mv_box (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  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+        rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
+    \<Longrightarrow> \<exists>ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 3 \<and> bp = ab"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "recursive.mv_box (n - Suc 0) (max (Suc (Suc n)) (max bc ba)) 3" in exI, simp)
+apply(rule_tac x = "recursive.mv_box (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]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\<rbrakk>
+    \<Longrightarrow> rs_pos = Suc n"
+apply(simp add: ci_pr_para_eq)
+done
+
+
+lemma [simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\<rbrakk>
+    \<Longrightarrow> 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) \<Longrightarrow> 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) \<Longrightarrow> 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 \<Longrightarrow> 
+       butlast lm @ (last lm - xa) # (rsa::nat) # 0 # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm =
+       butlast lm @ (last lm - xa) # rsa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm"
+apply(simp add: replicate_Suc[THEN sym])
+done
+
+lemma pr_cycle_part_ind: 
+  assumes g_ind: 
+  "\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
+                    (length a, lm @ rs # 0\<up>(ba - Suc aa) @ 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 \<le> last lm" 
+         "rec_ci g = (a, aa, ba)"
+         "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsa"
+         "lm \<noteq> []"
+  shows 
+  "\<exists>stp. abc_steps_l 
+     (0, butlast lm @ (last lm - Suc xa) # rsxa # 
+               0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) ap stp =
+     (0, butlast lm @ (last lm - xa) # rsa
+                 # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
+proof -
+  have k1: "\<exists>stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) #
+    rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) ap stp =
+         (length a + 4, butlast lm @ (last lm - xa) # 0 # rsa #
+                           0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ 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\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm" in abc_append_exc2,
+      auto)
+  proof -
+    show 
+      "\<exists>astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa 
+                   # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) 
+              [Dec (a_md - Suc 0)(length a + 7)] astp =
+      (Suc 0, butlast lm @ (last lm - Suc xa) # 
+             rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ 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 \<and> rs_pos = Suc n \<and>
+                              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\<up>(a_md - Suc (Suc rs_pos))" 
+                 "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\<up>(a_md - Suc (Suc rs_pos))" 
+                    "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 "\<exists>bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # 
+           rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ 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\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
+      apply(rule_tac as = "length a" and
+               bm = "butlast lm @ (last lm - Suc xa) # rsxa # rsa #
+                     0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm" 
+        in abc_append_exc2, simp_all)
+    proof -
+      from h have j1: "aa = Suc rs_pos \<and> a_md > ba \<and> 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 
+        "\<exists>astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) #
+                rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) a astp =
+        (length a, butlast lm @ (last lm - Suc xa) # rsxa # rsa 
+                         # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
+	apply(insert g_ind[of
+          "butlast lm @ (last lm - Suc xa) # [rsxa]" rsa 
+          "0\<up>(a_md - ba - Suc 0) @ 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 \<and> 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
+      "\<exists>bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa #
+                     rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ 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\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
+	apply(insert pr_cycle_part_middle_inv[of "butlast lm" 
+          "rs_pos - Suc 0" "(last lm - Suc xa)" rsxa 
+          "rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm"], simp)
+	done
+    qed
+  qed
+  from h have k2: 
+    "\<exists>stp. abc_steps_l (length a + 4, butlast lm @ (last lm - xa) # 0 
+           # rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm) ap stp =
+    (0, butlast lm @ (last lm - xa) # rsa # 0\<up>(a_md - Suc (Suc rs_pos)) @ 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\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm"])
+    apply(simp add: h ap_def)
+    apply(subgoal_tac "length lm = Suc n \<and> 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: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+    rec_ci g = (a, aa, ba);
+    rec_ci f = (ab, ac, bc)\<rbrakk>
+\<Longrightarrow> \<exists>ap bp. length ap = 6 + length ab \<and>
+    aprog = ap [+] bp \<and>
+    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.mv_box n (max (Suc (Suc (Suc n)))
+    (max bc ba)) [+] ab [+] recursive.mv_box n (Suc n)" in exI,
+     simp)
+apply(auto simp add: abc_append_commute numeral_3_eq_3)
+done
+
+lemma pr_cycle_part:
+  "\<lbrakk>\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow>
+     \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
+                        (length a, lm @ rs # 0\<up>(ba - Suc aa) @ 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 \<noteq> [];
+  x \<le> last lm\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) #
+              rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ x # suf_lm) aprog stp =
+  (6 + length ab, butlast lm @ last lm # rs #
+                                0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
+proof -
+  assume g_ind:
+    "\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
+    \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp =
+                      (length a, lm @ rs # 0\<up>(ba - Suc aa) @ 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 \<noteq> []"
+           "x \<le> last lm" 
+           "rec_ci f = (ab, ac, bc)" 
+  from h show 
+    "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) # 
+            rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ x # suf_lm) aprog stp =
+    (6 + length ab, butlast lm @ last lm # rs #
+                               0\<up>(a_md - Suc (Suc rs_pos)) @ 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 \<noteq> [] \<and> rs_pos = Suc n", simp)
+      apply(rule_tac rec_calc_inj, simp, simp)
+      apply(simp)
+      done
+    thus "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @  last lm # 
+             rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm) aprog stp =
+      (6 + length ab, butlast lm @ last lm # rs #
+                               0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
+      by(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
+  next
+    fix xa rsxa
+    assume ind:
+   "\<And>rsx. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsx 
+  \<Longrightarrow> \<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - xa) #
+             rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) aprog stp =
+      (6 + length ab, butlast lm @ last lm # rs # 
+                               0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
+      and g: "rec_calc_rel (Pr n f g) 
+                      (butlast lm @ [last lm - Suc xa]) rsxa"
+      "Suc xa \<le> 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 \<noteq> []"
+    from g have k1: 
+      "\<exists> 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 
+      "\<exists>stp. abc_steps_l (6 + length ab, 
+           butlast lm @ (last lm - Suc xa) # rsxa # 
+              0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) aprog stp =
+              (6 + length ab, butlast lm @ last lm # rs # 
+                                0\<up>(a_md - Suc (Suc rs_pos)) @ 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
+      "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ 
+       (last lm - Suc xa) # rsxa # 
+               0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) aprog stp
+        = (6 + length ab, butlast lm @ (last lm - xa) # rsa # 
+                               0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
+	proof -
+	  from g have k2_1: 
+            "\<exists> ap bp. length ap = 6 + length ab \<and>
+                   aprog = ap [+] bp \<and> 
+                   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 \<and> 
+                 aprog = ap [+] bp \<and> 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\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm" bp 0
+                  "butlast lm @ (last lm - xa) # rsa #
+                0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm" "length ap" ap],
+                 simp)
+		apply(subgoal_tac 
+                "\<exists>stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa)
+                           # rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # 
+                              suf_lm) bp stp =
+	          (0, butlast lm @ (last lm - xa) # rsa #
+                           0\<up>(a_md - Suc (Suc rs_pos)) @ 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: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
+    rec_ci g = (a, aa, ba);  
+    rec_ci f = (ab, ac, bc)\<rbrakk>
+    \<Longrightarrow>  length aprog = 13 + length ab + length a"
+apply(auto simp: rec_ci.simps)
+done
+
+fun mv_box_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool"
+  where
+  "mv_box_inv (as, lm) m n initlm = 
+         (let plus = initlm ! m + initlm ! n in
+           length initlm > max m n \<and> m \<noteq> n \<and> 
+              (if as = 0 then \<exists> k l. lm = initlm[m := k, n := l] \<and> 
+                    k + l = plus \<and> k \<le> initlm ! m 
+              else if as = 1 then \<exists> k l. lm = initlm[m := k, n := l]
+                             \<and> k + l + 1 = plus \<and> k < initlm ! m 
+              else if as = 2 then \<exists> k l. lm = initlm[m := k, n := l] 
+                              \<and> k + l = plus \<and> k \<le> initlm ! m
+              else if as = 3 then lm = initlm[m := 0, n := plus]
+              else False))"
+
+fun mv_box_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "mv_box_stage1 (as, lm) m  = 
+            (if as = 3 then 0 
+             else 1)"
+
+fun mv_box_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "mv_box_stage2 (as, lm) m = (lm ! m)"
+
+fun mv_box_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "mv_box_stage3 (as, lm) m = (if as = 1 then 3 
+                                else if as = 2 then 2
+                                else if as = 0 then 1 
+                                else 0)"
+ 
+fun mv_box_measure :: "((nat \<times> nat list) \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
+  where
+  "mv_box_measure ((as, lm), m) = 
+     (mv_box_stage1 (as, lm) m, mv_box_stage2 (as, lm) m,
+      mv_box_stage3 (as, lm) m)"
+
+definition lex_pair :: "((nat \<times> nat) \<times> nat \<times> nat) set"
+  where
+  "lex_pair = less_than <*lex*> less_than"
+
+definition lex_triple :: 
+ "((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set"
+  where
+  "lex_triple \<equiv> less_than <*lex*> lex_pair"
+
+definition mv_box_LE :: 
+ "(((nat \<times> nat list) \<times> nat) \<times> ((nat \<times> nat list) \<times> nat)) set"
+  where 
+  "mv_box_LE \<equiv> (inv_image lex_triple mv_box_measure)"
+
+lemma wf_lex_triple: "wf lex_triple"
+  by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def)
+
+lemma wf_mv_box_le[intro]: "wf mv_box_LE"
+by(auto intro:wf_inv_image wf_lex_triple simp: mv_box_LE_def)
+
+declare mv_box_inv.simps[simp del]
+
+lemma mv_box_inv_init:  
+"\<lbrakk>m < length initlm; n < length initlm; m \<noteq> n\<rbrakk> \<Longrightarrow> 
+  mv_box_inv (0, initlm) m n initlm"
+apply(simp add: abc_steps_l.simps mv_box_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.mv_box m n) = Some (Dec m 3)"
+apply(simp add: mv_box.simps abc_fetch.simps)
+done
+
+lemma [simp]: "abc_fetch (Suc 0) (recursive.mv_box m n) =
+               Some (Inc n)"
+apply(simp add: mv_box.simps abc_fetch.simps)
+done
+
+lemma [simp]: "abc_fetch 2 (recursive.mv_box m n) = Some (Goto 0)"
+apply(simp add: mv_box.simps abc_fetch.simps)
+done
+
+lemma [simp]: "abc_fetch 3 (recursive.mv_box m n) = None"
+apply(simp add: mv_box.simps abc_fetch.simps)
+done
+
+lemma [simp]: 
+  "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm;
+    k + l = initlm ! m + initlm ! n; k \<le> initlm ! m; 0 < k\<rbrakk>
+ \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, m := k - Suc 0] = 
+     initlm[m := ka, n := la] \<and>
+     Suc (ka + la) = initlm ! m + initlm ! n \<and> 
+     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]:
+  "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm; 
+    Suc (k + l) = initlm ! m + initlm ! n;
+    k < initlm ! m\<rbrakk>
+    \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, n := Suc l] = 
+                initlm[m := ka, n := la] \<and> 
+                ka + la = initlm ! m + initlm ! n \<and> 
+                ka \<le> initlm ! m"
+apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto)
+done
+
+lemma [simp]: 
+  "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
+   \<forall>na. \<not> (\<lambda>(as, lm) m. as = 3) 
+    (abc_steps_l (0, initlm) (recursive.mv_box m n) na) m \<and> 
+  mv_box_inv (abc_steps_l (0, initlm) 
+           (recursive.mv_box m n) na) m n initlm \<longrightarrow>
+  mv_box_inv (abc_steps_l (0, initlm) 
+           (recursive.mv_box m n) (Suc na)) m n initlm \<and>
+  ((abc_steps_l (0, initlm) (recursive.mv_box m n) (Suc na), m),
+   abc_steps_l (0, initlm) (recursive.mv_box m n) na, m) \<in> mv_box_LE"
+apply(rule allI, rule impI, simp add: abc_steps_ind)
+apply(case_tac "(abc_steps_l (0, initlm) (recursive.mv_box m n) na)",
+      simp)
+apply(auto split:if_splits simp add:abc_steps_l.simps mv_box_inv.simps)
+apply(auto simp add: mv_box_LE_def lex_triple_def lex_pair_def 
+                     abc_step_l.simps abc_steps_l.simps
+                     mv_box_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 mv_box_inv_halt: 
+  "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
+  \<exists> stp. (\<lambda> (as, lm). as = 3 \<and> 
+  mv_box_inv (as, lm) m n initlm) 
+             (abc_steps_l (0::nat, initlm) (mv_box m n) stp)"
+thm halt_lemma2
+apply(insert halt_lemma2[of mv_box_LE
+    "\<lambda> ((as, lm), m). mv_box_inv (as, lm) m n initlm"
+    "\<lambda> stp. (abc_steps_l (0, initlm) (recursive.mv_box m n) stp, m)"
+    "\<lambda> ((as, lm), m). as = (3::nat)"
+    ])
+apply(insert wf_mv_box_le)
+apply(simp add: mv_box_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.mv_box m n) na)",
+      simp, auto)
+done
+
+lemma mv_box_halt_cond:
+  "\<lbrakk>m \<noteq> n; mv_box_inv (a, b) m n lm; a = 3\<rbrakk> \<Longrightarrow> 
+  b = lm[n := lm ! m + lm ! n, m := 0]"
+apply(simp add: mv_box_inv.simps, auto)
+apply(simp add: list_update_swap)
+done
+
+lemma mv_box_ex:
+  "\<lbrakk>length lm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
+  \<exists> stp. abc_steps_l (0::nat, lm) (mv_box m n) stp
+  = (3, (lm[n := (lm ! m + lm ! n)])[m := 0::nat])"
+apply(drule mv_box_inv_halt, simp, erule_tac exE)
+apply(rule_tac x = stp in exI)
+apply(case_tac "abc_steps_l (0, lm) (recursive.mv_box m n) stp",
+      simp)
+apply(erule_tac mv_box_halt_cond, auto)
+done
+
+lemma [simp]: 
+  "\<lbrakk>a_md = Suc (max (Suc (Suc n)) (max bc ba)); 
+   length lm = rs_pos \<and> rs_pos = n \<and> n > 0\<rbrakk>
+  \<Longrightarrow> n - Suc 0 < length lm + 
+  (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm) \<and>
+   Suc (Suc n) < length lm + (Suc (max (Suc (Suc n)) (max bc ba)) -
+  rs_pos + length suf_lm) \<and> bc < length lm + (Suc (max (Suc (Suc n)) 
+ (max bc ba)) - rs_pos + length suf_lm) \<and> ba < length lm + 
+  (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm)"
+apply(arith)
+done
+
+lemma [simp]:
+  "\<lbrakk>a_md = Suc (max (Suc (Suc n)) (max bc ba)); 
+   length lm = rs_pos \<and> rs_pos = n \<and> n > 0\<rbrakk>
+ \<Longrightarrow> n - Suc 0 < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \<and>
+     Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \<and> 
+     bc < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \<and> 
+     ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))"
+apply(arith)
+done
+
+lemma [simp]: "n - Suc 0 \<noteq> max (Suc (Suc n)) (max bc ba)"
+apply(arith)
+done
+
+lemma [simp]: 
+  "a_md \<ge> Suc bc \<and> rs_pos > 0 \<and> bc \<ge> rs_pos \<Longrightarrow> 
+ bc - (rs_pos - Suc 0) + a_md - Suc bc = Suc (a_md - rs_pos - Suc 0)"
+apply(arith)
+done
+
+lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < rs_pos \<and> 
+                                                  Suc rs_pos < a_md 
+       \<Longrightarrow> n - Suc 0 < Suc (Suc (a_md + length suf_lm - Suc (Suc 0))) 
+        \<and> n < Suc (Suc (a_md + length suf_lm - Suc (Suc 0)))"
+apply(arith)
+done
+     
+lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < rs_pos \<and> 
+               Suc rs_pos < a_md \<Longrightarrow> n - Suc 0 \<noteq> n"
+by arith
+
+lemma ci_pr_ex2: 
+  "\<lbrakk>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)\<rbrakk>
+  \<Longrightarrow> \<exists>ap bp. aprog = ap [+] bp \<and> 
+         ap = mv_box n (max (Suc (Suc (Suc n))) (max bc ba))"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "(ab [+] (recursive.mv_box 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 numeral_3_eq_3)
+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
+
+thm nth_replicate
+(*
+lemma exp_nth[simp]: "n < m \<Longrightarrow> a\<up>m ! n = a"
+apply(sim)
+done
+*)
+lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < n \<Longrightarrow> 
+                      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]: "\<lbrakk>length lm = n; 0 < n\<rbrakk>  \<Longrightarrow> 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\<up>b @ [a] = a\<up>(b + Suc 0)"
+apply(simp add: exp_ind del: replicate.simps)
+done
+
+lemma less_not_less[simp]: "n > 0 \<Longrightarrow> \<not> n < n - Suc 0"
+by auto
+
+lemma [simp]:
+  "Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \<and> 
+  bc < Suc (length suf_lm + max (Suc (Suc n)) 
+  (max bc ba)) \<and> 
+  ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))"
+  by arith
+
+lemma [simp]: "length lm = n \<and> rs_pos = n \<and> n > 0 \<Longrightarrow> 
+(lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) 
+  [max (Suc (Suc n)) (max bc ba) :=
+   (lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) ! (n - Suc 0) + 
+       (lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) ! 
+                   max (Suc (Suc n)) (max bc ba), n - Suc 0 := 0::nat]
+ = butlast lm @ 0 # 0\<up>(max (Suc (Suc n)) (max bc ba) - n) @ last lm # suf_lm"
+apply(simp add: nth_append nth_replicate list_update_append)
+apply(insert list_update_append[of "0\<up>((max (Suc (Suc n)) (max bc ba)) - n)"
+         "[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\<up>a = c\<up>b)"
+apply(auto)
+done
+
+lemma [simp]:
+  "\<lbrakk>length lm = n; 0 < n;  Suc n < a_md\<rbrakk> \<Longrightarrow> 
+   (butlast lm @ rsa # 0\<up>(a_md - Suc n) @ last lm # suf_lm)
+    [n := (butlast lm @ rsa # 0\<up>(a_md - Suc n) @ last lm # suf_lm) ! 
+        (n - Suc 0) + (butlast lm @ rsa # (0::nat)\<up>(a_md - Suc n) @ 
+                                last lm # suf_lm) ! n, n - Suc 0 := 0]
+ = butlast lm @ 0 # rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm"
+apply(simp add: nth_append list_update_append)
+apply(case_tac "a_md - Suc n", auto)
+done
+
+lemma [simp]: 
+  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos
+  \<Longrightarrow> a_md - Suc 0 < 
+          Suc (Suc (Suc (a_md + length suf_lm - Suc (Suc (Suc 0)))))"
+by arith
+
+lemma [simp]: 
+  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos \<Longrightarrow> 
+                                   \<not> a_md - Suc 0 < rs_pos - Suc 0"
+by arith
+
+lemma [simp]: "Suc (Suc rs_pos) \<le> a_md \<Longrightarrow> 
+                                \<not> a_md - Suc 0 < rs_pos - Suc 0"
+by arith
+
+lemma [simp]: "\<lbrakk>Suc (Suc rs_pos) \<le> a_md\<rbrakk> \<Longrightarrow> 
+               \<not> a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))"
+by arith 
+
+lemma [simp]: 
+  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos
+ \<Longrightarrow> (abc_lm_v (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @
+        0 # suf_lm) (a_md - Suc 0) = 0 \<longrightarrow>
+      abc_lm_s (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 
+        0 # suf_lm) (a_md - Suc 0) 0 = 
+         lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) \<and>
+     abc_lm_v (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 
+               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\<up>(a_md - Suc (Suc rs_pos))" 
+               "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]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
+      rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
+    \<Longrightarrow> \<exists>cp. aprog = recursive.mv_box n (max (n + 3) 
+                    (max bc ba)) [+] cp"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "(ab [+] (recursive.mv_box 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]: "mv_box m n \<noteq> []"
+by (simp add: mv_box.simps)
+(*
+lemma [simp]: "\<lbrakk>rs_pos = n; 0 < rs_pos ; Suc rs_pos < a_md\<rbrakk> \<Longrightarrow> 
+                        n - Suc 0 < a_md + length suf_lm"
+by arith
+*)
+lemma [intro]: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
+    rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> 
+   \<exists>ap. (\<exists>cp. aprog = ap [+] ab [+] cp) \<and> length ap = 3"
+apply(case_tac "rec_ci g", simp add: rec_ci.simps)
+apply(rule_tac x = "mv_box n 
+              (max (n + 3) (max bc c))" in exI, simp)
+apply(rule_tac x = "recursive.mv_box 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]: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+    rec_ci g = (a, aa, ba); 
+    rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> 
+    \<exists>ap. (\<exists>cp. aprog = ap [+] recursive.mv_box n (Suc n) [+] cp)
+      \<and> length ap = 3 + length ab"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "recursive.mv_box 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 [intro]: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+    rec_ci g = (a, aa, ba); 
+    rec_ci f = (ab, ac, bc)\<rbrakk>
+    \<Longrightarrow> \<exists>ap. (\<exists>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) \<and>
+             length ap = 6 + length ab"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "recursive.mv_box n
+    (max (n + 3) (max bc ba)) [+] ab [+] 
+     recursive.mv_box n (Suc n)" in exI, simp)
+apply(rule_tac x = "[]" in exI, auto)
+apply(simp add: abc_append_commute)
+done
+
+lemma [simp]: 
+  "n < Suc (max (n + 3) (max bc ba) + length suf_lm) \<and> 
+   Suc (Suc n) < max (n + 3) (max bc ba) + length suf_lm \<and> 
+   bc < Suc (max (n + 3) (max bc ba) + length suf_lm) \<and> 
+   ba < Suc (max (n + 3) (max bc ba) + length suf_lm)"
+by arith
+
+lemma [simp]: "n \<noteq> max (n + (3::nat)) (max bc ba)"
+by arith
+
+lemma [simp]:"length lm = Suc n \<Longrightarrow> lm[n := (0::nat)] = butlast lm @ [0]"
+apply(subgoal_tac "\<exists> 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 \<Longrightarrow> lm ! n =last lm"
+apply(subgoal_tac "lm \<noteq> []")
+apply(simp add: last_conv_nth, case_tac lm, simp_all)
+done
+
+lemma [simp]: "length lm = Suc n \<Longrightarrow> 
+      (lm @ (0::nat)\<up>(max (n + 3) (max bc ba) - n) @ suf_lm)
+           [max (n + 3) (max bc ba) := (lm @ 0\<up>(max (n + 3) (max bc ba) - n) @ suf_lm) ! n + 
+                  (lm @ 0\<up>(max (n + 3) (max bc ba) - n) @ suf_lm) ! max (n + 3) (max bc ba), n := 0]
+       = butlast lm @ 0 # 0\<up>(max (n + 3) (max bc ba) - Suc n) @ last lm # suf_lm"
+apply(auto simp: list_update_append nth_append)
+apply(subgoal_tac "(0\<up>(max (n + 3) (max bc ba) - n)) = 0\<up>(max (n + 3) (max bc ba) - Suc n) @ [0::nat]")
+apply(simp add: list_update_append)
+apply(simp add: exp_suc_iff)
+done
+
+lemma [simp]: "Suc (Suc n) < a_md \<Longrightarrow>  
+      n < Suc (Suc (a_md + length suf_lm - 2)) \<and>
+        n < Suc (a_md + length suf_lm - 2)"
+by(arith)
+
+lemma [simp]: "\<lbrakk>length lm = Suc n; Suc (Suc n) < a_md\<rbrakk>
+        \<Longrightarrow>(butlast lm @ (rsa::nat) # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm)
+          [Suc n := (butlast lm @ rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm) ! n +
+                  (butlast lm @ rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm) ! Suc n, n := 0]
+    = butlast lm @ 0 # rsa # 0\<up>(a_md - Suc (Suc (Suc n))) @ last lm # suf_lm"
+apply(auto simp: list_update_append)
+apply(subgoal_tac "(0\<up>(a_md - Suc (Suc n))) = (0::nat) # (0\<up>(a_md - Suc (Suc (Suc n))))", simp add: nth_append)
+apply(simp add: replicate_Suc[THEN sym])
+done
+
+lemma pr_case:
+  assumes nf_ind:
+  "\<And> lm rs suf_lm. rec_calc_rel f lm rs \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(bc - ac) @ suf_lm) ab stp = 
+                (length ab, lm @ rs # 0\<up>(bc - Suc ac) @ suf_lm)"
+  and ng_ind: "\<And> lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
+        \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
+                       (length a, lm @ rs # 0\<up>(ba - Suc aa) @ 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 "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+proof -
+  from h have k1: "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
+    = (3, butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) @ last lm # suf_lm)"
+  proof -
+    have "\<exists>bp cp. aprog = bp [+] cp \<and> bp = mv_box 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 
+           "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
+              ([] [+] recursive.mv_box n
+                  (max (n + 3) (max bc ba)) [+] cp) stp =
+             (0 + 3, butlast lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ 
+                                        last lm # suf_lm)", simp)
+      apply(rule_tac abc_append_exc1, simp_all)
+      apply(insert mv_box_ex[of "n" "(max (n + 3) 
+                 (max bc ba))" "lm @ 0\<up>(a_md - rs_pos) @ suf_lm"], simp)
+      apply(subgoal_tac "a_md = Suc (max (n + 3) (max bc ba))",
+            simp)
+      apply(subgoal_tac "length lm = Suc n \<and> 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: 
+  "\<exists> stp. abc_steps_l (3,  butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) @ 
+             last lm # suf_lm) aprog stp 
+    = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+  proof -
+    from h have k2_1: "\<exists> 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: 
+       "\<exists> stp. abc_steps_l (3, butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) 
+                 @ last lm # suf_lm) aprog stp
+        = (3 + length ab, butlast lm @ rsa # 0\<up>(a_md - rs_pos - 1) @ 
+                                             last lm # suf_lm)"
+      proof -
+	from h have j1: "
+          \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 3 \<and> 
+              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 \<ge> Suc bc \<and> rs_pos > 0 \<and> bc \<ge> 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\<up>(a_md - bc - Suc 0) @ 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: 
+      "\<exists> stp. abc_steps_l (3 + length ab, butlast lm @ rsa # 
+                  0\<up>(a_md - rs_pos - 1) @ last lm # suf_lm) aprog stp
+        = (6 + length ab, butlast lm @ 0 # rsa # 
+                       0\<up>(a_md - rs_pos - 2) @ last lm # suf_lm)"
+      proof -	     
+	from h have "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
+          length ap = 3 + length ab \<and> bp = recursive.mv_box 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 \<and> length ap = 3 + 
+                    length ab \<and> bp = recursive.mv_box n (Suc n)"
+	  thus "?thesis"
+	    apply(simp del: abc_append_commute)
+	    apply(subgoal_tac 
+              "\<exists>stp. abc_steps_l (3 + length ab, 
+               butlast lm @ rsa # 0\<up>(a_md - Suc rs_pos) @
+                 last lm # suf_lm) (ap [+] 
+                   recursive.mv_box n (Suc n) [+] cp) stp =
+              ((3 + length ab) + 3, butlast lm @ 0 # rsa # 
+                  0\<up>(a_md - Suc (Suc rs_pos)) @ last lm # suf_lm)", simp)
+	    apply(rule_tac abc_append_exc1, simp_all)
+	    apply(insert mv_box_ex[of n "Suc n" 
+                    "butlast lm @ rsa # 0\<up>(a_md - Suc rs_pos) @ 
+                          last lm # suf_lm"], simp)
+	    apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n \<and> a_md > Suc (Suc n)", simp)
+	    apply(insert h, simp)
+            done
+	qed
+      qed
+      from h have k2_3: "lm \<noteq> []"
+	apply(rule_tac calc_pr_para_not_null, simp)
+	done
+      from h and k2_2 and k2_3 have k2_2_3: 
+      "\<exists> stp. abc_steps_l (6 + length ab, butlast lm @ 
+          (last lm - last lm) # rsa # 
+            0\<up>(a_md - (Suc (Suc rs_pos))) @ last lm # suf_lm) aprog stp
+        = (6 + length ab, butlast lm @ last lm # rs # 
+                        0\<up>(a_md - Suc (Suc (rs_pos))) @ 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: 
+       "\<exists> stp. abc_steps_l (6 + length ab,
+             butlast lm @ last lm # rs # 0\<up>(a_md - rs_pos - 2) @
+                  0 # suf_lm) aprog stp
+        = (13 + length ab + length a,
+                   lm @ rs # 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+      proof -
+	from h have 
+        "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
+                     length ap = 6 + length ab \<and> 
+                    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  
+            "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ 
+                last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 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\<up>(a_md - Suc rs_pos) @ 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) \<and> 
+                            length lm = rs_pos \<and> 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 
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
+               = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ 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 \<Longrightarrow> y = x"
+by simp
+
+lemma [simp]: 
+  "\<lbrakk>rec_ci f = (a, aa, ba); 
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk> \<Longrightarrow> \<exists>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) \<Longrightarrow> rs_pos = n"
+apply(case_tac "rec_ci f", simp add: rec_ci.simps)
+done
+(*
+lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md); rec_calc_rel (Mn n f) lm rs\<rbrakk> \<Longrightarrow> 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) \<Longrightarrow> aa < ba"
+apply(simp add: ci_ad_ge_paras)
+done
+
+lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); 
+                rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> ba \<le> a_md"
+apply(simp add: rec_ci.simps)
+by arith
+
+lemma mn_calc_f: 
+  assumes ind: 
+  "\<And>aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk>  
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp    
+           = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ 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 "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
+                  aprog stp = (length a, 
+                   lm @ x # rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ suf_lm)"
+proof -
+  from h have k1: "\<exists> ap bp. aprog = ap @ bp \<and> 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 
+      "\<exists>astp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc n) @ suf_lm) a astp
+      = (length a, lm @ x # rsx # 0\<up>(a_md - Suc (Suc n)) @ suf_lm)"
+      apply(insert ind[of a "Suc n" ba  "lm @ [x]" rsx 
+             "0\<up>(a_md - ba) @ suf_lm"], simp add: exponent_add_iff h k2)
+      apply(subgoal_tac "ba > aa \<and> a_md \<ge> ba \<and> aa = Suc n", 
+            insert h, auto)
+      done
+  qed
+qed
+
+fun mn_ind_inv ::
+  "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> nat list \<Rightarrow> 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 
+                 \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
+            else if as = ss + 2 then 
+                 \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> 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 \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 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 \<or> as = ss + 1 then 3
+             else if as = ss then 4
+             else 0
+)"
+
+fun mn_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "mn_stage2 (as, lm) ss n = 
+            (if as = ss + 1 \<or> as = ss + 2 then (lm ! (Suc n))
+             else 0)"
+
+fun mn_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)"
+
+ 
+fun mn_measure :: "((nat \<times> nat list) \<times> nat \<times> nat) \<Rightarrow>
+                                                (nat \<times> nat \<times> 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 \<times> nat list) \<times> nat \<times> nat) \<times>
+                     ((nat \<times> nat list) \<times> nat \<times> nat)) set"
+  where "mn_LE \<equiv> (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) \<Longrightarrow> 
+  \<not> (\<lambda>(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]: 
+  "\<lbrakk>rec_ci f = (a, aa, ba); 
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> 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]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> 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]:
+  "\<lbrakk>rec_ci f = (a, aa, ba);
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> 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]: 
+  "\<lbrakk>rec_ci f = (a, aa, ba);
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> 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]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> 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
+   \<Longrightarrow> \<exists>y. (lm @ x # rsx # suf_lm)[Suc (length lm) := rsx - Suc 0]   
+    = lm @ x # y # suf_lm \<and> y \<le> rsx"
+apply(case_tac rsx, simp, simp)
+apply(rule_tac x = nat in exI, simp add: list_update_append)
+done
+
+lemma [simp]: 
+  "\<lbrakk>y \<le> rsx; 0 < y\<rbrakk>
+   \<Longrightarrow> \<exists>ya. (lm @ x # y # suf_lm)[Suc (length lm) := y - Suc 0] 
+          = lm @ x # ya # suf_lm \<and> ya \<le> rsx"
+apply(case_tac y, simp, simp)
+apply(rule_tac x = nat in exI, simp add: list_update_append)
+done
+
+lemma mn_halt_lemma: 
+  "\<lbrakk>rec_ci f = (a, aa, ba);
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md);
+     0 < rsx; length lm = n\<rbrakk>
+    \<Longrightarrow>
+  \<forall>na. \<not> (\<lambda>(as, lm') (ss, n). as = 0)
+  (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na) 
+                                                       (length a, n)
+ \<and> mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm)
+                       aprog na) (length a) x rsx suf_lm lm 
+\<longrightarrow> mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 
+                         (Suc na)) (length a) x rsx suf_lm lm
+ \<and> ((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) \<in> 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:
+  "\<lbrakk>rec_ci f = (a, aa, ba);
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md);
+    0 < rsx; length lm = n\<rbrakk>
+    \<Longrightarrow> \<exists> stp. (\<lambda> (as, lm'). (as = 0 \<and> 
+           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
+  "\<lambda> ((as, lm'), ss, n). mn_ind_inv (as, lm') ss x rsx suf_lm lm"
+  "\<lambda> stp. (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp, 
+  length a, n)"
+  "\<lambda> ((as, lm'), ss, n). as = 0"], 
+   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 \<Longrightarrow> 
+                Suc (a_md - Suc (Suc rs_pos)) = a_md - Suc rs_pos"
+by arith
+
+term rec_ci
+(*
+lemma [simp]: "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md); rec_calc_rel (Mn n f) lm rs\<rbrakk>  \<Longrightarrow> Suc rs_pos < a_md"
+apply(case_tac "rec_ci f")
+apply(subgoal_tac "c > b \<and> b = Suc rs_pos \<and> a_md \<ge> c")
+apply(arith, auto)
+done
+*)
+lemma mn_ind_step: 
+  assumes ind:  
+  "\<And>aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md);
+   rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
+            = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ 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 "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
+             aprog stp = (0, lm @ Suc x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+thm abc_add_exc1
+proof -
+  have k1: 
+    "\<exists> stp. abc_steps_l (0, lm @ x #  0\<up>(a_md - Suc (rs_pos)) @ suf_lm)
+         aprog stp = 
+       (length a, lm @ x # rsx # 0\<up>(a_md  - Suc (Suc rs_pos)) @ 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: 
+    "\<exists> stp. abc_steps_l (length a,
+       lm @ x # rsx # 0\<up>(a_md  - Suc (Suc rs_pos)) @ suf_lm) aprog stp = 
+        (0, lm @ Suc x # 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+    apply(frule_tac x = x and 
+       suf_lm = "0\<up>(a_md - Suc (Suc rs_pos)) @ suf_lm" in mn_halt, auto)
+    apply(rule_tac x = "stp" in exI, 
+          simp add: mn_ind_inv.simps rec_ci_not_null)
+    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]: 
+  "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md);
+    rec_calc_rel (Mn n f) lm rs\<rbrakk> \<Longrightarrow> 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]: "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md);      
+                rec_calc_rel (Mn n f) lm rs\<rbrakk>  \<Longrightarrow> Suc rs_pos < a_md"
+apply(case_tac "rec_ci f")
+apply(subgoal_tac "c > b \<and> b = Suc rs_pos \<and> a_md \<ge> c")
+apply(arith, auto)
+done
+
+lemma mn_ind_steps:  
+  assumes ind:
+  "\<And>aprog a_md rs_pos rs suf_lm lm. 
+  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+              (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ 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" 
+  "\<forall>x<rs. (\<exists> v. rec_calc_rel f (lm @ [x]) v \<and> 0 < v)"
+  "n = length lm" 
+  "x \<le> rs"
+  shows "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+                 aprog stp = (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+apply(insert h, induct x, 
+      rule_tac x = 0 in exI, simp add: abc_steps_zero, simp)
+proof -
+  fix x
+  assume k1: 
+    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+                aprog stp = (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ 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" 
+          "\<forall>x<rs.(\<exists> v. rec_calc_rel f (lm @ [x]) v \<and> v > 0)" 
+          "n = length lm" 
+          "Suc x \<le> rs" 
+          "rec_ci f = (a, aa, ba)"
+  hence k2:
+    "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - rs_pos - 1) @ suf_lm) aprog
+               stp = (0, lm @ Suc x # 0\<up>(a_md - rs_pos - 1) @ 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 
+    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+          aprog stp = (0, lm @ Suc x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    apply(auto)
+    apply(rule_tac x = "stp + stpa" in exI, simp only: abc_steps_add)
+    done
+qed
+    
+lemma [simp]: 
+"\<lbrakk>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\<rbrakk>
+ \<Longrightarrow> abc_lm_v (lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) (Suc n) = 0"
+apply(auto simp: abc_lm_v.simps nth_append)
+done
+
+lemma [simp]: 
+  "\<lbrakk>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\<rbrakk>
+    \<Longrightarrow> abc_lm_s (lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) (Suc n) 0 =
+                           lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm"
+apply(auto simp: abc_lm_s.simps list_update_append)
+done
+
+lemma mn_length: 
+  "\<lbrakk>rec_ci f = (a, aa, ba);
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+  \<Longrightarrow> 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:
+  "\<And>aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); 
+  rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+              (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ 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 "\<exists>stp. abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
+     aprog stp = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+proof -
+  from h and ind have k1:
+    "\<exists>stp.  abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
+        aprog stp = (length a,  lm @ rs # 0\<up>(a_md - Suc rs_pos) @ 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: 
+  "\<exists>stp. abc_steps_l (length a, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+    aprog stp = (length a + 5, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ 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: 
+  "\<And>aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+  and h: "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
+         "rec_calc_rel (Mn n f) lm rs"
+  shows "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
+  = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ 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" 
+            "\<forall>x<rs. \<exists>v. rec_calc_rel f (lm @ [x]) v \<and> 0 < v"
+            "n = length lm"
+  hence k1:
+    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog
+                  stp = (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    thm mn_ind_steps
+    apply(auto intro: mn_ind_steps ind)
+    done
+  from h have k2: 
+    "\<exists>stp. abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog
+         stp = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    apply(auto intro: mn_final_step ind)
+    done
+  from k1 and k2 show 
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+  (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ 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 \<Longrightarrow> rs = 0"
+apply(rule_tac calc_z_reverse, auto)
+done
+
+lemma z_case:
+  "\<lbrakk>rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+           (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ 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
+
+fun addition_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>     
+                     nat list \<Rightarrow> bool"
+  where
+  "addition_inv (as, lm') m n p lm = 
+        (let sn = lm ! n in
+         let sm = lm ! m in
+         lm ! p = 0 \<and>
+             (if as = 0 then \<exists> x. x \<le> lm ! m \<and> lm' = lm[m := x,
+                                    n := (sn + sm - x), p := (sm - x)]
+             else if as = 1 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
+                            n := (sn + sm - x - 1), p := (sm - x - 1)]
+             else if as = 2 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x, 
+                               n := (sn + sm - x), p := (sm - x - 1)]
+             else if as = 3 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
+                                   n := (sn + sm - x), p := (sm - x)]
+             else if as = 4 then \<exists> x. x \<le> lm ! m \<and> lm' = lm[m := x,
+                                       n := (sn + sm), p := (sm - x)] 
+             else if as = 5 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x, 
+                                  n := (sn + sm), p := (sm - x - 1)] 
+             else if as = 6 then \<exists> x. x < lm ! m \<and> 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 \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "addition_stage1 (as, lm) m p = 
+          (if as = 0 \<or> as = 1 \<or> as = 2 \<or> as = 3 then 2 
+           else if as = 4 \<or> as = 5 \<or> as = 6 then 1
+           else 0)"
+
+fun addition_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow>  nat \<Rightarrow> nat"
+  where
+  "addition_stage2 (as, lm) m p = 
+              (if 0 \<le> as \<and> as \<le> 3 then lm ! m
+               else if 4 \<le> as \<and> as \<le> 6 then lm ! p
+               else 0)"
+
+fun addition_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 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 \<times> nat list) \<times> nat \<times> nat) \<Rightarrow> 
+                                                 (nat \<times> nat \<times> 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 \<times> nat list) \<times> nat \<times> nat) \<times> 
+                          ((nat \<times> nat list) \<times> nat \<times> nat)) set"
+  where "addition_LE \<equiv> (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: 
+  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
+                                   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]:
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x \<le> lm ! m; 0 < x\<rbrakk>
+ \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - x, 
+                    p := lm ! m - x, m := x - Suc 0] =
+                 lm[m := xa, n := lm ! n + lm ! m - Suc xa,
+                    p := lm ! m - Suc xa]"
+apply(case_tac x, simp, simp)
+apply(rule_tac x = nat in exI, simp add: list_update_swap 
+                                         list_update_overwrite)
+done
+
+lemma [simp]:
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
+   \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - Suc x,
+                      p := lm ! m - Suc x, n := lm ! n + lm ! m - x]
+                 = lm[m := xa, n := lm ! n + lm ! m - xa, 
+                      p := lm ! m - Suc xa]"
+apply(rule_tac x = x in exI, 
+      simp add: list_update_swap list_update_overwrite)
+done
+
+lemma [simp]: 
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
+   \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - x, 
+                          p := lm ! m - Suc 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 add: list_update_overwrite)
+done
+
+lemma [simp]: 
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = (0::nat); m < p; n < p; x < lm ! m\<rbrakk>
+  \<Longrightarrow> \<exists>xa\<le>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]: 
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p;
+    x \<le> lm ! m; lm ! m \<noteq> x\<rbrakk>
+  \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m, 
+                       p := lm ! m - x, p := lm ! m - Suc x] 
+               = lm[m := xa, n := lm ! n + lm ! m, 
+                       p := lm ! m - Suc xa]"
+apply(rule_tac x = x in exI, simp add: list_update_overwrite)
+done
+
+lemma [simp]:
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
+  \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m,
+                             p := lm ! m - Suc x, m := Suc x]
+                = lm[m := Suc xa, n := lm ! n + lm ! m, 
+                             p := lm ! m - Suc xa]"
+apply(rule_tac x = x in exI, 
+     simp add: list_update_swap list_update_overwrite)
+done
+
+lemma [simp]: 
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
+  \<Longrightarrow> \<exists>xa\<le>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: 
+  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
+  \<forall>na. \<not> (\<lambda>(as, lm') (m, p). as = 7) 
+        (abc_steps_l (0, lm) (addition m n p) na) (m, p) \<and> 
+  addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm 
+\<longrightarrow> addition_inv (abc_steps_l (0, lm) (addition m n p) 
+                                 (Suc na)) m n p lm 
+  \<and> ((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) \<in> 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: 
+  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow> 
+  \<exists> stp. (\<lambda> (as, lm'). as = 7 \<and> 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
+  "\<lambda> ((as, lm'), m, p). addition_inv (as, lm') m n p lm"
+  "\<lambda> stp. (abc_steps_l (0, lm) (addition m n p) stp, m, p)"
+  "\<lambda> ((as, lm'), m, p). as = 7"], 
+  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 = [] \<Longrightarrow> RR"
+by(simp add: addition.simps)
+
+lemma [simp]: "(0\<up>2)[0 := n] = [n, 0::nat]"
+apply(subgoal_tac "2 = Suc 1", 
+      simp only: replicate.simps)
+apply(auto)
+done
+
+lemma [simp]: 
+  "\<exists>stp. abc_steps_l (0, n # 0\<up>2 @ 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\<up>2 @ 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\<up>2 @ 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:
+  "\<lbrakk>rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+apply(simp add: rec_ci.simps rec_ci_s_def, auto)
+apply(rule_tac calc_s_reverse, auto)
+done
+
+lemma [simp]: 
+  "\<lbrakk>n < length lm; lm ! n = rs\<rbrakk>
+    \<Longrightarrow> \<exists>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\<up>2 = [0, 0::nat]"
+apply(auto simp:numeral_2_eq_2)
+done
+
+lemma id_case: 
+  "\<lbrakk>rec_ci (id m n) = (aprog, rs_pos, a_md); 
+    rec_calc_rel (id m n) lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ 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:
+  "\<lbrakk>P []; \<And>a list. P list \<Longrightarrow> P (list @ [a::'a])\<rbrakk> \<Longrightarrow> 
+                                            P ((list::'a list))"
+apply(case_tac "length list", simp)
+proof -
+  fix nat
+  assume ind: "\<And>a list. P list \<Longrightarrow> 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: "\<And>list. \<lbrakk>length list = Suc nat; P []\<rbrakk> \<Longrightarrow> 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      
+  
+lemma nth_eq_butlast_nth: "\<lbrakk>length ys > Suc k\<rbrakk> \<Longrightarrow> 
+                                        ys ! k = butlast ys ! k"
+apply(subgoal_tac "\<exists> 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]: 
+"\<lbrakk>\<forall>k<Suc (length list). rec_calc_rel ((list @ [a]) ! k) lm (ys ! k);
+  length ys = Suc (length list)\<rbrakk>
+   \<Longrightarrow> \<forall>k<length list. rec_calc_rel (list ! k) lm (butlast ys ! k)"
+apply(rule allI, rule impI)
+apply(erule_tac  x = k in allE, simp add: nth_append)
+apply(subgoal_tac "ys ! k = butlast ys ! k", simp)
+apply(rule_tac nth_eq_butlast_nth, arith)
+done
+
+lemma cn_merge_gs_tl_app: 
+  "cn_merge_gs (gs @ [g]) pstr = 
+        cn_merge_gs gs pstr [+] cn_merge_gs [g] (pstr + length gs)"
+apply(induct gs arbitrary: pstr, simp add: cn_merge_gs.simps, simp)
+apply(case_tac a, simp add: abc_append_commute)
+done
+
+lemma cn_merge_gs_length: 
+  "length (cn_merge_gs (map rec_ci list) pstr) = 
+      (\<Sum>(ap, pos, n)\<leftarrow>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 \<le> pstr \<Longrightarrow> pstr + x - n > 0"
+by arith
+
+lemma [simp]:
+  "\<lbrakk>Suc (pstr + length list) \<le> a_md; 
+    length ys = Suc (length list);
+    length lm = n;
+     Suc n \<le> pstr\<rbrakk>
+   \<Longrightarrow>  (ys ! length list # 0\<up>(pstr - Suc n) @ butlast ys @
+             0\<up>(a_md - (pstr + length list)) @ suf_lm) ! 
+                      (pstr + length list - n) = (0 :: nat)"
+apply(insert nth_append[of "ys ! length list # 0\<up>(pstr - Suc n) @
+     butlast ys" "0\<up>(a_md - (pstr + length list)) @ suf_lm"
+      "(pstr + length list - n)"], simp add: nth_append)
+done
+
+lemma [simp]:
+  "\<lbrakk>Suc (pstr + length list) \<le> a_md; 
+    length ys = Suc (length list);
+    length lm = n;
+     Suc n \<le> pstr\<rbrakk>
+    \<Longrightarrow> (lm @ last ys # 0\<up>(pstr - Suc n) @ butlast ys @
+         0\<up>(a_md - (pstr + length list)) @ suf_lm)[pstr + length list := 
+                                        last ys, n := 0] =
+        lm @ (0::nat)\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm"
+apply(insert list_update_length[of 
+   "lm @ last ys # 0\<up>(pstr - Suc n) @ butlast ys" 0 
+   "0\<up>(a_md - Suc (pstr + length list)) @ suf_lm" "last ys"], simp)
+apply(simp add: exponent_cons_iff)
+apply(insert list_update_length[of "lm" 
+        "last ys" "0\<up>(pstr - Suc n) @ butlast ys @ 
+      last ys # 0\<up>(a_md - Suc (pstr + length list)) @ 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: 
+  "\<lbrakk>\<And>x aprog a_md rs_pos rs suf_lm lm.
+    \<lbrakk>x \<in> set gs; rec_ci x = (aprog, rs_pos, a_md);
+     rec_calc_rel x lm rs\<rbrakk>
+     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
+           = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm); 
+   pstr + length gs\<le> a_md;
+   \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
+   length ys = length gs; length lm = n;
+   pstr \<ge> Max (set (Suc n # map (\<lambda>(aprog, p, n). n) (map rec_ci gs)))\<rbrakk>
+  \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm)
+                   (cn_merge_gs (map rec_ci gs) pstr) stp 
+   = (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) gs) +
+  3 * length gs, lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length gs)) @ 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: "\<And>x aprog a_md rs_pos rs suf_lm lm.
+    \<lbrakk>x = a \<or> x \<in> set list; rec_ci x = (aprog, rs_pos, a_md); 
+     rec_calc_rel x lm rs\<rbrakk>
+    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+                (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+  and ind2: 
+    "\<And>ys. \<lbrakk>\<And>x aprog a_md rs_pos rs suf_lm lm.
+    \<lbrakk>x \<in> set list; rec_ci x = (aprog, rs_pos, a_md);
+     rec_calc_rel x lm rs\<rbrakk>
+    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
+        = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm);
+    \<forall>k<length list. rec_calc_rel (list ! k) lm (ys ! k); 
+    length ys = length list\<rbrakk>
+    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm) 
+                   (cn_merge_gs (map rec_ci list) pstr) stp =
+    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
+     3 * length list,
+                lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)"
+    and h: "Suc (pstr + length list) \<le> a_md" 
+            "\<forall>k<Suc (length list). 
+                   rec_calc_rel ((list @ [a]) ! k) lm (ys ! k)" 
+            "length ys = Suc (length list)" 
+            "length lm = n"
+            "Suc n \<le> pstr \<and> (\<lambda>(aprog, p, n). n) (rec_ci a) \<le> pstr \<and> 
+            (\<forall>a\<in>set list. (\<lambda>(aprog, p, n). n) (rec_ci a) \<le> pstr)"
+  from h have k1: 
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm)
+                     (cn_merge_gs (map rec_ci list) pstr) stp =
+    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
+     3 * length list, lm @ 0\<up>(pstr - n) @ butlast ys @
+                               0\<up>(a_md - (pstr + length list)) @ suf_lm) "
+    apply(rule_tac ind2)
+    apply(rule_tac ind, auto)
+    done
+  from k1 and h show 
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm) 
+          (cn_merge_gs (map rec_ci list @ [rec_ci a]) pstr) stp =
+        (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) + 
+        (\<lambda>(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list),
+             lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
+    apply(simp add: cn_merge_gs_tl_app)
+    thm abc_append_exc2
+    apply(rule_tac as = 
+  "(\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list"    
+      and bm = "lm @ 0\<up>(pstr - n) @ butlast ys @ 
+                              0\<up>(a_md - (pstr + length list)) @ suf_lm" 
+      and bs = "(\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3" 
+      and bm' = "lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ 
+                                  suf_lm" in abc_append_exc2, simp)
+    apply(simp add: cn_merge_gs_length)
+  proof -
+    from h show 
+      "\<exists>bstp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ butlast ys @ 
+                                  0\<up>(a_md - (pstr + length list)) @ suf_lm) 
+              ((\<lambda>(gprog, gpara, gn). gprog [+] recursive.mv_box gpara 
+              (pstr + length list)) (rec_ci a)) bstp =
+              ((\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3, 
+             lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
+      apply(case_tac "rec_ci a", simp)
+      apply(rule_tac as = "length aa" and 
+                     bm = "lm @ (ys ! (length list)) # 
+          0\<up>(pstr - Suc n) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm" 
+        and bs = "3" and bm' = "lm @ 0\<up>(pstr - n) @ ys @
+             0\<up>(a_md - Suc (pstr + length list)) @ 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 
+        "\<exists>astp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ butlast ys @ 
+                         0\<up>(a_md - (pstr + length list)) @ suf_lm) aa astp =
+        (length aa, lm @ ys ! length list # 0\<up>(pstr - Suc n) @ 
+                       butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)"
+	apply(subgoal_tac "c \<ge> Suc n")
+	apply(insert ind[of a aa b c lm "ys ! length list" 
+          "0\<up>(pstr - c) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ 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     
+      "\<exists>bstp. abc_steps_l (0, lm @ ys ! length list #
+          0\<up>(pstr - Suc n) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)
+                 (recursive.mv_box b (pstr + length list)) bstp =
+       (3, lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
+	apply(insert mv_box_ex [of b "pstr + length list" 
+         "lm @ ys ! length list # 0\<up>(pstr - Suc n) @ butlast ys @ 
+         0\<up>(a_md - (pstr + length list)) @ 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.mv_box 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 "mv_box b (pstr + length list) \<noteq> []" 
+        by(simp add: mv_box.simps)
+    qed
+  next
+    show "(\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3 = 
+        length ((\<lambda>(gprog, gpara, gn). gprog [+]
+           recursive.mv_box gpara (pstr + length list)) (rec_ci a))"
+      by(case_tac "rec_ci a", simp)
+  next
+    show "listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
+      (\<lambda>(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list)=
+      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list + 
+                ((\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3)" by simp
+  next
+    show "(\<lambda>(gprog, gpara, gn). gprog [+] 
+      recursive.mv_box gpara (pstr + length list)) (rec_ci a) \<noteq> []"
+      by(case_tac "rec_ci a", 
+         simp add: abc_append.simps abc_shift.simps)
+  qed
+qed
+ 
+lemma [simp]: "length (mv_boxes aa ba n) = 3*n"
+by(induct n, auto simp: mv_boxes.simps)
+
+lemma exp_suc: "a\<up>Suc b = a\<up>b @ [a]"
+by(simp add: exp_ind del: replicate.simps)
+
+lemma [simp]: 
+  "\<lbrakk>Suc n \<le> ba - aa;  length lm2 = Suc n;
+    length lm3 = ba - Suc (aa + n)\<rbrakk>
+  \<Longrightarrow> (last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = (0::nat)"
+proof -
+  assume h: "Suc n \<le> 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 \<Longrightarrow>
+      (lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (aa + n) = last lm2"
+apply(simp add: nth_append)
+done
+
+lemma [simp]: "\<lbrakk>Suc n \<le> ba - aa; aa < ba\<rbrakk> \<Longrightarrow> 
+                    (ba < Suc (aa + (ba - Suc (aa + n) + n))) = False"
+apply arith
+done
+
+lemma [simp]: "\<lbrakk>Suc n \<le> ba - aa; aa < ba; length lm1 = aa; 
+       length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\<rbrakk>
+     \<Longrightarrow> (lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba + n) = 0"
+using nth_append[of "lm1 @ (0\<Colon>'a)\<up>n @ last lm2 # lm3 @ butlast lm2" 
+                     "(0\<Colon>'a) # lm4" "ba + n"]
+apply(simp)
+done
+
+lemma [simp]: 
+ "\<lbrakk>Suc n \<le> ba - aa; aa < ba; length lm1 = aa; length lm2 = Suc n;
+                 length lm3 = ba - Suc (aa + n)\<rbrakk>
+  \<Longrightarrow> (lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ (0::nat) # lm4)
+  [ba + n := last lm2, aa + n := 0] = 
+  lm1 @ 0 # 0\<up>n @ lm3 @ lm2 @ lm4"
+using list_update_append[of "lm1 @ 0\<up>n @ 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:
+  "\<lbrakk>n \<le> ba - aa; ba > aa; length lm1 = aa; 
+    length (lm2::nat list) = n; length lm3 = ba - aa - n\<rbrakk>
+     \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ lm4)
+       (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<up>n @ 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\<up>n @ 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:
+    "\<And>lm2 lm3 lm4. \<lbrakk>length lm2 = n; length lm3 = ba - (aa + n)\<rbrakk> \<Longrightarrow>
+    \<exists>stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ lm4) 
+       (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<up>n @ lm3 @ lm2 @ lm4)"
+  and h: "Suc n \<le> 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 
+    "\<exists>astp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ 0 # lm4) 
+                       (mv_boxes aa ba n) astp = 
+        (3 * n, lm1 @ 0\<up>n @ 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\<up>n @ 
+              0 # lm4 = lm1 @ lm2 @ lm3 @ 0\<up>n @ 0 # lm4", simp, simp)
+    apply(case_tac "lm2 = []", simp, simp)
+    done
+next
+  fix n lm2 lm3 lm4
+  assume h: "Suc n \<le> ba - aa"
+            "aa < ba" 
+            "length (lm1::nat list) = aa" 
+            "length (lm2::nat list) = Suc n" 
+            "length (lm3::nat list) = ba - Suc (aa + n)"
+  thus " \<exists>bstp. abc_steps_l (0, lm1 @ 0\<up>n @ last lm2 # lm3 @
+                       butlast lm2 @ 0 # lm4) 
+                         (recursive.mv_box (aa + n) (ba + n)) bstp
+               = (3, lm1 @ 0 # 0\<up>n @ lm3 @ lm2 @ lm4)"
+    apply(insert mv_box_ex[of "aa + n" "ba + n" 
+       "lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4"], simp)
+    done
+qed
+(*    
+lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; 
+                ba < aa; 
+               length lm2 = aa - Suc (ba + n)\<rbrakk>
+      \<Longrightarrow> ((0::nat) # lm2 @ 0\<up>n @ last lm3 # lm4) ! (aa - ba)
+         = last lm3"
+proof -
+  assume h: "Suc n \<le> 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\<up>n @ last lm3 # lm4) ! (aa - ba) = last lm3"
+    apply(simp,  simp add: nth_append)
+    done
+qed
+*)
+
+lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
+        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk>
+   \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4) ! (aa + n) = last lm3"
+using nth_append[of "lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n" "last lm3 # lm4" "aa + n"]
+apply(simp)
+done
+
+lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
+        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk>
+     \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4) ! (ba + n) = 0"
+apply(simp add: nth_append)
+done
+
+lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
+        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk> 
+     \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4)[ba + n := last lm3, aa + n := 0]
+      = lm1 @ lm3 @ lm2 @ 0 # 0\<up>n @ lm4"
+using list_update_append[of "lm1 @ butlast lm3" "(0\<Colon>'a) # lm2 @ (0\<Colon>'a)\<up>n @ last lm3 # lm4"]
+apply(simp)
+using list_update_append[of "lm1 @ butlast lm3 @ last lm3 # lm2 @ (0\<Colon>'a)\<up>n"
+                            "last lm3 # lm4" "aa + n" "0"]
+apply(simp)
+apply(simp only: replicate_Suc[THEN sym] exp_suc, simp)
+apply(case_tac lm3, simp, simp)
+done
+
+lemma mv_boxes_ex2:
+  "\<lbrakk>n \<le> aa - ba; 
+    ba < aa; 
+    length (lm1::nat list) = ba;
+    length (lm2::nat list) = aa - ba - n; 
+    length (lm3::nat list) = n\<rbrakk>
+     \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1 @ 0\<up>n @ lm2 @ lm3 @ lm4) 
+                (mv_boxes aa ba n) stp =
+                    (3 * n, lm1 @ lm3 @ lm2 @ 0\<up>n @ 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\<up>n @ 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: 
+"\<And>lm2 lm3 lm4. \<lbrakk>length lm2 = aa - (ba + n); length lm3 = n\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm1 @ 0\<up>n @ lm2 @ lm3 @ lm4) 
+                 (mv_boxes aa ba n) stp = 
+                            (3 * n, lm1 @ lm3 @ lm2 @ 0\<up>n @ lm4)"
+  and h: "Suc n \<le> 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
+    "\<exists>astp. abc_steps_l (0, lm1 @ 0\<up>n @ 0 # lm2 @ lm3 @ lm4)
+        (mv_boxes aa ba n) astp = 
+          (3 * n, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4)"
+    apply(insert ind[of "0 # lm2" "butlast lm3" "last lm3 # lm4"],
+          simp)
+    apply(subgoal_tac
+      "lm1 @ 0\<up>n @ 0 # lm2 @ butlast lm3 @ last lm3 # lm4 =
+           lm1 @ 0\<up>n @ 0 # lm2 @ lm3 @ lm4", simp, simp)
+    apply(case_tac "lm3 = []", simp, simp)
+    done
+next
+  fix n lm2 lm3 lm4
+  assume h:
+    "Suc n \<le> aa - ba" 
+    "ba < aa"
+    "length lm1 = ba"
+    "length (lm2::nat list) = aa - Suc (ba + n)" 
+    "length (lm3::nat list) = Suc n"
+  thus
+    "\<exists>bstp. abc_steps_l (0, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ 
+                               last lm3 # lm4) 
+           (recursive.mv_box (aa + n) (ba + n)) bstp =
+                 (3, lm1 @ lm3 @ lm2 @ 0 # 0\<up>n @ lm4)"
+    apply(insert mv_box_ex[of "aa + n" "ba + n" "lm1 @ butlast lm3 @ 
+                          0 # lm2 @ 0\<up>n @ last lm3 # lm4"], simp)
+    done
+qed
+
+lemma cn_merge_gs_len: 
+  "length (cn_merge_gs (map rec_ci gs) pstr) = 
+      (\<Sum>(ap, pos, n)\<leftarrow>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 \<Longrightarrow>
+     Suc (pstr + length ys - n) = Suc (pstr + length ys) - n"
+by arith
+
+lemma save_paras':  
+  "\<lbrakk>length lm = n; pstr > n; a_md > pstr + length ys + n\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ ys @
+               0\<up>(a_md - pstr - length ys) @ suf_lm) 
+                 (mv_boxes 0 (pstr + Suc (length ys)) n) stp
+        = (3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+thm mv_boxes_ex
+apply(insert mv_boxes_ex[of n "pstr + Suc (length ys)" 0 "[]" "lm" 
+         "0\<up>(pstr - n) @ ys @ [0]" "0\<up>(a_md - pstr - length ys - n - Suc 0) @ 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 (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs))))
+ = (Max (insert ba (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs)))"
+apply(rule min_max.sup_absorb2, auto)
+done
+
+lemma [simp]:
+  "((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs) = 
+                  (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs)"
+apply(induct gs)
+apply(simp, simp)
+done
+
+lemma ci_cn_md_def:  
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+  rec_ci f = (a, aa, ba)\<rbrakk>
+    \<Longrightarrow> a_md = max (Suc n) (Max (insert ba (((\<lambda>(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:
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+    rec_ci f = (a, aa, ba); 
+    pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                                    (map rec_ci (f # gs))))\<rbrakk>
+    \<Longrightarrow> \<exists>ap bp cp. 
+      aprog = ap [+] bp [+] cp \<and>
+      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+              3 * length gs \<and> 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 
+      (((\<lambda>(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 (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
+   0 (length gs) [+] a [+]recursive.mv_box aa (max (Suc n) 
+   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+   empty_boxes (length gs) [+] recursive.mv_box (max (Suc n) 
+  (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+   mv_boxes (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) 
+   ` set gs))) + length gs)) 0 n" in exI, auto)
+apply(simp add: abc_append_commute)
+done
+
+lemma save_paras: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+    rs_pos = n;
+    \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
+    length ys = length gs;
+    length lm = n;
+    rec_ci f = (a, aa, ba);
+    pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                                          (map rec_ci (f # gs))))\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+          3 * length gs, lm @ 0\<up>(pstr - n) @ ys @
+                 0\<up>(a_md - pstr - length ys) @ suf_lm) aprog stp = 
+           ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+                      3 * length gs + 3 * n, 
+             0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+  assume h:
+    "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
+    "rs_pos = n" 
+    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)" 
+    "length ys = length gs"  
+    "length lm = n"    
+    "rec_ci f = (a, aa, ba)"
+    and g: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                                        (map rec_ci (f # gs))))"
+  from h and g have k1: 
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
+    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+                3 *length gs \<and> bp = mv_boxes 0 (pstr + Suc (length ys)) n"
+    apply(drule_tac save_paras_prog_ex, auto)
+    done
+  from h have k2: 
+    "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ ys @ 
+                         0\<up>(a_md - pstr - length ys) @ suf_lm)
+         (mv_boxes 0 (pstr + Suc (length ys)) n) stp = 
+        (3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ 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 
+    "\<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+         3 * length gs, lm @ 0\<up>(pstr - n) @ ys @ 
+                 0\<up>(a_md - pstr - length ys) @ suf_lm) aprog stp =
+             ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+               3 * length gs + 3 * n, 
+                0\<up> pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ 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 \<and> length ap = 
+            (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs
+            \<and> 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) \<Longrightarrow> rs_pos = n"
+apply(simp add: rec_ci.simps, case_tac "rec_ci f", simp)
+done
+
+lemma calc_gs_prog_ex: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+    rec_ci f = (a, aa, ba);
+    Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                         (map rec_ci (f # gs)))) = pstr\<rbrakk>
+   \<Longrightarrow> \<exists>ap bp. aprog = ap [+] bp \<and> 
+                 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 (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
+   mv_boxes (max (Suc n) (Max (insert ba 
+  (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
+   a [+] recursive.mv_box aa (max (Suc n)
+    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+   empty_boxes (length gs) [+] recursive.mv_box (max (Suc n)
+    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+    mv_boxes (Suc (max (Suc n) (Max 
+    (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n"
+   in exI)
+apply(auto simp: abc_append_commute)
+done
+
+lemma cn_calc_gs: 
+  assumes ind: 
+  "\<And>x aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>x \<in> set gs; 
+   rec_ci x = (aprog, rs_pos, a_md); 
+   rec_calc_rel x lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+     (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+  and h:  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"  
+          "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+          "length ys = length gs" 
+          "length lm = n" 
+          "rec_ci f = (a, aa, ba)" 
+          "Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                               (map rec_ci (f # gs)))) = pstr"
+  shows  
+  "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs, 
+   lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md -pstr - length ys) @ suf_lm) "
+proof -
+  from h have k1:
+    "\<exists> ap bp. aprog = ap [+] bp \<and> 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 \<ge> 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:
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) 
+                         (cn_merge_gs (map rec_ci gs) pstr) stp 
+    = ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs, 
+                  lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - pstr - length ys) @ 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 
+      "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
+           (cn_merge_gs (map rec_ci gs) pstr [+] bp) stp =
+      (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) gs) +
+         3 * length gs, 
+         lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length ys)) @ suf_lm)"
+      apply(insert abc_append_exc1[of 
+        "lm @ 0\<up>(a_md - rs_pos) @ suf_lm" 
+        "(cn_merge_gs (map rec_ci gs) pstr)" 
+        "length (cn_merge_gs (map rec_ci gs) pstr)" 
+        "lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - pstr - length ys) @ suf_lm" 0 
+        "[]" bp], simp add: cn_merge_gs_len)
+      done      
+  qed
+qed
+
+lemma reset_new_paras': 
+  "\<lbrakk>length lm = n; 
+    pstr > 0; 
+    a_md \<ge> pstr + length ys + n;
+     pstr > length ys\<rbrakk> \<Longrightarrow>
+   \<exists>stp. abc_steps_l (0, 0\<up>pstr @ ys @ 0 # lm @  0\<up>(a_md - Suc (pstr + length ys + n)) @
+          suf_lm) (mv_boxes pstr 0 (length ys)) stp =
+  (3 * length ys, ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+thm mv_boxes_ex2
+apply(insert mv_boxes_ex2[of "length ys" "pstr" 0 "[]"
+     "0\<up>(pstr - length ys)" "ys" 
+     "0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm"], 
+     simp add: exponent_add_iff)
+done
+
+lemma [simp]:  
+  "\<lbrakk>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 (\<lambda>(aprog, p, n). n) 
+               (map rec_ci (f # gs))))\<rbrakk>
+  \<Longrightarrow> length ys < pstr"
+apply(subgoal_tac "length ys = aa", simp)
+apply(subgoal_tac "aa < ba \<and> ba \<le> 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: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+   rec_ci f = (a, aa, ba);
+   Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+  (map rec_ci (f # gs)))) = pstr\<rbrakk>
+  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
+  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+           3 *length gs + 3 * n \<and> 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 (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
+          mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
+           (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n" in exI, 
+       simp add: cn_merge_gs_len)
+apply(rule_tac x = "a [+]
+     recursive.mv_box aa (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+     empty_boxes (length gs) [+] recursive.mv_box 
+     (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n
+      [+] mv_boxes (Suc (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI,
+       auto simp: abc_append_commute)
+done
+
+lemma reset_new_paras:
+       "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+        rs_pos = n;
+        \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
+        length ys = length gs;
+        length lm = n;
+        length ys = aa;
+        rec_ci f = (a, aa, ba);
+        pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                                    (map rec_ci (f # gs))))\<rbrakk>
+\<Longrightarrow> \<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+                                               3 * length gs + 3 * n,
+        0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
+  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n,
+           ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+  assume h:
+    "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
+    "rs_pos = n" 
+    "length ys = aa"
+    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+    "length ys = length gs"  "length lm = n"    
+    "rec_ci f = (a, aa, ba)"
+    and g: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                                         (map rec_ci (f # gs))))"
+  thm rec_ci.simps
+  from h and g have k1:
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 
+    (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+          3 *length gs + 3 * n \<and> bp = mv_boxes pstr 0 (length ys)"
+    by(drule_tac reset_new_paras_prog_ex, auto)
+  from h have k2:
+    "\<exists> stp. abc_steps_l (0, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @
+              suf_lm) (mv_boxes pstr 0 (length ys)) stp = 
+    (3 * (length ys), 
+     ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ 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 \<and> aa < ba \<and> ba \<le> pstr", arith,
+          simp add: para_pattern)
+    apply(insert g, auto intro: min_max.le_supI2)
+    done
+  from k1 show 
+    "\<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3
+    * length gs + 3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ 
+     suf_lm) aprog stp =
+    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs +
+      3 * n, ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ 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 \<and> length ap = 
+      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs +
+                  3 * n \<and> bp = mv_boxes pstr 0 (length ys)"
+    from this and k2 show "?thesis"
+      apply(simp)
+      apply(drule_tac as = 
+        "(\<Sum>(ap, pos, n)\<leftarrow>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: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+    rec_ci f = (a, aa, ba);
+    Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                   (map rec_ci (f # gs)))) = pstr\<rbrakk>
+   \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
+  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+                                6 *length gs + 3 * n \<and> 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
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
+     mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
+            (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
+     mv_boxes (max (Suc n) (Max (insert ba 
+      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs)" in exI,
+     simp add: cn_merge_gs_len)
+apply(rule_tac x = "recursive.mv_box aa (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+     empty_boxes (length gs) [+] recursive.mv_box (max (Suc n) (
+     Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+     mv_boxes (Suc (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI,
+  auto simp: abc_append_commute)
+done
+
+lemma calc_cn_f:
+  assumes ind:
+  "\<And>x aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>x \<in> set (f # gs);
+  rec_ci x = (aprog, rs_pos, a_md); 
+  rec_calc_rel x lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+  (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ 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 (\<lambda>(aprog, p, n). n)
+                                (map rec_ci (f # gs))))"
+  shows "\<exists>stp. abc_steps_l   
+  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n,
+  ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
+  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 
+                3 * n + length a,
+  ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+  from h have k1: 
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
+    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+    6 *length gs + 3 * n \<and> 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 \<and> 
+      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+      6 * length gs + 3 * n \<and> 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\<up>(pstr - ba + length gs) @ 0 # lm @ 
+        0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp)
+      apply(subgoal_tac "ba > aa \<and> aa = length gs\<and> pstr \<ge> 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]: 
+  "\<lbrakk>pstr + length ys + n \<le> a_md; ys \<noteq> []\<rbrakk> \<Longrightarrow> 
+                          pstr < a_md + length suf_lm"
+apply(case_tac "length ys", simp)
+apply(arith)
+done
+*)
+
+lemma [simp]: 
+  "pstr > length ys 
+  \<Longrightarrow> (ys @ rs # 0\<up>pstr @ lm @
+  0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) ! pstr = (0::nat)"
+apply(simp add: nth_append)
+done
+
+(*
+lemma [simp]: "\<lbrakk>length ys < pstr; pstr - length ys = Suc x\<rbrakk>
+  \<Longrightarrow> pstr - Suc (length ys) = x"
+by arith
+*)
+
+lemma [simp]: "pstr > length ys \<Longrightarrow> 
+      (ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)
+                                         [pstr := rs, length ys := 0] =
+       ys @ 0\<up>(pstr - length ys) @ (rs::nat) # 0\<up>length ys @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm"
+apply(auto simp: list_update_append)
+apply(case_tac "pstr - length ys",simp_all)
+using list_update_length[of 
+  "0\<up>(pstr - Suc (length ys))" "0" "0\<up>length ys @ lm @ 
+  0\<up>(a_md - Suc (pstr + length ys + n)) @ 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': 
+  "\<lbrakk>pstr > length ys\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, ys @ rs # 0\<up>pstr @ lm @ 
+  0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) 
+  (recursive.mv_box (length ys) pstr) stp =
+  (3, ys @ 0\<up>(pstr - (length ys)) @ rs # 
+  0\<up>length ys  @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+apply(insert mv_box_ex[of "length ys" pstr 
+  "ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc(pstr + length ys + n)) @ suf_lm"], 
+  simp)
+done
+
+
+lemma save_rs_prog_ex:
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+  rec_ci f = (a, aa, ba);
+  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                        (map rec_ci (f # gs)))) = pstr\<rbrakk>
+  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
+  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+              6 *length gs + 3 * n + length a
+  \<and> bp = mv_box 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 
+   (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
+   [+] mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
+   (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
+   mv_boxes (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> 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.mv_box (max (Suc n) (Max (insert ba 
+    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+   mv_boxes (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> 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"
+  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+  "length ys = length gs" 
+  "rec_calc_rel f ys rs" 
+  "rec_ci f = (a, aa, ba)"  
+  "length lm = n"
+  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                                            (map rec_ci (f # gs))))"
+  shows "\<exists>stp. abc_steps_l
+           ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs
+          + 3 * n + length a, ys @ rs # 0\<up>pstr @ lm @
+             0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
+  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs 
+  + 3 * n + length a + 3,
+  ys @ 0\<up>(pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
+                               0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+  thm rec_ci.simps
+  from h and pdef have k1: 
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
+    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+    6 *length gs + 3 * n + length a \<and> bp = mv_box (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 
+    "\<exists>stp. abc_steps_l
+    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n
+    + length a, ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) 
+    @ suf_lm) aprog stp =
+    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n
+    + length a + 3, ys @ 0\<up>(pstr - length ys) @ rs # 
+    0\<up>length ys @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ 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 \<and> length ap = 
+      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 
+      3 * n + length a \<and> bp = recursive.mv_box (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 \<and> pstr \<ge> ba \<and> 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 mv_box_step_ex: "length lm = n \<Longrightarrow> 
+      \<exists>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 mv_box_ex': 
+  "\<lbrakk>length lm = n\<rbrakk> \<Longrightarrow> 
+  \<exists> 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 mv_box_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 \<Longrightarrow> list = drop (Suc n) lm"
+apply(induct n arbitrary: lm a list, simp)
+apply(case_tac "lm", simp, simp)
+done
+
+lemma empty_boxes_ex: "\<lbrakk>length lm \<ge> n\<rbrakk>
+     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm) (empty_boxes n) stp = 
+                                          (2*n, 0\<up>n @ 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 \<le> length lm"  "drop n lm = a # list"
+  thus "\<exists>bstp. abc_steps_l (0, 0\<up>n @ a # list) [Dec n 2, Goto 0] bstp =
+                       (Suc (Suc 0), 0 # 0\<up>n @ drop (Suc n) lm)"
+    apply(insert mv_box_ex'[of "0\<up>n" n a list], simp, erule_tac exE)
+    apply(rule_tac x = stp in exI, simp, simp only: exponent_cons_iff)
+    apply(simp add:exp_ind del: replicate.simps)
+    done
+qed
+
+
+lemma mv_box_paras_prog_ex:
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+  rec_ci f = (a, aa, ba); 
+  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                    (map rec_ci (f # gs)))) = pstr\<rbrakk>
+  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
+  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+  6 *length gs + 3 * n + length a + 3 \<and> 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 (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
+    mv_boxes 0 (Suc (max (Suc n) (Max 
+     (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n
+    [+] mv_boxes (max (Suc n) (Max (insert ba 
+    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
+     a [+] recursive.mv_box aa (max (Suc n) 
+   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))" 
+    in exI, simp add: cn_merge_gs_len)
+apply(rule_tac x = " recursive.mv_box (max (Suc n) (Max (insert ba
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+     mv_boxes (Suc (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI, 
+  auto simp: abc_append_commute)
+done
+
+lemma mv_box_paras: 
+ assumes h: 
+  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
+  "rec_calc_rel (Cn n f gs) lm rs" 
+  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+  "length ys = length gs" 
+  "rec_calc_rel f ys rs" 
+  "rec_ci f = (a, aa, ba)" 
+  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                                             (map rec_ci (f # gs))))"
+  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+                              6 * length gs + 3 * n + length a + 3"
+  shows "\<exists>stp. abc_steps_l
+           (ss, ys @ 0\<up>(pstr - length ys) @ rs # 0\<up>length ys 
+               @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
+   (ss + 2 * length gs, 0\<up>pstr @ rs # 0\<up>length ys  @ lm @ 
+                                0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+  from h and pdef and starts have k1: 
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
+    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+                               6 *length gs + 3 * n + length a + 3
+    \<and> bp = empty_boxes (length ys)"
+    by(drule_tac mv_box_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 \<le> 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 \<and> 
+      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+      6 * length gs + 3 * n + length a + 3 \<and> 
+      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\<up>(pstr - (length gs)) @ rs #
+        0\<up>length gs @ lm @ 0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], 
+        simp add: h)
+      apply(erule_tac exE, rule_tac x = stp in exI, 
+        simp add: replicate.simps[THEN sym]
+        replicate_add[THEN sym] del: replicate.simps)
+      apply(subgoal_tac "pstr >(length gs)", simp)
+      apply(subgoal_tac "ba > aa \<and> length gs = aa \<and> pstr \<ge> ba", simp)
+      apply(simp add: j1 j2 j3)
+      done     
+  qed
+qed
+
+lemma restore_rs_prog_ex:
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+  rec_ci f = (a, aa, ba);
+  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+  (map rec_ci (f # gs)))) = pstr;
+  ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+  8 * length gs + 3 * n + length a + 3\<rbrakk>
+  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
+                                           bp = mv_box 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 (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
+      mv_boxes 0 (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n)
+        \<circ> rec_ci) ` set gs))) + length gs)) n [+]
+     mv_boxes (max (Suc n) (Max (insert ba 
+      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
+     a [+] recursive.mv_box aa (max (Suc n)
+       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> 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 (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) 
+        + length gs)) 0 n" 
+  in exI, auto simp: abc_append_commute)
+done
+
+lemma exp_add: "a\<up>(b+c) = a\<up>b @ a\<up>c"
+apply(simp add:replicate_add)
+done
+
+lemma [simp]: "n < pstr \<Longrightarrow> (0\<up>pstr)[n := rs] @ [0::nat] = 0\<up>n @ rs # 0\<up>(pstr - n)"
+using list_update_length[of "0\<up>n" "0::nat" "0\<up>(pstr - Suc n)" rs]
+apply(simp add: replicate_Suc[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" 
+  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+  "length ys = length gs"
+  "rec_calc_rel f ys rs" 
+  "rec_ci f = (a, aa, ba)" 
+  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                                        (map rec_ci (f # gs))))"
+  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+                              8 * length gs + 3 * n + length a + 3" 
+  shows "\<exists>stp. abc_steps_l
+           (ss, 0\<up>pstr @ rs # 0\<up>length ys  @ lm @
+                    0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
+  (ss + 3, 0\<up>n @ rs # 0\<up>(pstr - n) @ 0\<up>length ys  @ lm @ 
+                                   0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+ from h and pdef and starts have k1:
+   "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
+                                            bp = mv_box 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 \<and> length ap = ss \<and> 
+                                 bp = recursive.mv_box pstr n"
+   thus"?thesis"
+     apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
+     apply(insert mv_box_ex[of pstr n "0\<up>pstr @ rs # 0\<up>length gs @
+                     lm @ 0\<up>(a_md - Suc (pstr + length gs + n)) @ 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 \<noteq> [] \<Longrightarrow> length xs \<ge> Suc 0"
+by(case_tac xs, auto)
+
+lemma  [simp]: "n < max (Suc n) (max ba (Max (((\<lambda>(aprog, p, n). n) o 
+                                                  rec_ci) ` set gs)))"
+by(simp)
+
+lemma restore_paras_prog_ex: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+  rec_ci f = (a, aa, ba);
+  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                          (map rec_ci (f # gs)))) = pstr;
+  ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+                         8 * length gs + 3 * n + length a + 6\<rbrakk>
+  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
+                      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 (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
+      [+] mv_boxes 0 (Suc (max (Suc n) 
+       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) 
+     + length gs)) n [+] mv_boxes (max (Suc n) 
+    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
+     a [+] recursive.mv_box aa (max (Suc n) 
+      (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+     empty_boxes (length gs) [+]
+     recursive.mv_box (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> 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" 
+  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+  "length ys = length gs"
+  "rec_calc_rel f ys rs" 
+  "rec_ci f = (a, aa, ba)"
+  and pdef: 
+  "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                         (map rec_ci (f # gs))))"
+  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+                              8 * length gs + 3 * n + length a + 6" 
+  shows "\<exists>stp. abc_steps_l (ss, 0\<up>n @ rs # 0\<up>(pstr - n+ length ys) @
+                         lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)
+  aprog stp = (ss + 3 * n, lm @ rs # 0\<up>(a_md - Suc n) @ suf_lm)"
+proof -
+  thm rec_ci.simps
+  from h and pdef and starts have k1:
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and>
+                     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 \<and> length ap = ss \<and> 
+                              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\<up>(pstr - n + length gs)" "lm" 
+        "0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp)
+      apply(subgoal_tac "pstr > n \<and> 
+        a_md > pstr + length gs + n \<and> 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:
+  "\<lbrakk>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)\<rbrakk>
+  \<Longrightarrow> length aprog = (\<Sum>(ap, pos, n)\<leftarrow>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:
+  "\<And>x aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>x \<in> set (f # gs);
+  rec_ci x = (aprog, rs_pos, a_md);
+  rec_calc_rel x lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ 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 "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
+  = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ 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 (\<lambda>(aprog, p, n). n) 
+                                         (map rec_ci (f # gs)))))"  
+  let ?gs_len = "listsum (map (\<lambda> (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"
+    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)" 
+    "length ys = length gs" 
+    "rec_calc_rel f ys rs"
+    "n = length lm"
+    "rec_ci f = (a, b, c)"  
+  hence k1:
+    "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+    (?gs_len + 3 * length gs, lm @ 0\<up>(?pstr - n) @ ys @
+                               0\<up>(a_md - ?pstr - length ys) @ 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: 
+    "\<exists> stp. abc_steps_l (?gs_len + 3 * length gs,  lm @ 
+        0\<up>(?pstr - n) @ ys @ 0\<up>(a_md - ?pstr - length ys) @ suf_lm) aprog stp = 
+    (?gs_len + 3 * length gs + 3 * n, 0\<up>?pstr @ ys @ 0 # lm @ 
+                              0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)"
+    thm save_paras
+    apply(erule_tac ba = c in save_paras, auto intro: ci_cn_para_eq)
+    done
+  from g have k3: 
+    "\<exists> stp. abc_steps_l (?gs_len + 3 * length gs + 3 * n,
+    0\<up>?pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
+    (?gs_len + 6 * length gs + 3 * n,  
+           ys @ 0\<up>?pstr @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ 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: 
+    "\<exists>stp. abc_steps_l  (?gs_len + 6 * length gs + 3 * n,  
+    ys @ 0\<up>?pstr @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
+    (?gs_len + 6 * length gs + 3 * n + length a, 
+   ys @ rs # 0\<up>?pstr  @ lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ 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:
+    "\<exists> stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + length a,
+    ys @ rs # 0\<up>?pstr @ lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)
+    aprog stp =
+    (?gs_len + 6 * length gs + 3 * n + length a + 3,
+    ys @ 0\<up>(?pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
+    0\<up>(a_md  - Suc (?pstr + length ys + n)) @ suf_lm)"
+    apply(rule_tac save_rs, auto simp: h)
+    done
+  from g have k6: 
+    "\<exists> stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + 
+    length a + 3, ys @ 0\<up>(?pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
+    0\<up>(a_md  - Suc (?pstr + length ys + n)) @ suf_lm) 
+    aprog stp =
+    (?gs_len + 8 * length gs + 3 *n + length a + 3,
+    0\<up>?pstr @ rs # 0\<up>length ys @ lm @ 
+                        0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm)"
+    apply(drule_tac suf_lm = suf_lm in mv_box_paras, auto)
+    apply(rule_tac x = stp in exI, simp)
+    done
+  from g have k7: 
+    "\<exists> stp. abc_steps_l (?gs_len + 8 * length gs + 3 *n + 
+    length a + 3, 0\<up>?pstr  @ rs # 0\<up>length ys @ lm @ 
+    0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
+    (?gs_len + 8 * length gs + 3 * n + length a + 6, 
+    0\<up>n @ rs # 0\<up>(?pstr  - n) @ 0\<up>length ys @ lm @
+                        0\<up>(a_md -Suc (?pstr + length ys + n)) @ 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: "\<exists> stp. abc_steps_l (?gs_len + 8 * length gs + 
+    3 * n + length a + 6,
+    0\<up>n @ rs # 0\<up>(?pstr  - n) @ 0\<up>length ys @ lm @
+                      0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
+    (?gs_len + 8 * length gs + 6 * n + length a + 6,
+                           lm @ rs # 0\<up>(a_md - Suc n) @ 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 
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+    (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ 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 recursive_compile_correct:
+  "\<lbrakk>rec_ci recf = (ap, arity, fp);
+    rec_calc_rel recf args r\<rbrakk>
+  \<Longrightarrow> (\<exists> stp. (abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp) = 
+              (length ap, args@[r]@0\<up>(fp - arity - 1) @ 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:
+    "\<And>ap fp arity r anything args.
+    \<lbrakk>aa = ap \<and> ba = arity \<and> ca = fp; rec_calc_rel f args r\<rbrakk> \<Longrightarrow> 
+    \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ r # 0\<up>(fp - Suc arity) @ anything)"
+    and g_ind:
+    "\<And>x xa y xb ya ap fp arity r anything args.
+    \<lbrakk>x = (aa, ba, ca); xa = aa \<and> y = (ba, ca); xb = ba \<and> ya = ca; 
+    a = ap \<and> b = arity \<and> c = fp; rec_calc_rel g args r\<rbrakk>
+    \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ r # 0\<up>(fp - Suc arity) @ 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: 
+    "\<And> args r anything. rec_calc_rel f args r \<Longrightarrow> 
+    \<exists>stp. abc_steps_l (0, args @ 0\<up>(ca - ba) @ anything) aa stp = 
+    (length aa, args @ r # 0\<up>(ca - Suc ba) @ anything)"
+    and ng_ind: 
+    "\<And> args r anything. rec_calc_rel g args r \<Longrightarrow> 
+    \<exists>stp. abc_steps_l (0, args @ 0\<up>(c - b) @ anything) a stp = 
+         (length a, args @ r # 0\<up>(c - Suc b)  @ 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 
+    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp = 
+    (length ap, args @ r # 0\<up>(fp - Suc arity) @ 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 "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ 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 
+    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ 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 "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp 
+    = (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+    by(erule_tac id_case)
+next
+  fix n f gs ap fp arity r anything args
+  assume ind: "\<And>x ap fp arity r anything args.
+    \<lbrakk>x \<in> set (f # gs); 
+    rec_ci x = (ap, arity, fp); 
+    rec_calc_rel x args r\<rbrakk>
+    \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+  and h: "rec_ci (Cn n f gs) = (ap, arity, fp)" 
+    "rec_calc_rel (Cn n f gs) args r"
+  from h show
+    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) 
+       ap stp = (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+    apply(rule_tac cn_case, rule_tac ind, auto)
+    done
+next
+  fix n f ap fp arity r anything args
+  assume ind:
+    "\<And>ap fp arity r anything args.
+    \<lbrakk>rec_ci f = (ap, arity, fp); rec_calc_rel f args r\<rbrakk> \<Longrightarrow> 
+    \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+  and h: "rec_ci (Mn n f) = (ap, arity, fp)" 
+    "rec_calc_rel (Mn n f) args r"
+  from h show 
+    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp = 
+              (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+    apply(rule_tac mn_case, rule_tac ind, auto)
+    done    
+qed
+
+lemma abc_append_uhalt1:
+  "\<lbrakk>\<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp);
+    p = ap [+] bp [+] cp\<rbrakk>
+  \<Longrightarrow> \<forall> stp. (\<lambda> (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:
+  "\<lbrakk>abc_steps_l (0, am) ap stp = (length ap, lm); bp \<noteq> [];
+  \<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp);
+  p = ap [+] bp [+] cp\<rbrakk>
+  \<Longrightarrow> \<forall> stp. (\<lambda> (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 \<noteq> []"
+    "\<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp)"
+    "p = ap [+] bp [+] cp"
+  have "\<exists> 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: "\<forall> stp. (\<lambda> (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
+    "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
+                    (abc_steps_l (0, am) p (stpa + stp))"
+    apply(simp add: abc_steps_add)
+    done
+  thus "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
+                           (abc_steps_l (0, am) p stp)"
+    apply(rule_tac allI, auto)
+    apply(case_tac "stp \<ge>  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)" 
+                "\<not> stpa \<le> stp"
+    thus "a < length p"
+      using g1 h
+      apply(case_tac "a < length p", simp, simp)
+      apply(subgoal_tac "\<exists> 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: "\<forall> y. (\<exists> rs. rec_calc_rel f (lm @ [y]) rs \<and> rs \<noteq> 0)"
+  shows "\<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
+               aprog stp of (ss, e) \<Rightarrow> 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')" 
+            "\<forall>y. \<exists>rs. rec_calc_rel f (lm @ [y]) rs \<and> rs \<noteq> 0" "length lm = n"
+  thm mn_ind_step
+  have "\<exists>stpa \<ge> stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog stpa 
+         = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+  proof(induct stp, auto)
+    show "\<exists>stpa. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
+          aprog stpa = (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+      apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
+      done
+  next
+    fix stp stpa
+    assume g1: "stp \<le> stpa"
+      and g2: "abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+                            aprog stpa
+               = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    have "\<exists>rs. rec_calc_rel f (lm @ [stp]) rs \<and> rs \<noteq> 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 \<and> rs \<noteq> 0" ..
+    hence "\<exists> stpb. abc_steps_l (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @
+                     suf_lm) aprog stpb 
+      = (0, lm @ Suc stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+      using h
+      apply(rule_tac mn_ind_step)
+      apply(rule_tac recursive_compile_correct, 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 "\<exists>stpa\<ge>Suc stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @
+                 suf_lm) aprog stpa 
+      = (0, lm @ Suc stp # 0\<up>(a_md - Suc rs_pos) @ 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 \<le> stpa \<and> abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+         aprog stpa = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)" ..
+  thus "case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
+    of (ss, e) \<Rightarrow> ss < length aprog"
+    apply(case_tac "abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog
+      stp", simp, case_tac "a \<ge> length aprog", 
+        simp, simp)
+    apply(subgoal_tac "\<exists> d. stpa = stp + d", erule_tac exE)
+    apply(subgoal_tac "lm @ 0\<up>(a_md - rs_pos) @ suf_lm = lm @ 0 # 
+             0\<up>(a_md - Suc rs_pos) @ 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: replicate_Suc[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!]: 
+  "\<lbrakk>ap = bp; cp = dp\<rbrakk> \<Longrightarrow> ap [+] cp = bp [+] dp"
+by simp 
+
+lemma cn_merge_gs_split: 
+  "\<lbrakk>i < length gs; rec_ci (gs!i) = (ga, gb, gc)\<rbrakk> \<Longrightarrow> 
+     cn_merge_gs (map rec_ci gs) p = 
+        cn_merge_gs (map rec_ci (take i gs)) p [+] ga [+] 
+       mv_box 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 \<dots>gi, gi+1, \<dots> gn"}, this lemma describes what 
+  happens when every one of @{text "g1, g2, \<dots> gi"} terminates, but 
+  @{text "gi+1"} does not terminate, so that whole function @{text "Cn f g1 g2 \<dots>gi, gi+1, \<dots> 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: "\<forall> j < i. (\<exists> rs. rec_calc_rel (gs!j) lm rs)" 
+  and unhalt_condition: "\<And> slm. \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(gc - gb) @ slm) 
+     ga stp of (se, e) \<Rightarrow> se < length ga"
+  shows " \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) aprog
+  stp of (ss, e) \<Rightarrow> 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: "\<forall>j<i. \<exists>rs. rec_calc_rel (gs ! j) lm rs"
+    "i < length gs"
+  and ind:
+    "\<And> slm. \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(gc - n) @ slm) ga stp of (se, e) \<Rightarrow> se < length ga"
+  have h3: "rs_pos = n"
+    using h1
+    by(rule_tac ci_cn_para_eq, simp)
+  let ?ggs = "take i gs"
+  have "\<exists> ys. (length ys = i \<and> 
+    (\<forall> 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 \<and> (\<forall> k < i. rec_calc_rel (?ggs ! k)
+                        lm (ys ! k)))" ..
+  let ?pstr = "Max (set (Suc n # c # map (\<lambda>(aprog, p, n). n)
+    (map rec_ci (f # gs))))"
+  have "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suflm) 
+    (cn_merge_gs (map rec_ci ?ggs) ?pstr) stp =
+    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) ?ggs) +
+    3 * length ?ggs, lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @
+    suflm) "
+    apply(rule_tac  cn_merge_gs_ex)
+    apply(rule_tac  recursive_compile_correct, 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 \<in> 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\<up>(a_md - n) @ suflm) 
+    (cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa =
+    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) ?ggs) +
+    3 * length ?ggs, lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @
+    suflm)" ..
+  moreover have 
+    "\<exists> cp. aprog = (cn_merge_gs
+    (map rec_ci ?ggs) ?pstr) [+] ga [+] cp"
+    using h1
+    apply(simp add: rec_ci.simps)
+    apply(rule_tac x = "mv_box 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 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) +
+      length gs)) n [+] mv_boxes (max (Suc n) (Max (insert c 
+      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
+      a [+] recursive.mv_box b (max (Suc n) 
+      (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+     empty_boxes (length gs) [+] recursive.mv_box (max (Suc n) 
+      (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+      mv_boxes (Suc (max (Suc n) (Max (insert c 
+    (((\<lambda>(aprog, p, n). n) \<circ> 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 (((\<lambda>(aprog, p, n). n) \<circ> 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 "\<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) 
+    aprog stp of (ss, e) \<Rightarrow> ss < length aprog"
+  proof(rule_tac abc_append_unhalt2)
+    show "abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) (
+      cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa =
+         (length ((cn_merge_gs (map rec_ci ?ggs) ?pstr)),  
+          lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @ suflm)"
+      using h3 g2
+      apply(simp add: cn_merge_gs_length)
+      done
+  next
+    show "ga \<noteq> []"
+      using h1
+      apply(simp add: rec_ci_not_null)
+      done
+  next
+    show "\<forall>stp. case abc_steps_l (0, lm @ 0\<up>(?pstr - n) @ ys
+      @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm) ga  stp of
+          (ss, e) \<Rightarrow> ss < length ga"
+      using ind[of "0\<up>(?pstr - gc) @ ys @ 0\<up>(a_md - (?pstr + length (take i gs)))
+        @ suflm"]
+      apply(subgoal_tac "lm @ 0\<up>(?pstr - n) @ ys
+        @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm
+                       = lm @ 0\<up>(gc - n) @ 
+        0\<up>(?pstr - gc) @ ys @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm", simp)
+      apply(simp add: replicate_add[THEN sym])
+      apply(subgoal_tac "gc > n \<and> ?pstr \<ge> 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 recursive_compile_correct_spec: 
+  "\<lbrakk>rec_ci re = (ap, ary, fp); 
+    rec_calc_rel re args r\<rbrakk>
+  \<Longrightarrow> (\<exists> stp. (abc_steps_l (0, args @ 0\<up>(fp - ary)) ap stp) = 
+                     (length ap, args@[r]@0\<up>(fp - ary - 1)))"
+using recursive_compile_correct[of re ap ary fp args r "[]"]
+by simp
+
+definition dummy_abc :: "nat \<Rightarrow> abc_inst list"
+where
+"dummy_abc k = [Inc k, Dec k 0, Goto 3]"
+
+definition abc_list_crsp:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"
+  where
+  "abc_list_crsp xs ys = (\<exists> n. xs = ys @ 0\<up>n \<or> ys = xs @ 0\<up>n)"
+
+lemma [intro]: "abc_list_crsp (lm @ 0\<up>m) lm"
+apply(auto simp: abc_list_crsp_def)
+done
+
+lemma abc_list_crsp_lm_v: 
+  "abc_list_crsp lma lmb \<Longrightarrow> abc_lm_v lma n = abc_lm_v lmb n"
+apply(auto simp: abc_list_crsp_def abc_lm_v.simps 
+                 nth_append)
+done
+
+lemma  rep_app_cons_iff: 
+  "k < n \<Longrightarrow> 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 \<Longrightarrow> 
+      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)
+proof -
+  fix na
+  assume h: "m < length lmb + na" " \<not> 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 "\<exists>nb. replicate na 0[m - length lmb := n] =
+                 replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
+                 replicate (m - length lmb) 0 @ [n] =
+                 replicate na 0[m - length lmb := n] @ replicate nb 0"
+    apply(auto)
+    done
+next
+  fix na
+  assume h: "\<not> m < length lmb + na"
+  show 
+    "\<exists>nb. replicate na 0 @ replicate (m - (length lmb + na)) 0 @ [n] =
+           replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
+          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: "\<not> 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 "\<exists>nb. replicate (m - length lma) 0 @ [n] =
+                 replicate na 0[m - length lma := n] @ replicate nb 0 
+           \<or> replicate na 0[m - length lma := n] =
+                 replicate (m - length lma) 0 @ n # replicate nb 0"
+    apply(auto)
+    done
+next
+  fix na
+  assume "\<not> m < length lma + na"
+  thus " \<exists>nb. replicate (m - length lma) 0 @ [n] =
+            replicate na 0 @
+            replicate (m - (length lma + na)) 0 @ n # replicate nb 0 
+        \<or>   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: 
+  "\<lbrakk>abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); 
+    abc_step_l (aa, lmb) i = (a', lmb')\<rbrakk>
+    \<Longrightarrow> a' = a \<and> abc_list_crsp lma' lmb'"
+apply(case_tac i, auto simp: abc_step_l.simps 
+       abc_list_crsp_lm_s abc_list_crsp_lm_v Let_def 
+                       split: abc_inst.splits if_splits)
+done
+
+lemma abc_list_crsp_steps: 
+  "\<lbrakk>abc_steps_l (0, lm @ 0\<up>m) aprog stp = (a, lm'); aprog \<noteq> []\<rbrakk> 
+      \<Longrightarrow> \<exists> lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and> 
+                                          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\<up>m) aprog stp", 
+      simp add: abc_step_red)
+proof -
+  fix stp a lm' aa b
+  assume ind:
+    "\<And>a lm'. aa = a \<and> b = lm' \<Longrightarrow> 
+     \<exists>lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and>
+                                          abc_list_crsp lm' lma"
+    and h: "abc_steps_l (0, lm @ 0\<up>m) aprog (Suc stp) = (a, lm')" 
+           "abc_steps_l (0, lm @ 0\<up>m) aprog stp = (aa, b)" 
+           "aprog \<noteq> []"
+  hence g1: "abc_steps_l (0, lm @ 0\<up>m) aprog (Suc stp)
+          = abc_step_l (aa, b) (abc_fetch aa aprog)"
+    apply(rule_tac abc_step_red, simp)
+    done
+  have "\<exists>lma. abc_steps_l (0, lm) aprog stp = (aa, lma) \<and> 
+              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) \<and> 
+     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_step_red, simp)
+    done
+  show "\<exists>lma. abc_steps_l (0, lm) aprog (Suc stp) = (a, lma) \<and> abc_list_crsp lm' lma"
+    using g1 g2 g3 h
+    apply(auto)
+    apply(case_tac "abc_step_l (aa, b) (abc_fetch aa aprog)",
+          case_tac "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp)
+    apply(rule_tac abc_list_crsp_step, auto)
+    done
+qed
+
+lemma recursive_compile_correct_norm: 
+  "\<lbrakk>rec_ci re = (aprog, rs_pos, a_md);  
+   rec_calc_rel re lm rs\<rbrakk>
+  \<Longrightarrow> (\<exists> stp lm' m. (abc_steps_l (0, lm) aprog stp) = 
+  (length aprog, lm') \<and> abc_list_crsp lm' (lm @ rs # 0\<up>m))"
+apply(frule_tac recursive_compile_correct_spec, 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) \<noteq> []"
+apply(simp add: dummy_abc_def)
+done
+
+lemma dummy_abc_steps_ex: 
+  "\<exists>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 [simp]: 
+  "\<lbrakk>Suc (length lm) - length lm' \<le> n; \<not> length lm < length lm'; lm @ rs # 0 \<up> m = lm' @ 0 \<up> n\<rbrakk> 
+  \<Longrightarrow> lm' @ 0 \<up> Suc (length lm - length lm') = lm @ [rs]"
+apply(subgoal_tac "n > m")
+apply(subgoal_tac "\<exists> d. n = d + m", erule_tac exE)
+apply(simp add: replicate_add)
+apply(drule_tac length_equal, simp)
+apply(simp add: replicate_Suc[THEN sym] del: replicate_Suc)
+apply(rule_tac x = "n - m" in exI, simp)
+apply(drule_tac length_equal, simp)
+done
+
+lemma [elim]: 
+  "lm @ rs # 0\<up>m = lm' @ 0\<up>n \<Longrightarrow> 
+  \<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) = 
+                            lm @ rs # 0\<up>m"
+proof(cases "length lm' > length lm")
+  case True 
+  assume h: "lm @ rs # 0\<up>m = lm' @ 0\<up>n" "length lm < length lm'"
+  hence "m \<ge> n"
+    apply(drule_tac length_equal)
+    apply(simp)
+    done
+  hence "\<exists> 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 
+                     replicate_add)
+    done
+next
+  case False
+  assume h:"lm @ rs # 0\<up>m = lm' @ 0\<up>n" 
+    and    g: "\<not> length lm < length lm'"
+  have "take (Suc (length lm)) (lm @ rs # 0\<up>m) = 
+                        take (Suc (length lm)) (lm' @ 0\<up>n)"
+    using h by simp
+  moreover have "n \<ge> (Suc (length lm) - length lm')"
+    using h g
+    apply(drule_tac length_equal)
+    apply(simp)
+    done
+  ultimately show 
+    "\<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) =
+                                                       lm @ rs # 0\<up>m"
+    using g h
+    apply(simp add: abc_lm_s.simps abc_lm_v.simps 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\<up>m)
+  \<Longrightarrow> \<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) 
+             = lm @ rs # 0\<up>m"
+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: replicate_add)
+done
+
+lemma abc_append_dummy_complie:
+  "\<lbrakk>rec_ci recf = (ap, ary, fp);  
+    rec_calc_rel recf args r; 
+    length args = k\<rbrakk>
+  \<Longrightarrow> (\<exists> stp m. (abc_steps_l (0, args) (ap [+] dummy_abc k) stp) = 
+                  (length ap + 3, args @ r # 0\<up>m))"
+apply(drule_tac recursive_compile_correct_norm, 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\<up>m)"
+  thm abc_append_exc2
+  thm abc_lm_s.simps
+  have "\<exists>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 "\<exists>stp m. abc_steps_l (0, args) (ap [+] 
+    dummy_abc (length args)) stp = (Suc (Suc (Suc (length ap))), args @ r # 0\<up>m)"
+    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 \<Longrightarrow> abc_lm_v (args @ r # 0\<up>m) k = r "
+apply(simp add: abc_lm_v.simps nth_append)
+done
+
+lemma [simp]: "crsp (layout_of (ap [+] dummy_abc k)) (0, args)
+  (Suc 0, Bk # Bk # ires, <args> @ Bk \<up> rn) ires"
+apply(auto simp: crsp.simps start_of.simps)
+done
+
+lemma recursive_compile_to_tm_correct: 
+  "\<lbrakk>rec_ci recf = (ap, ary, fp); 
+    rec_calc_rel recf args r;
+    length args = k;
+    ly = layout_of (ap [+] dummy_abc k);
+    tp = tm_of (ap [+] dummy_abc k)\<rbrakk>
+  \<Longrightarrow> \<exists> stp m l. steps0 (Suc 0, Bk # Bk # ires, <args> @ Bk\<up>rn)
+  (tp @ shift (mopup k) (length tp div 2)) stp
+  = (0, Bk\<up>m @ Bk # Bk # ires, Oc\<up>Suc r @ Bk\<up>l)"
+  using abc_append_dummy_complie[of recf ap ary fp args r k]
+apply(simp)
+apply(erule_tac exE)+
+apply(frule_tac tp = tp and n = k 
+               and ires = ires in compile_correct_halt, simp_all add: length_append)
+apply(simp_all add: length_append)
+done
+
+lemma [simp]:
+  "list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))) xs \<Longrightarrow>
+  list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))))) xs"
+apply(induct xs, simp, simp)
+apply(case_tac a, simp)
+done
+
+lemma shift_append: "shift (xs @ ys) n = shift xs n @ shift ys n"
+apply(simp add: shift.simps)
+done
+
+lemma [simp]: "length (shift (mopup n) ss) = 4 * n + 12"
+apply(auto simp: mopup.simps shift_append mopup_b_def)
+done
+
+lemma length_tm_even[intro]: "length (tm_of ap) mod 2 = 0"
+apply(simp add: tm_of.simps)
+done
+
+lemma [simp]: "k < length ap \<Longrightarrow> 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 start_of_suc_inc:
+  "\<lbrakk>k < length ap; ap ! k = Inc n\<rbrakk> \<Longrightarrow> start_of (layout_of ap) (Suc k) =
+                        start_of (layout_of ap) k + 2 * n + 9"
+apply(rule_tac start_of_Suc1, auto simp: abc_fetch.simps)
+done
+
+lemma start_of_suc_dec:
+  "\<lbrakk>k < length ap; ap ! k = (Dec n e)\<rbrakk> \<Longrightarrow> start_of (layout_of ap) (Suc k) =
+                        start_of (layout_of ap) k + 2 * n + 16"
+apply(rule_tac start_of_Suc2, auto simp: abc_fetch.simps)
+done
+
+lemma inc_state_all_le:
+  "\<lbrakk>k < length ap; ap ! k = Inc n; 
+       (a, b) \<in> set (shift (shift tinc_b (2 * n)) 
+                            (start_of (layout_of ap) k - Suc 0))\<rbrakk>
+       \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
+apply(subgoal_tac "b \<le> start_of (layout_of ap) (Suc k)")
+apply(subgoal_tac "start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap) ")
+apply(arith)
+apply(case_tac "Suc k = length ap", simp)
+apply(rule_tac start_of_less, simp)
+apply(auto simp: tinc_b_def shift.simps start_of_suc_inc length_of.simps startof_not0)
+done
+
+lemma findnth_le[elim]: 
+  "(a, b) \<in> set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))
+  \<Longrightarrow> b \<le> Suc (start_of (layout_of ap) k + 2 * n)"
+apply(induct n, simp add: findnth.simps shift.simps)
+apply(simp add: findnth.simps shift_append, auto)
+apply(auto simp: shift.simps)
+done
+
+lemma findnth_state_all_le1:
+  "\<lbrakk>k < length ap; ap ! k = Inc n;
+  (a, b) \<in> 
+  set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\<rbrakk> 
+  \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
+apply(subgoal_tac "b \<le> start_of (layout_of ap) (Suc k)")
+apply(subgoal_tac "start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap) ")
+apply(arith)
+apply(case_tac "Suc k = length ap", simp)
+apply(rule_tac start_of_less, simp)
+apply(subgoal_tac "b \<le> start_of (layout_of ap) k + 2*n + 1 \<and> 
+     start_of (layout_of ap) k + 2*n + 1 \<le>  start_of (layout_of ap) (Suc k)", auto)
+apply(auto simp: tinc_b_def shift.simps length_of.simps startof_not0 start_of_suc_inc)
+done
+
+lemma start_of_eq: "length ap < as \<Longrightarrow> 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 \<le> start_of (layout_of ap) (length ap)"
+apply(subgoal_tac "as > length ap \<or> as = length ap \<or> as < length ap", 
+      auto simp: start_of_eq start_of_less)
+done
+
+lemma findnth_state_all_le2: 
+  "\<lbrakk>k < length ap; 
+  ap ! k = Dec n e;
+  (a, b) \<in> set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\<rbrakk>
+  \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
+apply(subgoal_tac "b \<le> start_of (layout_of ap) k + 2*n + 1 \<and> 
+     start_of (layout_of ap) k + 2*n + 1 \<le>  start_of (layout_of ap) (Suc k) \<and>
+      start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap)", auto)
+apply(subgoal_tac "start_of (layout_of ap) (Suc k) = 
+  start_of  (layout_of ap)  k + 2*n + 16", simp)
+apply(simp add: start_of_suc_dec)
+apply(rule_tac start_of_all_le)
+done
+
+lemma dec_state_all_le[simp]:
+  "\<lbrakk>k < length ap; ap ! k = Dec n e; 
+  (a, b) \<in> set (shift (shift tdec_b (2 * n))
+  (start_of (layout_of ap) k - Suc 0))\<rbrakk>
+       \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
+apply(subgoal_tac "2*n + start_of (layout_of ap) k + 16 \<le> start_of (layout_of ap) (length ap) \<and> start_of (layout_of ap) k > 0")
+prefer 2
+apply(subgoal_tac "start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2*n + 16
+                 \<and> start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap)")
+apply(simp add: startof_not0, rule_tac conjI)
+apply(simp add: start_of_suc_dec)
+apply(rule_tac start_of_all_le)
+apply(auto simp: tdec_b_def shift.simps)
+done
+
+lemma tms_any_less: 
+  "\<lbrakk>k < length ap; (a, b) \<in> set (tms_of ap ! k)\<rbrakk> \<Longrightarrow> 
+  b \<le> start_of (layout_of ap) (length ap)"
+apply(case_tac "ap!k", auto simp: tms_of.simps tpairs_of.simps ci.simps shift_append sete.simps)
+apply(erule_tac findnth_state_all_le1, simp_all)
+apply(erule_tac inc_state_all_le, simp_all)
+apply(erule_tac findnth_state_all_le2, simp_all)
+apply(rule_tac start_of_all_le)
+apply(rule_tac dec_state_all_le, simp_all)
+apply(rule_tac start_of_all_le)
+done
+
+lemma concat_in: "i < length (concat xs) \<Longrightarrow> \<exists>k < length xs. concat xs ! i \<in> 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
+
+declare length_concat[simp]
+
+lemma in_tms: "i < length (tm_of ap) \<Longrightarrow> \<exists> k < length ap. (tm_of ap ! i) \<in> set (tms_of ap ! k)"
+apply(simp only: tm_of.simps)
+using concat_in[of i "tms_of ap"]
+apply(auto)
+done
+
+lemma all_le_start_of: "list_all (\<lambda>(acn, s). 
+  s \<le> start_of (layout_of ap) (length ap)) (tm_of ap)"
+apply(simp only: list_all_length)
+apply(rule_tac allI, rule_tac impI)
+apply(drule_tac in_tms, auto elim: tms_any_less)
+done
+
+lemma length_ci: 
+"\<lbrakk>k < length ap; length (ci ly y (ap ! k)) = 2 * qa\<rbrakk>
+      \<Longrightarrow> layout_of ap ! k = qa"
+apply(case_tac "ap ! k")
+apply(auto simp: layout_of.simps ci.simps 
+  length_of.simps tinc_b_def tdec_b_def length_findnth sete.simps)
+done
+
+lemma [intro]: "length (ci ly y i) mod 2 = 0"
+apply(case_tac i, auto simp: ci.simps length_findnth
+  tinc_b_def sete.simps tdec_b_def)
+done
+
+lemma [intro]: "listsum (map (length \<circ> (\<lambda>(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) \<le> length ap \<Longrightarrow>
+  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: "\<And>(ap::'a list). length ys \<le> length ap \<Longrightarrow> 
+    zip ys ap = zip ys (take (length ys) ap)"
+  and h: "length (a # ys) \<le> 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 length_start_of_tm: "start_of (layout_of ap) (length ap) = Suc (length (tm_of ap)  div 2)"
+using tpa_states[of "tm_of ap"  "length ap" ap]
+apply(simp add: tm_of.simps)
+done
+
+lemma [elim]: "list_all (\<lambda>(acn, s). s \<le> Suc q) xs
+        \<Longrightarrow> list_all (\<lambda>(acn, s). s \<le> q + (2 * n + 6)) xs"
+apply(simp add: list_all_length)
+apply(auto)
+done
+
+lemma [simp]: "length mopup_b = 12"
+apply(simp add: mopup_b_def)
+done
+(*
+lemma [elim]: "\<lbrakk>na < 4 * n; tshift (mop_bef n) q ! na = (a, b)\<rbrakk> \<Longrightarrow> 
+  b \<le> 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 \<or> na = 1 + 4*n \<or> na = 2 + 4*n \<or> na = 3 + 4*n",
+  auto simp: shift_length)
+apply(simp_all add: tshift.simps)
+done
+*)
+
+lemma mp_up_all_le: "list_all  (\<lambda>(acn, s). s \<le> 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 [simp]: "(a, b) \<in> set (mopup_a n) \<Longrightarrow> b \<le> 2 * n + 6"
+apply(induct n, auto simp: mopup_a.simps)
+done
+
+lemma [simp]: "(a, b) \<in> set (shift (mopup n) (listsum (layout_of ap)))
+  \<Longrightarrow> b \<le> (2 * listsum (layout_of ap) + length (mopup n)) div 2"
+apply(auto simp: mopup.simps shift_append shift.simps)
+apply(auto simp: mopup_a.simps mopup_b_def)
+done
+
+lemma [intro]: " 2 \<le> 2 * listsum (layout_of ap) + length (mopup n)"
+apply(simp add: mopup.simps)
+done
+
+lemma [intro]: " (2 * listsum (layout_of ap) + length (mopup n)) mod 2 = 0"
+apply(auto simp: mopup.simps)
+apply arith
+done
+
+lemma [simp]: "b \<le> Suc x
+          \<Longrightarrow> b \<le> (2 * x + length (mopup n)) div 2"
+apply(auto simp: mopup.simps)
+done
+
+lemma t_compiled_correct: 
+  "\<lbrakk>tp = tm_of ap; ly = layout_of ap; mop_ss = start_of ly (length ap)\<rbrakk> \<Longrightarrow> 
+    tm_wf (tp @ shift( mopup n) (length tp div 2), 0)"
+  using length_start_of_tm[of ap] all_le_start_of[of ap]
+apply(auto simp: tm_wf.simps List.list_all_iff)
+done
+
+end
+
+    
+  
+
+
+  
+
--- a/thys/uncomputable.thy	Wed Jan 23 17:02:23 2013 +0100
+++ b/thys/uncomputable.thy	Wed Jan 23 20:18:40 2013 +0100
@@ -25,8 +25,6 @@
   and "5 = Suc 4" 
   and "6 = Suc 5" by arith+
 
-
-
 text {*
   The {\em Copying} TM, which duplicates its input. 
 *}
@@ -61,11 +59,11 @@
 
 fun inv_init1 :: "nat \<Rightarrow> tape \<Rightarrow> bool"
   where
-   "inv_init1 x (l, r) = (l = [] \<and> r = Oc \<up> x )"
+   "inv_init1 x (l, r) = (l = [] \<and> r = Oc \<up> x)"
 
 fun inv_init2 :: "nat \<Rightarrow> tape \<Rightarrow> bool"
   where 
-  "inv_init2 x (l, r) = (\<exists> i j. i > 0 \<and> i + j = x \<and> l = Oc \<up> i \<and> r = Oc\<up>j)"
+  "inv_init2 x (l, r) = (\<exists> i j. i > 0 \<and> i + j = x \<and> l = Oc \<up> i \<and> r = Oc \<up> j)"
 
 fun inv_init3 :: "nat \<Rightarrow> tape \<Rightarrow> bool"
   where
@@ -77,14 +75,14 @@
 
 fun inv_init0 :: "nat \<Rightarrow> tape \<Rightarrow> bool"
   where
-  "inv_init0 x (l, r) = ((x >  Suc 0 \<and> l = Oc\<up>(x - 2) \<and> r = [Oc, Oc, Bk, Oc]) \<or>   
-                        (x = 1 \<and> l = [] \<and> r = [Bk, Oc, Bk, Oc]))"
+  "inv_init0 x (l, r) = ((x > 1 \<and> l = Oc \<up> (x - 2) \<and> r = [Oc, Oc, Bk, Oc]) \<or>   
+                         (x = 1 \<and> l = [] \<and> r = [Bk, Oc, Bk, Oc]))"
 
 fun inv_init :: "nat \<Rightarrow> config \<Rightarrow> bool"
   where
-  "inv_init x (s, l, r) = (
-         if s = 0 then inv_init0 x (l, r) 
-         else if s = Suc 0 then inv_init1 x (l, r)
+  "inv_init x (s, l, r) = 
+        (if s = 0 then inv_init0 x (l, r) 
+         else if s = 1 then inv_init1 x (l, r)
          else if s = 2 then inv_init2 x (l, r)
          else if s = 3 then inv_init3 x (l, r)
          else if s = 4 then inv_init4 x (l, r)
@@ -1077,6 +1075,13 @@
   where
   "dither \<equiv> [(W0, 1), (R, 2), (L, 1), (L, 0)] "
 
+(* invariants of dither *)
+abbreviation
+  "dither_halt_inv \<equiv> \<lambda>(l, r). (\<exists>nd. l = Bk \<up> nd) \<and> r = [Oc, Oc]"
+
+abbreviation
+  "dither_unhalt_inv \<equiv> \<lambda>(l, r). (\<exists>nd. l = Bk \<up> nd) \<and> r = [Oc]"
+
 lemma dither_unhalt_state: 
   "(steps0 (1, Bk \<up> m, [Oc]) dither stp = (1, Bk \<up> m, [Oc])) \<or> 
    (steps0 (1, Bk \<up> m, [Oc]) dither stp = (2, Oc # Bk \<up> m, []))"
@@ -1092,25 +1097,22 @@
 by auto
 
 lemma dither_loops:
-  shows "{(\<lambda>(l, r). (\<exists>nd. l = Bk \<up> nd) \<and> r = [Oc])} dither \<up>" 
+  shows "{dither_unhalt_inv} dither \<up>" 
 apply(rule Hoare_unhaltI)
 using dither_unhalt_rs
 apply(auto)
 done
 
 lemma dither_halt_rs: 
-  "\<exists>stp. steps0 (Suc 0, Bk \<up> m, [Oc, Oc]) dither stp = (0, Bk \<up> m, [Oc, Oc])"
+  "\<exists>stp. steps0 (1, Bk \<up> m, [Oc, Oc]) dither stp = (0, Bk \<up> m, [Oc, Oc])"
 unfolding dither_def
 apply(rule_tac x = "3" in exI)
 apply(simp add: steps.simps step.simps fetch.simps numeral)
 done 
 
-definition
-  "dither_halt_inv \<equiv> (\<lambda>(l, r). (\<exists>nd. l = Bk \<up> nd) \<and> r = [Oc, Oc])"
 
 lemma dither_halts:
   shows "{dither_halt_inv} dither {dither_halt_inv}" 
-unfolding dither_halt_inv_def
 apply(rule HoareI)
 using dither_halt_rs
 apply(auto)
@@ -1218,33 +1220,29 @@
   The TM @{text "H"} is the one which is assummed being able to solve the Halting problem.
   *}
   and H :: "instr list"
-  assumes h_wf[intro]: "tm_wf (H, 0)"
+  assumes h_wf[intro]: "tm_wf0 H"
   -- {*
   The following two assumptions specifies that @{text "H"} does solve the Halting problem.
   *}
   and h_case: 
-  "\<And> M n. \<lbrakk>(haltP (M, 0) lm)\<rbrakk> \<Longrightarrow> 
-             \<exists> na nb. (steps (Suc 0, [], <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc]))"
+  "\<And> M lm. (haltP0 M lm) \<Longrightarrow> \<exists> na nb. (steps0 (1, [], <code M # lm>) H na = (0, Bk \<up> nb, [Oc]))"
   and nh_case: 
-  "\<And> M n. \<lbrakk>(\<not> haltP (M, 0) lm)\<rbrakk> \<Longrightarrow>
-             \<exists> na nb. (steps (Suc 0, [],  <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc, Oc]))"
+  "\<And> M lm. (\<not> haltP0 M lm) \<Longrightarrow>
+             \<exists> na nb. (steps0 (1, [],  <code M # lm>) H na = (0, Bk \<up> nb, [Oc, Oc]))"
 begin
 
 lemma h_newcase: 
-  "\<And> M n. haltP (M, 0) lm \<Longrightarrow> 
-  \<exists> na nb. (steps (Suc 0, Bk\<up>x , <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc]))"
+  "\<And> M lm. haltP0 M lm \<Longrightarrow> \<exists> na nb. (steps0 (1, Bk \<up> x , <code M # lm>) H na = (0, Bk \<up> nb, [Oc]))"
 proof -
-  fix M n
+  fix M lm
   assume "haltP (M, 0) lm"
-  hence "\<exists> na nb. (steps (Suc 0, [], <code M # lm>) (H, 0) na
+  hence "\<exists> na nb. (steps (1, [], <code M # lm>) (H, 0) na
             = (0, Bk\<up>nb, [Oc]))"
-    apply(erule_tac h_case)
-    done
+    by (rule h_case)
   from this obtain na nb where 
-    cond1:"(steps (Suc 0, [], <code M # lm>) (H, 0) na
-            = (0, Bk\<up>nb, [Oc]))" by blast
-  thus "\<exists> na nb. (steps (Suc 0, Bk\<up>x, <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc]))"
-  proof(rule_tac x = na in exI, case_tac "steps (Suc 0, Bk\<up>x, <code M # lm>) (H, 0) na", simp)
+    cond1:"(steps (1, [], <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc]))" by blast
+  thus "\<exists> na nb. (steps (1, Bk\<up>x, <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc]))"
+  proof(rule_tac x = na in exI, case_tac "steps (1, Bk\<up>x, <code M # lm>) (H, 0) na", simp)
     fix a b c
     assume cond2: "steps (Suc 0, Bk\<up>x, <code M # lm>) (H, 0) na = (a, b, c)"
     have "tinres (Bk\<up>nb) b \<and> [Oc] = c \<and> 0 = a"
@@ -1254,32 +1252,29 @@
         apply(auto)
         done
     next
-      show "(steps (Suc 0, [], <code M # lm>) (H, 0) na
+      show "(steps (1, [], <code M # lm>) (H, 0) na
             = (0, Bk\<up>nb, [Oc]))"
-        by(simp add: cond1)
+        by(simp add: cond1[simplified])
     next
-      show "steps (Suc 0, Bk\<up>x, <code M # lm>) (H, 0) na = (a, b, c)"
-        by(simp add: cond2)
+      show "steps (1, Bk\<up>x, <code M # lm>) (H, 0) na = (a, b, c)"
+        by(simp add: cond2[simplified])
     qed
     thus "a = 0 \<and> (\<exists>nb. b = (Bk\<up>nb)) \<and> c = [Oc]"
       by(auto elim: tinres_ex1)
   qed
 qed
 
-lemma nh_newcase: "\<And> M n. \<lbrakk>\<not> (haltP (M, 0) lm)\<rbrakk> \<Longrightarrow> 
-             \<exists> na nb. (steps (Suc 0, Bk\<up>x, <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc, Oc]))"
+lemma nh_newcase: 
+  "\<And> M lm. \<not> (haltP (M, 0) lm) \<Longrightarrow> \<exists> na nb. (steps0 (1, Bk\<up>x, <code M # lm>) H na = (0, Bk\<up>nb, [Oc, Oc]))"
 proof -
-  fix M n
+  fix M lm
   assume "\<not> haltP (M, 0) lm"
-  hence "\<exists> na nb. (steps (Suc 0, [], <code M # lm>) (H, 0) na
-            = (0, Bk\<up>nb, [Oc, Oc]))"
-    apply(erule_tac nh_case)
-    done
+  hence "\<exists> na nb. (steps (1, [], <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc, Oc]))"
+    by (rule nh_case)
   from this obtain na nb where 
-    cond1: "(steps (Suc 0, [], <code M # lm>) (H, 0) na
-            = (0, Bk\<up>nb, [Oc, Oc]))" by blast
-  thus "\<exists> na nb. (steps (Suc 0, Bk\<up>x, <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc, Oc]))"
-  proof(rule_tac x = na in exI, case_tac "steps (Suc 0, Bk\<up>x, <code M # lm>) (H, 0) na", simp)
+    cond1: "(steps (1, [], <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc, Oc]))" by blast
+  thus "\<exists> na nb. (steps (1, Bk\<up>x, <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc, Oc]))"
+  proof(rule_tac x = na in exI, case_tac "steps (1, Bk\<up>x, <code M # lm>) (H, 0) na", simp)
     fix a b c
     assume cond2: "steps (Suc 0, Bk\<up>x, <code M # lm>) (H, 0) na = (a, b, c)"
     have "tinres (Bk\<up>nb) b \<and> [Oc, Oc] = c \<and> 0 = a"
@@ -1288,17 +1283,59 @@
         apply(auto simp: tinres_def)
         done
     next
-      show "(steps (Suc 0, [], <code M # lm>) (H, 0) na
-            = (0, Bk\<up>nb, [Oc, Oc]))"
-        by(simp add: cond1)
+      show "(steps (Suc 0, [], <code M # lm>) (H, 0) na = (0, Bk\<up>nb, [Oc, Oc]))"
+        by(simp add: cond1[simplified])
     next
       show "steps (Suc 0, Bk\<up>x, <code M # lm>) (H, 0) na = (a, b, c)"
-        by(simp add: cond2)
+        by(simp add: cond2[simplified])
     qed
     thus "a = 0 \<and> (\<exists>nb. b =  Bk\<up>nb) \<and> c = [Oc, Oc]"
       by(auto elim: tinres_ex1)
   qed
 qed
+
+
+(* invariants for H *)
+abbreviation
+  "pre_H_inv M n \<equiv> \<lambda>(l::cell list, r::cell list). (l = [Bk] \<and> r = <[code M, n]>)"
+
+abbreviation
+  "post_H_halt_inv \<equiv> \<lambda>(l, r). (\<exists>nd. l = Bk \<up> nd) \<and> r = [Oc, Oc]"
+
+abbreviation
+  "post_H_unhalt_inv \<equiv> \<lambda>(l, r). (\<exists>nd. l = Bk \<up> nd) \<and> r = [Oc]"
+
+
+lemma H_halt_inv:
+  assumes "\<not> haltP0 M [n]" 
+  shows "{pre_H_inv M n} H {post_H_halt_inv}"
+proof -
+  obtain stp i
+    where "steps0 (1, Bk \<up> 1, <[code M, n]>) H stp = (0, Bk \<up> i, [Oc, Oc])"
+    using nh_newcase[of "M" "[n]" "1", OF assms] by auto 
+  then show "{pre_H_inv M n} H {post_H_halt_inv}"
+    unfolding Hoare_def
+    apply(auto)
+    apply(rule_tac x = stp in exI)
+    apply(auto)
+    done
+qed
+
+lemma H_unhalt_inv:
+  assumes "haltP0 M [n]" 
+  shows "{pre_H_inv M n} H {post_H_unhalt_inv}"
+proof -
+  obtain stp i
+    where "steps0 (1, Bk \<up> 1, <[code M, n]>) H stp = (0, Bk \<up> i, [Oc])"
+    using h_newcase[of "M" "[n]" "1", OF assms] by auto 
+  then show "{pre_H_inv M n} H {post_H_unhalt_inv}"
+    unfolding Hoare_def
+    apply(auto)
+    apply(rule_tac x = stp in exI)
+    apply(auto)
+    done
+qed
+
    
 (* TM that produces the contradiction and its code *)
 abbreviation 
@@ -1314,7 +1351,7 @@
   (* invariants *)
   def P1 \<equiv> "\<lambda>(l::cell list, r::cell list). (l = [] \<and> r = <[code_tcontra]>)"
   def P2 \<equiv> "\<lambda>(l::cell list, r::cell list). (l = [Bk] \<and> r = <[code_tcontra, code_tcontra]>)"
-  def P3 \<equiv> dither_halt_inv
+  def P3 \<equiv> "\<lambda>(l, r). (\<exists>nd. l = Bk \<up> nd) \<and> r = [Oc, Oc]"
 
   (*
   {P1} tcopy {P2}  {P2} H {P3} 
@@ -1333,18 +1370,10 @@
     show "{P1} tcopy {P2}" unfolding P1_def P2_def
       by (rule tcopy_correct2)
   next
-    case B_halt
-    obtain n i
-      where "steps0 (1, Bk \<up> 1, <[code_tcontra, code_tcontra]>) H n = (0, Bk \<up> i, [Oc, Oc])"
-      using nh_newcase[of "tcontra" "[code_tcontra]" 1, OF assms]
-      by (auto)
-    then show "{P2} H {P3}"
-      unfolding P2_def P3_def dither_halt_inv_def
-      unfolding Hoare_def
-      apply(auto)
-      apply(rule_tac x = n in exI)
-      apply(simp)
-      done
+    case B_halt (* of H *)
+    show "{P2} H {P3}"
+      unfolding P2_def P3_def
+      using assms by (rule H_halt_inv)
   qed (simp)
 
   (* {P3} dither {P3} *)
@@ -1356,7 +1385,7 @@
     by (rule Hoare_plus_halt_simple[OF first second H_wf])
 
   with assms show "False"
-    unfolding P1_def P3_def dither_halt_inv_def 
+    unfolding P1_def P3_def
     unfolding haltP_def
     unfolding Hoare_def 
     apply(auto)
@@ -1364,8 +1393,8 @@
     apply(case_tac "steps0 (Suc 0, [], <[code tcontra]>) tcontra n")
     apply(simp, auto)
     apply(erule_tac x = 2 in allE, erule_tac x = 0 in allE)
-    apply(simp)
-    by (smt replicate_0 replicate_Suc)
+    apply(simp add: numeral)
+    done
 qed
 
 (* asumme tcontra halts on its code *)
@@ -1395,18 +1424,10 @@
     show "{P1} tcopy {P2}" unfolding P1_def P2_def
       by (rule tcopy_correct2)
   next
-    case B_halt
-    obtain n i
-      where "steps0 (1, Bk \<up> 1, <[code_tcontra, code_tcontra]>) H n = (0, Bk \<up> i, [Oc])"
-      using h_newcase[of "tcontra" "[code_tcontra]" 1, OF assms]
-      by (auto)
+    case B_halt (* of H *)
     then show "{P2} H {P3}"
       unfolding P2_def P3_def
-      unfolding Hoare_def
-      apply(auto)
-      apply(rule_tac x = n in exI)
-      apply(simp)
-      done
+      using assms by (rule H_unhalt_inv)
   qed (simp)
 
   (* {P3} dither loops *)
@@ -1418,8 +1439,8 @@
     by (rule Hoare_plus_unhalt_simple[OF first second H_wf])
 
   with assms show "False"
+    unfolding P1_def
     unfolding haltP_def
-    unfolding P1_def
     unfolding Hoare_unhalt_def
     by (auto, metis is_final_eq)  
 qed