# HG changeset patch # User Christian Urban # Date 1364377622 0 # Node ID d8e6f0798e235a171e1474160a69611e9d178c6e # Parent e9ef4ada308bc017defb223c35c943c478d8bb54 much simplified version of Recursive.thy diff -r e9ef4ada308b -r d8e6f0798e23 thys/Abacus.thy diff -r e9ef4ada308b -r d8e6f0798e23 thys/Abacus_Hoare.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys/Abacus_Hoare.thy Wed Mar 27 09:47:02 2013 +0000 @@ -0,0 +1,442 @@ +theory Abacus_Hoare +imports Abacus +begin + +type_synonym abc_assert = "nat list \ bool" + +definition + assert_imp :: "('a \ bool) \ ('a \ bool) \ bool" ("_ \ _" [0, 0] 100) +where + "assert_imp P Q \ \xs. P xs \ Q xs" + +fun abc_holds_for :: "(nat list \ bool) \ (nat \ nat list) \ bool" ("_ abc'_holds'_for _" [100, 99] 100) +where + "P abc_holds_for (s, lm) = P lm" + +(* Hoare Rules *) +(* halting case *) +(*consts abc_Hoare_halt :: "abc_assert \ abc_prog \ abc_assert \ bool" ("({(1_)}/ (_)/ {(1_)})" 50)*) + +fun abc_final :: "(nat \ nat list) \ abc_prog \ bool" + where + "abc_final (s, lm) p = (s = length p)" + +fun abc_notfinal :: "abc_conf \ abc_prog \ bool" + where + "abc_notfinal (s, lm) p = (s < length p)" + +definition + abc_Hoare_halt :: "abc_assert \ abc_prog \ abc_assert \ bool" ("({(1_)}/ (_)/ {(1_)})" 50) +where + "abc_Hoare_halt P p Q \ \lm. P lm \ (\n. abc_final (abc_steps_l (0, lm) p n) p \ Q abc_holds_for (abc_steps_l (0, lm) p n))" + +lemma abc_Hoare_haltI: + assumes "\lm. P lm \ \n. abc_final (abc_steps_l (0, lm) p n) p \ Q abc_holds_for (abc_steps_l (0, lm) p n)" + shows "{P} (p::abc_prog) {Q}" +unfolding abc_Hoare_halt_def +using assms by auto + +text {* + {P} A {Q} {Q} B {S} + ----------------------------------------- + {P} A [+] B {S} +*} + +definition + abc_Hoare_unhalt :: "abc_assert \ abc_prog \ bool" ("({(1_)}/ (_)) \" 50) +where + "abc_Hoare_unhalt P p \ \args. P args \ (\ n .abc_notfinal (abc_steps_l (0, args) p n) p)" + +lemma abc_Hoare_unhaltI: + assumes "\args n. P args \ abc_notfinal (abc_steps_l (0, args) p n) p" + shows "{P} (p::abc_prog) \" +unfolding abc_Hoare_unhalt_def +using assms by auto + +fun abc_inst_shift :: "abc_inst \ nat \ abc_inst" + where + "abc_inst_shift (Inc m) n = Inc m" | + "abc_inst_shift (Dec m e) n = Dec m (e + n)" | + "abc_inst_shift (Goto m) n = Goto (m + n)" + +fun abc_shift :: "abc_inst list \ nat \ abc_inst list" + where + "abc_shift xs n = map (\ x. abc_inst_shift x n) xs" + +fun abc_comp :: "abc_inst list \ abc_inst list \ + abc_inst list" (infixl "[+]" 99) + where + "abc_comp al bl = (let al_len = length al in + al @ abc_shift bl al_len)" + +lemma abc_comp_first_step_eq_pre: + "s < length A + \ abc_step_l (s, lm) (abc_fetch s (A [+] B)) = + abc_step_l (s, lm) (abc_fetch s A)" +by(simp add: abc_step_l.simps abc_fetch.simps nth_append) + +lemma abc_before_final: + "\abc_final (abc_steps_l (0, lm) p n) p; p \ []\ + \ \ n'. abc_notfinal (abc_steps_l (0, lm) p n') p \ + abc_final (abc_steps_l (0, lm) p (Suc n')) p" +proof(induct n) + case 0 + thus "?thesis" + by(simp add: abc_steps_l.simps) +next + case (Suc n) + have ind: " \abc_final (abc_steps_l (0, lm) p n) p; p \ []\ \ + \n'. abc_notfinal (abc_steps_l (0, lm) p n') p \ abc_final (abc_steps_l (0, lm) p (Suc n')) p" + by fact + have final: "abc_final (abc_steps_l (0, lm) p (Suc n)) p" by fact + have notnull: "p \ []" by fact + show "?thesis" + proof(cases "abc_final (abc_steps_l (0, lm) p n) p") + case True + have "abc_final (abc_steps_l (0, lm) p n) p" by fact + then have "\n'. abc_notfinal (abc_steps_l (0, lm) p n') p \ abc_final (abc_steps_l (0, lm) p (Suc n')) p" + using ind notnull + by simp + thus "?thesis" + by simp + next + case False + have "\ abc_final (abc_steps_l (0, lm) p n) p" by fact + from final this have "abc_notfinal (abc_steps_l (0, lm) p n) p" + by(case_tac "abc_steps_l (0, lm) p n", simp add: abc_step_red2 + abc_step_l.simps abc_fetch.simps split: if_splits) + thus "?thesis" + using final + by(rule_tac x = n in exI, simp) + qed +qed + +lemma notfinal_Suc: + "abc_notfinal (abc_steps_l (0, lm) A (Suc n)) A \ + abc_notfinal (abc_steps_l (0, lm) A n) A" +apply(case_tac "abc_steps_l (0, lm) A n") +apply(simp add: abc_step_red2 abc_fetch.simps abc_step_l.simps split: if_splits) +done + +lemma abc_comp_frist_steps_eq_pre: + assumes notfinal: "abc_notfinal (abc_steps_l (0, lm) A n) A" + and notnull: "A \ []" + shows "abc_steps_l (0, lm) (A [+] B) n = abc_steps_l (0, lm) A n" +using notfinal +proof(induct n) + case 0 + thus "?case" + by(simp add: abc_steps_l.simps) +next + case (Suc n) + have ind: "abc_notfinal (abc_steps_l (0, lm) A n) A \ abc_steps_l (0, lm) (A [+] B) n = abc_steps_l (0, lm) A n" + by fact + have h: "abc_notfinal (abc_steps_l (0, lm) A (Suc n)) A" by fact + then have a: "abc_notfinal (abc_steps_l (0, lm) A n) A" + by(simp add: notfinal_Suc) + then have b: "abc_steps_l (0, lm) (A [+] B) n = abc_steps_l (0, lm) A n" + using ind by simp + obtain s lm' where c: "abc_steps_l (0, lm) A n = (s, lm')" + by (metis prod.exhaust) + then have d: "s < length A \ abc_steps_l (0, lm) (A [+] B) n = (s, lm')" + using a b by simp + thus "?case" + using c + by(simp add: abc_step_red2 abc_fetch.simps abc_step_l.simps nth_append) +qed + +declare abc_shift.simps[simp del] abc_comp.simps[simp del] +lemma halt_steps2: "st \ length A \ abc_steps_l (st, lm) A stp = (st, lm)" +apply(induct stp) +by(simp_all add: abc_step_red2 abc_steps_l.simps abc_step_l.simps abc_fetch.simps) + +lemma halt_steps: "abc_steps_l (length A, lm) A n = (length A, lm)" +apply(induct n, simp add: abc_steps_l.simps) +apply(simp add: abc_step_red2 abc_step_l.simps nth_append abc_fetch.simps) +done + +lemma abc_steps_add: + "abc_steps_l (as, lm) ap (m + n) = + abc_steps_l (abc_steps_l (as, lm) ap m) ap n" +apply(induct m arbitrary: n as lm, simp add: abc_steps_l.simps) +proof - + fix m n as lm + assume ind: + "\n as lm. abc_steps_l (as, lm) ap (m + n) = + abc_steps_l (abc_steps_l (as, lm) ap m) ap n" + show "abc_steps_l (as, lm) ap (Suc m + n) = + abc_steps_l (abc_steps_l (as, lm) ap (Suc m)) ap n" + apply(insert ind[of as lm "Suc n"], simp) + apply(insert ind[of as lm "Suc 0"], simp add: abc_steps_l.simps) + apply(case_tac "(abc_steps_l (as, lm) ap m)", simp) + apply(simp add: abc_steps_l.simps) + apply(case_tac "abc_step_l (a, b) (abc_fetch a ap)", + simp add: abc_steps_l.simps) + done +qed + +lemma equal_when_halt: + assumes exc1: "abc_steps_l (s, lm) A na = (length A, lma)" + and exc2: "abc_steps_l (s, lm) A nb = (length A, lmb)" + shows "lma = lmb" +proof(cases "na > nb") + case True + then obtain d where "na = nb + d" + by (metis add_Suc_right less_iff_Suc_add) + thus "?thesis" using assms halt_steps + by(simp add: abc_steps_add) +next + case False + then obtain d where "nb = na + d" + by (metis add.comm_neutral less_imp_add_positive nat_neq_iff) + thus "?thesis" using assms halt_steps + by(simp add: abc_steps_add) +qed + +lemma abc_comp_frist_steps_halt_eq': + assumes final: "abc_steps_l (0, lm) A n = (length A, lm')" + and notnull: "A \ []" + shows "\ n'. abc_steps_l (0, lm) (A [+] B) n' = (length A, lm')" +proof - + have "\ n'. abc_notfinal (abc_steps_l (0, lm) A n') A \ + abc_final (abc_steps_l (0, lm) A (Suc n')) A" + using assms + by(rule_tac n = n in abc_before_final, simp_all) + then obtain na where a: + "abc_notfinal (abc_steps_l (0, lm) A na) A \ + abc_final (abc_steps_l (0, lm) A (Suc na)) A" .. + obtain sa lma where b: "abc_steps_l (0, lm) A na = (sa, lma)" + by (metis prod.exhaust) + then have c: "abc_steps_l (0, lm) (A [+] B) na = (sa, lma)" + using a abc_comp_frist_steps_eq_pre[of lm A na B] assms + by simp + have d: "sa < length A" using b a by simp + then have e: "abc_step_l (sa, lma) (abc_fetch sa (A [+] B)) = + abc_step_l (sa, lma) (abc_fetch sa A)" + by(rule_tac abc_comp_first_step_eq_pre) + from a have "abc_steps_l (0, lm) A (Suc na) = (length A, lm')" + using final equal_when_halt + by(case_tac "abc_steps_l (0, lm) A (Suc na)" , simp) + then have "abc_steps_l (0, lm) (A [+] B) (Suc na) = (length A, lm')" + using a b c e + by(simp add: abc_step_red2) + thus "?thesis" + by blast +qed + +lemma abc_exec_null: "abc_steps_l sam [] n = sam" +apply(cases sam) +apply(induct n) +apply(auto simp: abc_step_red2) +apply(auto simp: abc_step_l.simps abc_steps_l.simps abc_fetch.simps) +done + +lemma abc_comp_frist_steps_halt_eq: + assumes final: "abc_steps_l (0, lm) A n = (length A, lm')" + shows "\ n'. abc_steps_l (0, lm) (A [+] B) n' = (length A, lm')" +using final +apply(case_tac "A = []") +apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps abc_exec_null) +apply(rule_tac abc_comp_frist_steps_halt_eq', simp_all) +done + +lemma abc_comp_second_step_eq: + assumes exec: "abc_step_l (s, lm) (abc_fetch s B) = (sa, lma)" + shows "abc_step_l (s + length A, lm) (abc_fetch (s + length A) (A [+] B)) + = (sa + length A, lma)" +using assms +apply(auto simp: abc_step_l.simps abc_fetch.simps nth_append abc_comp.simps abc_shift.simps split : if_splits ) +apply(case_tac [!] "B ! s", auto simp: Let_def) +done + +lemma abc_comp_second_steps_eq: + assumes exec: "abc_steps_l (0, lm) B n = (sa, lm')" + shows "abc_steps_l (length A, lm) (A [+] B) n = (sa + length A, lm')" +using assms +proof(induct n arbitrary: sa lm') + case 0 + thus "?case" + by(simp add: abc_steps_l.simps) +next + case (Suc n) + have ind: "\sa lm'. abc_steps_l (0, lm) B n = (sa, lm') \ + abc_steps_l (length A, lm) (A [+] B) n = (sa + length A, lm')" by fact + have exec: "abc_steps_l (0, lm) B (Suc n) = (sa, lm')" by fact + obtain sb lmb where a: " abc_steps_l (0, lm) B n = (sb, lmb)" + by (metis prod.exhaust) + then have "abc_steps_l (length A, lm) (A [+] B) n = (sb + length A, lmb)" + using ind by simp + moreover have "abc_step_l (sb + length A, lmb) (abc_fetch (sb + length A) (A [+] B)) = (sa + length A, lm') " + using a exec abc_comp_second_step_eq + by(simp add: abc_step_red2) + ultimately show "?case" + by(simp add: abc_step_red2) +qed + +lemma length_abc_comp[simp, intro]: + "length (A [+] B) = length A + length B" +by(auto simp: abc_comp.simps abc_shift.simps) + +lemma abc_Hoare_plus_halt : + assumes A_halt : "{P} (A::abc_prog) {Q}" + and B_halt : "{Q} (B::abc_prog) {S}" + shows "{P} (A [+] B) {S}" +proof(rule_tac abc_Hoare_haltI) + fix lm + assume a: "P lm" + then obtain na lma where + "abc_final (abc_steps_l (0, lm) A na) A" + and b: "abc_steps_l (0, lm) A na = (length A, lma)" + and c: "Q abc_holds_for (length A, lma)" + using A_halt unfolding abc_Hoare_halt_def + by (metis (full_types) abc_final.simps abc_holds_for.simps prod.exhaust) + have "\ n. abc_steps_l (0, lm) (A [+] B) n = (length A, lma)" + using abc_comp_frist_steps_halt_eq b + by(simp) + then obtain nx where h1: "abc_steps_l (0, lm) (A [+] B) nx = (length A, lma)" .. + from c have "Q lma" + using c unfolding abc_holds_for.simps + by simp + then obtain nb lmb where + "abc_final (abc_steps_l (0, lma) B nb) B" + and d: "abc_steps_l (0, lma) B nb = (length B, lmb)" + and e: "S abc_holds_for (length B, lmb)" + using B_halt unfolding abc_Hoare_halt_def + by (metis (full_types) abc_final.simps abc_holds_for.simps prod.exhaust) + have h2: "abc_steps_l (length A, lma) (A [+] B) nb = (length B + length A, lmb)" + using d abc_comp_second_steps_eq + by simp + thus "\n. abc_final (abc_steps_l (0, lm) (A [+] B) n) (A [+] B) \ + S abc_holds_for abc_steps_l (0, lm) (A [+] B) n" + using h1 e + by(rule_tac x = "nx + nb" in exI, simp add: abc_steps_add) +qed + +lemma abc_unhalt_append_eq: + assumes unhalt: "{P} (A::abc_prog) \" + and P: "P args" + shows "abc_steps_l (0, args) (A [+] B) stp = abc_steps_l (0, args) A stp" +proof(induct stp) + case 0 + thus "?case" + by(simp add: abc_steps_l.simps) +next + case (Suc stp) + have ind: "abc_steps_l (0, args) (A [+] B) stp = abc_steps_l (0, args) A stp" + by fact + obtain s nl where a: "abc_steps_l (0, args) A stp = (s, nl)" + by (metis prod.exhaust) + then have b: "s < length A" + using unhalt P + apply(auto simp: abc_Hoare_unhalt_def) + by (metis abc_notfinal.simps) + thus "?case" + using a ind + by(simp add: abc_step_red2 abc_step_l.simps abc_fetch.simps nth_append abc_comp.simps) +qed + +lemma abc_Hoare_plus_unhalt1: + "{P} (A::abc_prog) \ \ {P} (A [+] B) \" +apply(rule_tac abc_Hoare_unhaltI) +apply(frule_tac args = args and B = B and stp = n in abc_unhalt_append_eq) +apply(simp_all add: abc_Hoare_unhalt_def) +apply(erule_tac x = args in allE, simp) +apply(erule_tac x = n in allE) +apply(case_tac "(abc_steps_l (0, args) A n)", simp) +done + + +lemma notfinal_all_before: + "\abc_notfinal (abc_steps_l (0, args) A x) A; y\x \ + \ abc_notfinal (abc_steps_l (0, args) A y) A " +apply(subgoal_tac "\ d. x = y + d", auto) +apply(case_tac "abc_steps_l (0, args) A y",simp) +apply(rule_tac classical, simp add: abc_steps_add leI halt_steps2) +by arith + +lemma abc_Hoare_plus_unhalt2': + assumes unhalt: "{Q} (B::abc_prog) \" + and halt: "{P} (A::abc_prog) {Q}" + and notnull: "A \ []" + and P: "P args" + shows "abc_notfinal (abc_steps_l (0, args) (A [+] B) n) (A [+] B)" +proof - + obtain st nl stp where a: "abc_final (abc_steps_l (0, args) A stp) A" + and b: "Q abc_holds_for (length A, nl)" + and c: "abc_steps_l (0, args) A stp = (st, nl)" + using halt P unfolding abc_Hoare_halt_def + by (metis abc_holds_for.simps prod.exhaust) + thm abc_before_final + obtain stpa where d: + "abc_notfinal (abc_steps_l (0, args) A stpa) A \ abc_final (abc_steps_l (0, args) A (Suc stpa)) A" + using a notnull abc_before_final[of args A stp] + by(auto) + thus "?thesis" + proof(cases "n < Suc stpa") + case True + have h: "n < Suc stpa" by fact + then have "abc_notfinal (abc_steps_l (0, args) A n) A" + using d + by(rule_tac notfinal_all_before, auto) + moreover then have "abc_steps_l (0, args) (A [+] B) n = abc_steps_l (0, args) A n" + using notnull + by(rule_tac abc_comp_frist_steps_eq_pre, simp_all) + ultimately show "?thesis" + by(case_tac "abc_steps_l (0, args) A n", simp) + next + case False + have "\ n < Suc stpa" by fact + then obtain d where i1: "n = Suc stpa + d" + by (metis add_Suc less_iff_Suc_add not_less_eq) + have "abc_steps_l (0, args) A (Suc stpa) = (length A, nl)" + using d a c + apply(case_tac "abc_steps_l (0, args) A stp", simp add: equal_when_halt) + by(case_tac "abc_steps_l (0, args) A (Suc stpa)", simp add: equal_when_halt) + moreover have "abc_steps_l (0, args) (A [+] B) stpa = abc_steps_l (0, args) A stpa" + using notnull d + by(rule_tac abc_comp_frist_steps_eq_pre, simp_all) + ultimately have i2: "abc_steps_l (0, args) (A [+] B) (Suc stpa) = (length A, nl)" + using d + apply(case_tac "abc_steps_l (0, args) A stpa", simp) + by(simp add: abc_step_red2 abc_steps_l.simps abc_fetch.simps abc_comp.simps nth_append) + obtain s' nl' where i3:"abc_steps_l (0, nl) B d = (s', nl')" + by (metis prod.exhaust) + then have i4: "abc_steps_l (0, args) (A [+] B) (Suc stpa + d) = (length A + s', nl')" + using i2 apply(simp only: abc_steps_add) + using abc_comp_second_steps_eq[of nl B d s' nl'] + by simp + moreover have "s' < length B" + using unhalt b i3 + apply(simp add: abc_Hoare_unhalt_def) + apply(erule_tac x = nl in allE, simp) + by(erule_tac x = d in allE, simp) + ultimately show "?thesis" + using i1 + by(simp) + qed +qed + +lemma abc_comp_null_left[simp]: "[] [+] A = A" +apply(induct A) +apply(case_tac [2] a) +apply(auto simp: abc_comp.simps abc_shift.simps abc_inst_shift.simps) +done + +lemma abc_comp_null_right[simp]: "A [+] [] = A" +apply(induct A) +apply(case_tac [2] a) +apply(auto simp: abc_comp.simps abc_shift.simps abc_inst_shift.simps) +done + +lemma abc_Hoare_plus_unhalt2: + "\{Q} (B::abc_prog)\; {P} (A::abc_prog) {Q}\\ {P} (A [+] B) \" +apply(case_tac "A = []") +apply(simp add: abc_Hoare_halt_def abc_Hoare_unhalt_def abc_exec_null) +apply(rule_tac abc_Hoare_unhaltI) +apply(erule_tac abc_Hoare_plus_unhalt2', simp) +apply(simp, simp) +done + +end + + diff -r e9ef4ada308b -r d8e6f0798e23 thys/Abacus_Mopup.thy diff -r e9ef4ada308b -r d8e6f0798e23 thys/Rec_Def.thy --- a/thys/Rec_Def.thy Thu Mar 14 20:43:43 2013 +0000 +++ b/thys/Rec_Def.thy Wed Mar 27 09:47:02 2013 +0000 @@ -1,95 +1,50 @@ -(* Title: thys/Rec_Def.thy - Author: Jian Xu, Xingyuan Zhang, and Christian Urban -*) - -header {* Definition of Recursive Functions *} - - theory Rec_Def imports Main begin -section {* Recursive functions *} +datatype recf = z + | s + | id nat nat + | Cn nat recf "recf list" + | Pr nat recf recf + | Mn nat recf -datatype recf = - z | s | - -- {* The projection function, where @{text "id i j"} returns the @{text "j"}-th - argment out of the @{text "i"} arguments. *} - id nat nat | - -- {* The compostion operator, where "@{text "Cn n f [g1; g2; \ ;gm]"} - computes @{text "f (g1(x1, x2, \, xn), g2(x1, x2, \, xn), \ , - gm(x1, x2, \ , xn))"} for input argments @{text "x1, \, xn"}. *} - Cn nat recf "recf list" | - -- {* The primitive resursive operator, where @{text "Pr n f g"} computes: - @{text "Pr n f g (x1, x2, \, xn-1, 0) = f(x1, \, xn-1)"} - and @{text "Pr n f g (x1, x2, \, xn-1, k') = g(x1, x2, \, xn-1, k, - Pr n f g (x1, \, xn-1, k))"}. - *} - Pr nat recf recf | - -- {* The minimization operator, where @{text "Mn n f (x1, x2, \ , xn)"} - computes the first i such that @{text "f (x1, \, xn, i) = 0"} and for all - @{text "j"}, @{text "f (x1, x2, \, xn, j) > 0"}. *} - Mn nat recf - -(* -partial_function (tailrec) - rec_exec :: "recf \ nat list \ nat" -where - "rec_exec f ns = (case (f, ns) of - (z, xs) => 0 - | (s, xs) => Suc (xs ! 0) - | (id m n, xs) => (xs ! n) - | (Cn n f gs, xs) => - (let ys = (map (\ a. rec_exec a xs) gs) in - rec_exec f ys) - | (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])])) - | (Mn n f, xs) => (LEAST x. rec_exec f (xs @ [x]) = 0))" -*) +definition pred_of_nl :: "nat list \ nat list" + where + "pred_of_nl xs = butlast xs @ [last xs - 1]" -text {* - The semantis of recursive operators is given by an inductively defined - relation as follows, where - @{text "rec_calc_rel R [x1, x2, \, xn] r"} means the computation of - @{text "R"} over input arguments @{text "[x1, x2, \, xn"} terminates - and gives rise to a result @{text "r"} -*} +function rec_exec :: "recf \ nat list \ nat" + where + "rec_exec z xs = 0" | + "rec_exec s xs = (Suc (xs ! 0))" | + "rec_exec (id m n) xs = (xs ! n)" | + "rec_exec (Cn n f gs) xs = + rec_exec f (map (\ a. rec_exec a xs) gs)" | + "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 +apply(relation "measures [\ (r, xs). size r, (\ (r, xs). last xs)]") +apply(auto simp add: less_Suc_eq_le intro: trans_le_add2 list_size_estimation') +done -inductive rec_calc_rel :: "recf \ nat list \ nat \ bool" -where - calc_z: "rec_calc_rel z [n] 0" | - calc_s: "rec_calc_rel s [n] (Suc n)" | - calc_id: "\length args = i; j < i; args!j = r\ \ rec_calc_rel (id i j) args r" | - calc_cn: "\length args = n; - \ k < length gs. rec_calc_rel (gs ! k) args (rs ! k); - length rs = length gs; - rec_calc_rel f rs r\ - \ rec_calc_rel (Cn n f gs) args r" | - calc_pr_zero: - "\length args = n; - rec_calc_rel f args r0 \ - \ rec_calc_rel (Pr n f g) (args @ [0]) r0" | - calc_pr_ind: " - \ length args = n; - rec_calc_rel (Pr n f g) (args @ [k]) rk; - rec_calc_rel g (args @ [k] @ [rk]) rk'\ - \ rec_calc_rel (Pr n f g) (args @ [Suc k]) rk'" | - calc_mn: "\length args = n; - rec_calc_rel f (args@[r]) 0; - \ i < r. (\ ri. rec_calc_rel f (args@[i]) ri \ ri \ 0)\ - \ rec_calc_rel (Mn n f) args r" +inductive terminate :: "recf \ nat list \ bool" + where + termi_z: "terminate z [n]" +| termi_s: "terminate s [n]" +| termi_id: "\n < m; length xs = m\ \ terminate (id m n) xs" +| termi_cn: "\terminate f (map (\g. rec_exec g xs) gs); + \g \ set gs. terminate g xs; length xs = n\ \ terminate (Cn n f gs) xs" +| termi_pr: "\\ y < x. terminate g (xs @ y # [rec_exec (Pr n f g) (xs @ [y])]); + terminate f xs; + length xs = n\ + \ terminate (Pr n f g) (xs @ [x])" +| termi_mn: "\length xs = n; terminate f (xs @ [r]); + rec_exec f (xs @ [r]) = 0; + \ i < r. terminate f (xs @ [i]) \ rec_exec f (xs @ [i]) > 0\ \ terminate (Mn n f) xs" -inductive_cases calc_pr_reverse: "rec_calc_rel (Pr n f g) (lm) rSucy" -inductive_cases calc_z_reverse: "rec_calc_rel z lm x" - -inductive_cases calc_s_reverse: "rec_calc_rel s lm x" - -inductive_cases calc_id_reverse: "rec_calc_rel (id m n) lm x" - -inductive_cases calc_cn_reverse: "rec_calc_rel (Cn n f gs) lm x" - -inductive_cases calc_mn_reverse:"rec_calc_rel (Mn n f) lm x" end \ No newline at end of file diff -r e9ef4ada308b -r d8e6f0798e23 thys/Recursive.thy --- a/thys/Recursive.thy Thu Mar 14 20:43:43 2013 +0000 +++ b/thys/Recursive.thy Wed Mar 27 09:47:02 2013 +0000 @@ -1,30 +1,7 @@ -(* Title: thys/Recursive.thy - Author: Jian Xu, Xingyuan Zhang, and Christian Urban -*) - -header {* Translating Recursive Functions into Abacus Machines *} - theory Recursive -imports Rec_Def Abacus +imports Abacus Rec_Def Abacus_Hoare begin - -text {* - Some auxilliary Abacus machines used to construct the result Abacus machines. -*} - -text {* - @{text "get_paras_num recf"} returns the arity of recursive function @{text "recf"}. -*} -fun get_paras_num :: "recf \ nat" - where - "get_paras_num z = 1" | - "get_paras_num s = 1" | - "get_paras_num (id m n) = m" | - "get_paras_num (Cn n f gs) = n" | - "get_paras_num (Pr n f g) = Suc n" | - "get_paras_num (Mn n f) = n" - fun addition :: "nat \ nat \ nat \ abc_prog" where "addition m n p = [Dec m 4, Inc n, Inc p, Goto 0, Dec p 7, Inc m, Goto 4]" @@ -33,22 +10,6 @@ where "mv_box m n = [Dec m 3, Inc n, Goto 0]" -fun abc_inst_shift :: "abc_inst \ nat \ abc_inst" - where - "abc_inst_shift (Inc m) n = Inc m" | - "abc_inst_shift (Dec m e) n = Dec m (e + n)" | - "abc_inst_shift (Goto m) n = Goto (m + n)" - -fun abc_shift :: "abc_inst list \ nat \ abc_inst list" - where - "abc_shift xs n = map (\ x. abc_inst_shift x n) xs" - -fun abc_append :: "abc_inst list \ abc_inst list \ - abc_inst list" (infixl "[+]" 60) - where - "abc_append al bl = (let al_len = length al in - al @ abc_shift bl al_len)" - text {* The compilation of @{text "z"}-operator. *} definition rec_ci_z :: "abc_inst list" where @@ -124,2245 +85,23 @@ 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_boxes.simps[simp del] mv_box.simps[simp del] addition.simps[simp del] declare abc_steps_l.simps[simp del] abc_fetch.simps[simp del] abc_step_l.simps[simp del] -lemma abc_steps_add: - "abc_steps_l (as, lm) ap (m + n) = - abc_steps_l (abc_steps_l (as, lm) ap m) ap n" -apply(induct m arbitrary: n as lm, simp add: abc_steps_l.simps) -proof - - fix m n as lm - assume ind: - "\n as lm. abc_steps_l (as, lm) ap (m + n) = - abc_steps_l (abc_steps_l (as, lm) ap m) ap n" - show "abc_steps_l (as, lm) ap (Suc m + n) = - abc_steps_l (abc_steps_l (as, lm) ap (Suc m)) ap n" - apply(insert ind[of as lm "Suc n"], simp) - apply(insert ind[of as lm "Suc 0"], simp add: abc_steps_l.simps) - apply(case_tac "(abc_steps_l (as, lm) ap m)", simp) - apply(simp add: abc_steps_l.simps) - apply(case_tac "abc_step_l (a, b) (abc_fetch a ap)", - simp add: abc_steps_l.simps) - done -qed - -(*lemmas: rec_ci and rec_calc_rel*) - -lemma rec_calc_inj_case_z: - "\rec_calc_rel z l x; rec_calc_rel z l y\ \ x = y" -apply(auto elim: calc_z_reverse) -done - -lemma rec_calc_inj_case_s: - "\rec_calc_rel s l x; rec_calc_rel s l y\ \ x = y" -apply(auto elim: calc_s_reverse) -done - -lemma rec_calc_inj_case_id: - "\rec_calc_rel (recf.id nat1 nat2) l x; - rec_calc_rel (recf.id nat1 nat2) l y\ \ x = y" -apply(auto elim: calc_id_reverse) -done - -lemma rec_calc_inj_case_mn: - assumes ind: "\ l x y. \rec_calc_rel f l x; rec_calc_rel f l y\ - \ x = y" - and h: "rec_calc_rel (Mn n f) l x" "rec_calc_rel (Mn n f) l y" - shows "x = y" - apply(insert h) - apply(elim calc_mn_reverse) - apply(case_tac "x > y", simp) - apply(erule_tac x = "y" in allE, auto) -proof - - fix v va - assume "rec_calc_rel f (l @ [y]) 0" - "rec_calc_rel f (l @ [y]) v" - "0 < v" - thus "False" - apply(insert ind[of "l @ [y]" 0 v], simp) - done -next - fix v va - assume - "rec_calc_rel f (l @ [x]) 0" - "\xv. rec_calc_rel f (l @ [x]) v \ 0 < v" "\ y < x" - thus "x = y" - apply(erule_tac x = "x" in allE) - apply(case_tac "x = y", auto) - apply(drule_tac y = v in ind, simp, simp) - done -qed - -lemma rec_calc_inj_case_pr: - assumes f_ind: - "\l x y. \rec_calc_rel f l x; rec_calc_rel f l y\ \ x = y" - and g_ind: - "\x xa y xb ya l xc yb. - \x = rec_ci f; (xa, y) = x; (xb, ya) = y; - rec_calc_rel g l xc; rec_calc_rel g l yb\ \ xc = yb" - and h: "rec_calc_rel (Pr n f g) l x" "rec_calc_rel (Pr n f g) l y" - shows "x = y" - apply(case_tac "rec_ci f") -proof - - fix a b c - assume "rec_ci f = (a, b, c)" - hence ng_ind: - "\ l xc yb. \rec_calc_rel g l xc; rec_calc_rel g l yb\ - \ xc = yb" - apply(insert g_ind[of "(a, b, c)" "a" "(b, c)" b c], simp) - done - from h show "x = y" - apply(erule_tac calc_pr_reverse, erule_tac calc_pr_reverse) - apply(erule f_ind, simp, simp) - apply(erule_tac calc_pr_reverse, simp, simp) - proof - - fix la ya ry laa yaa rya - assume k1: "rec_calc_rel g (la @ [ya, ry]) x" - "rec_calc_rel g (la @ [ya, rya]) y" - and k2: "rec_calc_rel (Pr (length la) f g) (la @ [ya]) ry" - "rec_calc_rel (Pr (length la) f g) (la @ [ya]) rya" - from k2 have "ry = rya" - apply(induct ya arbitrary: ry rya) - apply(erule_tac calc_pr_reverse, - erule_tac calc_pr_reverse, simp) - apply(erule f_ind, simp, simp, simp) - apply(erule_tac calc_pr_reverse, simp) - apply(erule_tac rSucy = rya in calc_pr_reverse, simp, simp) - proof - - fix ya ry rya l y ryb laa yb ryc - assume ind: - "\ry rya. \rec_calc_rel (Pr (length l) f g) (l @ [y]) ry; - rec_calc_rel (Pr (length l) f g) (l @ [y]) rya\ \ ry = rya" - and j: "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryb" - "rec_calc_rel g (l @ [y, ryb]) ry" - "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryc" - "rec_calc_rel g (l @ [y, ryc]) rya" - from j show "ry = rya" - apply(insert ind[of ryb ryc], simp) - apply(insert ng_ind[of "l @ [y, ryc]" ry rya], simp) - done - qed - from k1 and this show "x = y" - apply(simp) - apply(insert ng_ind[of "la @ [ya, rya]" x y], simp) - done - qed -qed - -lemma Suc_nth_part_eq: - "\k \k<(length list). (xs) ! k = (list) ! k" -apply(rule allI, rule impI) -apply(erule_tac x = "Suc k" in allE, simp) -done - - -lemma list_eq_intro: - "\length xs = length ys; \ k < length xs. xs ! k = ys ! k\ - \ xs = ys" -apply(induct xs arbitrary: ys, simp) -apply(case_tac ys, simp, simp) -proof - - fix a xs ys aa list - assume ind: - "\ys. \length list = length ys; \k - \ xs = ys" - and h: "length xs = length list" - "\k xs = list" - apply(insert ind[of list], simp) - apply(frule Suc_nth_part_eq, simp) - apply(erule_tac x = "0" in allE, simp) - done -qed - -lemma rec_calc_inj_case_cn: - assumes ind: - "\x l xa y. - \x = f \ x \ set gs; rec_calc_rel x l xa; rec_calc_rel x l y\ - \ xa = y" - and h: "rec_calc_rel (Cn n f gs) l x" - "rec_calc_rel (Cn n f gs) l y" - shows "x = y" - apply(insert h, elim calc_cn_reverse) - apply(subgoal_tac "rs = rsa") - apply(rule_tac x = f and l = rsa and xa = x and y = y in ind, - simp, simp, simp) - apply(intro list_eq_intro, simp, rule allI, rule impI) - apply(erule_tac x = k in allE, rule_tac x = k in allE, simp, simp) - apply(rule_tac x = "gs ! k" in ind, simp, simp, simp) - done - -lemma rec_calc_inj: - "\rec_calc_rel f l x; - rec_calc_rel f l y\ \ x = y" -apply(induct f arbitrary: l x y rule: rec_ci.induct) -apply(simp add: rec_calc_inj_case_z) -apply(simp add: rec_calc_inj_case_s) -apply(simp add: rec_calc_inj_case_id) -apply (metis rec_calc_inj_case_cn) -apply(erule rec_calc_inj_case_pr, auto) -apply(erule rec_calc_inj_case_mn, auto) -done - - -lemma calc_rel_reverse_ind_step_ex: - "\rec_calc_rel (Pr n f g) (lm @ [Suc x]) rs\ - \ \ rs. rec_calc_rel (Pr n f g) (lm @ [x]) rs" -apply(erule calc_pr_reverse, simp, simp) -apply(rule_tac x = rk in exI, simp) -done - -lemma [simp]: "Suc x \ y \ Suc (y - Suc x) = y - x" -by arith - -lemma calc_pr_para_not_null: - "rec_calc_rel (Pr n f g) lm rs \ lm \ []" -apply(erule calc_pr_reverse, simp, simp) -done - -lemma calc_pr_less_ex: - "\rec_calc_rel (Pr n f g) lm rs; x \ last lm\ \ - \rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rs" -apply(subgoal_tac "lm \ []") -apply(induct x, rule_tac x = rs in exI, simp, simp, erule exE) -apply(rule_tac rs = xa in calc_rel_reverse_ind_step_ex, simp) -apply(simp add: calc_pr_para_not_null) -done - -lemma calc_pr_zero_ex: - "rec_calc_rel (Pr n f g) lm rs \ - \rs. rec_calc_rel f (butlast lm) rs" -apply(drule_tac x = "last lm" in calc_pr_less_ex, simp, - erule_tac exE, simp) -apply(erule_tac calc_pr_reverse, simp) -apply(rule_tac x = rs in exI, simp, simp) -done - - -lemma abc_steps_ind: - "abc_steps_l (as, am) ap (Suc stp) = - abc_steps_l (abc_steps_l (as, am) ap stp) ap (Suc 0)" -apply(insert abc_steps_add[of as am ap stp "Suc 0"], simp) -done - -lemma abc_steps_zero: "abc_steps_l asm ap 0 = asm" -apply(case_tac asm, simp add: abc_steps_l.simps) -done - -lemma abc_append_nth: - "n < length ap + length bp \ - (ap [+] bp) ! n = - (if n < length ap then ap ! n - else abc_inst_shift (bp ! (n - length ap)) (length ap))" -apply(simp add: abc_append.simps nth_append map_nth split: if_splits) -done - -lemma abc_state_keep: - "as \ length bp \ abc_steps_l (as, lm) bp stp = (as, lm)" -apply(induct stp, simp add: abc_steps_zero) -apply(simp add: abc_steps_ind) -apply(simp add: abc_steps_zero) -apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps) -done - -lemma abc_halt_equal: - "\abc_steps_l (0, lm) bp stpa = (length bp, lm1); - abc_steps_l (0, lm) bp stpb = (length bp, lm2)\ \ lm1 = lm2" -apply(case_tac "stpa - stpb > 0") -apply(insert abc_steps_add[of 0 lm bp stpb "stpa - stpb"], simp) -apply(insert abc_state_keep[of bp "length bp" lm2 "stpa - stpb"], - simp, simp add: abc_steps_zero) -apply(insert abc_steps_add[of 0 lm bp stpa "stpb - stpa"], simp) -apply(insert abc_state_keep[of bp "length bp" lm1 "stpb - stpa"], - simp) -done - -lemma abc_halt_point_ex: - "\\stp. abc_steps_l (0, lm) bp stp = (bs, lm'); - bs = length bp; bp \ []\ - \ \ stp. (\ (s, l). s < bs \ - (abc_steps_l (s, l) bp (Suc 0)) = (bs, lm')) - (abc_steps_l (0, lm) bp stp) " -apply(erule_tac exE) -proof - - fix stp - assume "bs = length bp" - "abc_steps_l (0, lm) bp stp = (bs, lm')" - "bp \ []" - thus - "\stp. (\(s, l). s < bs \ - abc_steps_l (s, l) bp (Suc 0) = (bs, lm')) - (abc_steps_l (0, lm) bp stp)" - apply(induct stp, simp add: abc_steps_zero, simp) - proof - - fix stpa - assume ind: - "abc_steps_l (0, lm) bp stpa = (length bp, lm') - \ \stp. (\(s, l). s < length bp \ abc_steps_l (s, l) bp - (Suc 0) = (length bp, lm')) (abc_steps_l (0, lm) bp stp)" - and h: "abc_steps_l (0, lm) bp (Suc stpa) = (length bp, lm')" - "abc_steps_l (0, lm) bp stp = (length bp, lm')" - "bp \ []" - from h show - "\stp. (\(s, l). s < length bp \ abc_steps_l (s, l) bp (Suc 0) - = (length bp, lm')) (abc_steps_l (0, lm) bp stp)" - apply(case_tac "abc_steps_l (0, lm) bp stpa", - case_tac "a = length bp") - apply(insert ind, simp) - apply(subgoal_tac "b = lm'", simp) - apply(rule_tac abc_halt_equal, simp, simp) - apply(rule_tac x = stpa in exI, simp add: abc_steps_ind) - apply(simp add: abc_steps_zero) - apply(rule classical, simp add: abc_steps_l.simps - abc_fetch.simps abc_step_l.simps) - done - qed -qed - - -lemma abc_append_empty_r[simp]: "[] [+] ab = ab" -apply(simp add: abc_append.simps abc_inst_shift.simps) -apply(induct ab, simp, simp) -apply(case_tac a, simp_all add: abc_inst_shift.simps) -done - -lemma abc_append_empty_l[simp]: "ab [+] [] = ab" -apply(simp add: abc_append.simps abc_inst_shift.simps) -done - - -lemma abc_append_length[simp]: - "length (ap [+] bp) = length ap + length bp" -apply(simp add: abc_append.simps) -done - -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]: - "\a < length bp; abc_steps_l (a, b) bp (Suc 0) = (length bp, lm')\ - \ abc_steps_l (length ap + a, b) (ap [+] bp [+] cp) (Suc 0) = - (length ap + length bp, lm')" -apply(simp add: abc_steps_l.simps abc_fetch.simps abc_append_nth) -apply(case_tac "bp ! a", - auto simp: abc_steps_l.simps abc_step_l.simps) -done - -lemma abc_step_state_in: - "\bs < length bp; abc_steps_l (a, b) bp (Suc 0) = (bs, l)\ - \ a < length bp" -apply(simp add: abc_steps_l.simps abc_fetch.simps) -apply(rule_tac classical, - simp add: abc_step_l.simps abc_steps_l.simps) -done - - -lemma abc_append_state_in_exc: - "\bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\ - \ abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = - (length ap + bs, l)" -apply(induct stpa arbitrary: bs l, simp add: abc_steps_zero) -proof - - fix stpa bs l - assume ind: - "\bs l. \bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\ - \ abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = - (length ap + bs, l)" - and h: "bs < length bp" - "abc_steps_l (0, lm) bp (Suc stpa) = (bs, l)" - from h show - "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) (Suc stpa) = - (length ap + bs, l)" - apply(simp add: abc_steps_ind) - apply(case_tac "(abc_steps_l (0, lm) bp stpa)", simp) - proof - - fix a b - assume g: "abc_steps_l (0, lm) bp stpa = (a, b)" - "abc_steps_l (a, b) bp (Suc 0) = (bs, l)" - from h and g have k1: "a < length bp" - apply(simp add: abc_step_state_in) - done - from h and g and k1 show - "abc_steps_l (abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa) - (ap [+] bp [+] cp) (Suc 0) = (length ap + bs, l)" - apply(insert ind[of a b], simp) - apply(simp add: abc_steps_l.simps abc_fetch.simps - abc_append_nth) - apply(case_tac "bp ! a", auto simp: - abc_steps_l.simps abc_step_l.simps) - done - qed -qed - -lemma [simp]: "abc_steps_l (0, am) [] stp = (0, am)" -apply(induct stp, simp add: abc_steps_zero) -apply(simp add: abc_steps_ind) -apply(simp add: abc_steps_zero abc_steps_l.simps - abc_fetch.simps abc_step_l.simps) -done - -lemma abc_append_exc1: - "\\ stp. abc_steps_l (0, lm) bp stp = (bs, lm'); - bs = length bp; - as = length ap\ - \ \ stp. abc_steps_l (as, lm) (ap [+] bp [+] cp) stp - = (as + bs, lm')" -apply(case_tac "bp = []", erule_tac exE, simp, - rule_tac x = 0 in exI, simp add: abc_steps_zero) -apply(frule_tac abc_halt_point_ex, simp, simp, - erule_tac exE, erule_tac exE) -apply(rule_tac x = "stpa + Suc 0" in exI) -apply(case_tac "(abc_steps_l (0, lm) bp stpa)", - simp add: abc_steps_ind) -apply(subgoal_tac - "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa - = (length ap + a, b)", simp) -apply(simp add: abc_steps_zero) -apply(rule_tac abc_append_state_in_exc, simp, simp) -done - -lemma abc_append_exc3: - "\\ stp. abc_steps_l (0, am) bp stp = (bs, bm); ss = length ap\ - \ \ stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)" -apply(erule_tac exE) -proof - - fix stp - assume h: "abc_steps_l (0, am) bp stp = (bs, bm)" "ss = length ap" - thus " \stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)" - proof(induct stp arbitrary: bs bm) - fix bs bm - assume "abc_steps_l (0, am) bp 0 = (bs, bm)" - thus "\stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)" - apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps) - done - next - fix stp bs bm - assume ind: - "\bs bm. \abc_steps_l (0, am) bp stp = (bs, bm); - ss = length ap\ \ - \stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)" - and g: "abc_steps_l (0, am) bp (Suc stp) = (bs, bm)" - from g show - "\stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)" - apply(insert abc_steps_add[of 0 am bp stp "Suc 0"], simp) - apply(case_tac "(abc_steps_l (0, am) bp stp)", simp) - proof - - fix a b - assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)" - "abc_steps_l (0, am) bp (Suc stp) = - abc_steps_l (a, b) bp (Suc 0)" - "abc_steps_l (0, am) bp stp = (a, b)" - thus "?thesis" - apply(insert ind[of a b], simp add: h, erule_tac exE) - apply(rule_tac x = "Suc stp" in exI) - apply(simp only: abc_steps_ind, simp add: abc_steps_zero) - proof - - fix stp - assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)" - thus "abc_steps_l (a + length ap, b) (ap [+] bp) (Suc 0) - = (bs + length ap, bm)" - apply(simp add: abc_steps_l.simps abc_steps_zero - abc_fetch.simps split: if_splits) - apply(case_tac "bp ! a", - simp_all add: abc_inst_shift.simps abc_append_nth - abc_steps_l.simps abc_steps_zero abc_step_l.simps) - apply(auto) - done - qed - qed - qed -qed - -lemma abc_add_equal: - "\ap \ []; - abc_steps_l (0, am) ap astp = (a, b); - a < length ap\ - \ (abc_steps_l (0, am) (ap @ bp) astp) = (a, b)" -apply(induct astp arbitrary: a b, simp add: abc_steps_l.simps, simp) -apply(simp add: abc_steps_ind) -apply(case_tac "(abc_steps_l (0, am) ap astp)") -proof - - fix astp a b aa ba - assume ind: - "\a b. \abc_steps_l (0, am) ap astp = (a, b); - a < length ap\ \ - abc_steps_l (0, am) (ap @ bp) astp = (a, b)" - and h: "abc_steps_l (abc_steps_l (0, am) ap astp) ap (Suc 0) - = (a, b)" - "a < length ap" - "abc_steps_l (0, am) ap astp = (aa, ba)" - from h show "abc_steps_l (abc_steps_l (0, am) (ap @ bp) astp) - (ap @ bp) (Suc 0) = (a, b)" - apply(insert ind[of aa ba], simp) - apply(subgoal_tac "aa < length ap", simp) - apply(simp add: abc_steps_l.simps abc_fetch.simps - nth_append abc_steps_zero) - apply(rule abc_step_state_in, auto) - done -qed - - -lemma abc_add_exc1: - "\\ astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap\ - \ \ stp. abc_steps_l (0, am) (ap @ bp) stp = (as, bm)" -apply(case_tac "ap = []", simp, - rule_tac x = 0 in exI, simp add: abc_steps_zero) -apply(drule_tac abc_halt_point_ex, simp, simp) -apply(erule_tac exE, case_tac "(abc_steps_l (0, am) ap astp)", simp) -apply(rule_tac x = "Suc astp" in exI, simp add: abc_steps_ind, auto) -apply(frule_tac bp = bp in abc_add_equal, simp, simp, simp) -apply(simp add: abc_steps_l.simps abc_steps_zero - abc_fetch.simps nth_append) -done - -declare abc_shift.simps[simp del] +inductive_cases terminate_pr_reverse: "terminate (Pr n f g) xs" -lemma abc_append_exc2: - "\\ astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap; - \ bstp. abc_steps_l (0, bm) bp bstp = (bs, bm'); bs = length bp; - cs = as + bs; bp \ []\ - \ \ stp. abc_steps_l (0, am) (ap [+] bp) stp = (cs, bm')" -apply(insert abc_append_exc1[of bm bp bs bm' as ap "[]"], simp) -apply(drule_tac bp = "abc_shift bp (length ap)" in abc_add_exc1, simp) -apply(subgoal_tac "ap @ abc_shift bp (length ap) = ap [+] bp", - simp, auto) -apply(rule_tac x = "stpa + stp" in exI, simp add: abc_steps_add) -apply(simp add: abc_append.simps) -done -lemma exponent_add_iff: "a\b @ a\c@ xs = a\(b+c) @ xs" -apply(auto simp: replicate_add) -done - -lemma exponent_cons_iff: "a # a\c @ xs = a\(Suc c) @ xs" -apply(auto simp: replicate_add) -done - -lemma [simp]: "length lm = n \ - abc_steps_l (Suc 0, lm @ Suc x # 0 # suf_lm) - [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0)) - = (3, lm @ Suc x # 0 # suf_lm)" -apply(simp add: abc_steps_l.simps abc_fetch.simps - abc_step_l.simps abc_lm_v.simps abc_lm_s.simps - nth_append list_update_append) -done - -lemma [simp]: - "length lm = n \ - abc_steps_l (Suc 0, lm @ Suc x # Suc y # suf_lm) - [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0)) - = (Suc 0, lm @ Suc x # y # suf_lm)" -apply(simp add: abc_steps_l.simps abc_fetch.simps - abc_step_l.simps abc_lm_v.simps abc_lm_s.simps - nth_append list_update_append) -done - -lemma pr_cycle_part_middle_inv: - "\length lm = n\ \ - \ stp. abc_steps_l (0, lm @ x # y # suf_lm) - [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp - = (3, lm @ Suc x # 0 # suf_lm)" -proof - - assume h: "length lm = n" - hence k1: "\ stp. abc_steps_l (0, lm @ x # y # suf_lm) - [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp - = (Suc 0, lm @ Suc x # y # suf_lm)" - apply(rule_tac x = "Suc 0" in exI) - apply(simp add: abc_steps_l.simps abc_step_l.simps - abc_lm_v.simps abc_lm_s.simps nth_append - list_update_append abc_fetch.simps) - done - from h have k2: - "\ stp. abc_steps_l (Suc 0, lm @ Suc x # y # suf_lm) - [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp - = (3, lm @ Suc x # 0 # suf_lm)" - apply(induct y) - apply(rule_tac x = "Suc (Suc 0)" in exI, simp, simp, - erule_tac exE) - apply(rule_tac x = "Suc (Suc 0) + stp" in exI, - simp only: abc_steps_add, simp) - done - from k1 and k2 show - "\ stp. abc_steps_l (0, lm @ x # y # suf_lm) - [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp - = (3, lm @ Suc x # 0 # suf_lm)" - apply(erule_tac exE, erule_tac exE) - apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add) - done -qed - -lemma [simp]: - "length lm = Suc n \ - (abc_steps_l (length ap, lm @ x # Suc y # suf_lm) - (ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length ap)]) - (Suc (Suc (Suc 0)))) - = (length ap, lm @ Suc x # y # suf_lm)" -apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps - abc_lm_v.simps list_update_append nth_append abc_lm_s.simps) -done - -lemma switch_para_inv: - assumes bp_def:"bp = ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto ss]" - and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" - "ss = length ap" - "length lm = Suc n" - shows " \stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp = - (0, lm @ (x + y) # 0 # suf_lm)" -apply(induct y arbitrary: x) -apply(rule_tac x = "Suc 0" in exI, - simp add: bp_def 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: - "\x. \stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp = - (0, lm @ (x + y) # 0 # suf_lm)" - show "\stp. abc_steps_l (ss, lm @ x # Suc y # suf_lm) bp stp = - (0, lm @ (x + Suc y) # 0 # suf_lm)" - apply(insert ind[of "Suc x"], erule_tac exE) - apply(rule_tac x = "Suc (Suc (Suc 0)) + stp" in exI, - simp only: abc_steps_add bp_def h) - apply(simp add: h) - done -qed - -lemma [simp]: - "length lm = rs_pos \ Suc (Suc rs_pos) < a_md \ 0 < rs_pos \ - a_md - Suc 0 < Suc (Suc (Suc (a_md + length suf_lm - - Suc (Suc (Suc 0)))))" -apply(arith) -done - -lemma [simp]: - "Suc (Suc rs_pos) < a_md \ 0 < rs_pos \ - \ a_md - Suc 0 < rs_pos - Suc 0" -apply(arith) -done - -lemma [simp]: - "Suc (Suc rs_pos) < a_md \ 0 < rs_pos \ - \ a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))" -apply(arith) -done - -lemma butlast_append_last: "lm \ [] \ lm = butlast lm @ [last lm]" -apply(auto) -done - -lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md) - \ (Suc (Suc rs_pos)) < a_md" -apply(simp add: rec_ci.simps) -apply(case_tac "rec_ci f", simp) -apply(case_tac "rec_ci g", simp) -apply(arith) -done - -lemma ci_pr_para_eq: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md) - \ rs_pos = Suc n" -apply(simp add: rec_ci.simps) -apply(case_tac "rec_ci g", case_tac "rec_ci f", simp) -done - -lemma [intro]: - "\rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm xs\ - \ length lm = rs_pos" -apply(simp add: rec_ci.simps rec_ci_z_def) -apply(erule_tac calc_z_reverse, simp) -done - -lemma [intro]: - "\rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm xs\ - \ length lm = rs_pos" -apply(simp add: rec_ci.simps rec_ci_s_def) -apply(erule_tac calc_s_reverse, simp) -done - -lemma [intro]: - "\rec_ci (recf.id nat1 nat2) = (aprog, rs_pos, a_md); - rec_calc_rel (recf.id nat1 nat2) lm xs\ \ length lm = rs_pos" -apply(simp add: rec_ci.simps rec_ci_id.simps) -apply(erule_tac calc_id_reverse, simp) -done - -lemma [intro]: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_calc_rel (Cn n f gs) lm xs\ \ length lm = rs_pos" -apply(erule_tac calc_cn_reverse, simp) -apply(simp add: rec_ci.simps) -apply(case_tac "rec_ci f", simp) -done - -lemma [intro]: - "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_calc_rel (Pr n f g) lm xs\ \ length lm = rs_pos" -apply(erule_tac calc_pr_reverse, simp) -apply(drule_tac ci_pr_para_eq, simp, simp) -apply(drule_tac ci_pr_para_eq, simp) -done - -lemma [intro]: - "\rec_ci (Mn n f) = (aprog, rs_pos, a_md); - rec_calc_rel (Mn n f) lm xs\ \ length lm = rs_pos" -apply(erule_tac calc_mn_reverse) -apply(simp add: rec_ci.simps) -apply(case_tac "rec_ci f", simp) -done - -lemma para_pattern: - "\rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm xs\ - \ length lm = rs_pos" -apply(case_tac f, auto) -done - -lemma ci_pr_g_paras: - "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_ci g = (a, aa, ba); - rec_calc_rel (Pr n f g) (lm @ [x]) rs; x > 0\ \ - aa = Suc rs_pos " -apply(erule calc_pr_reverse, simp) -apply(subgoal_tac "length (args @ [k, rk]) = aa", simp) -apply(subgoal_tac "rs_pos = Suc n", simp) -apply(simp add: ci_pr_para_eq) -apply(erule para_pattern, simp) -done - -lemma ci_pr_g_md_less: - "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_ci g = (a, aa, ba)\ \ ba < a_md" -apply(simp add: rec_ci.simps) -apply(case_tac "rec_ci f", auto) -done - -lemma [intro]: "rec_ci z = (ap, rp, ad) \ rp < ad" - by(simp add: rec_ci.simps) - -lemma [intro]: "rec_ci s = (ap, rp, ad) \ rp < ad" - by(simp add: rec_ci.simps) - -lemma [intro]: "rec_ci (recf.id nat1 nat2) = (ap, rp, ad) \ rp < ad" - by(simp add: rec_ci.simps) - -lemma [intro]: "rec_ci (Cn n f gs) = (ap, rp, ad) \ rp < ad" -apply(simp add: rec_ci.simps) -apply(case_tac "rec_ci f", simp) -done - -lemma [intro]: "rec_ci (Pr n f g) = (ap, rp, ad) \ rp < ad" -apply(simp add: rec_ci.simps) -by(case_tac "rec_ci f", case_tac "rec_ci g", auto) - -lemma [intro]: "rec_ci (Mn n f) = (ap, rp, ad) \ rp < ad" -apply(simp add: rec_ci.simps) -apply(case_tac "rec_ci f", simp) -apply(arith) -done - -lemma ci_ad_ge_paras: "rec_ci f = (ap, rp, ad) \ ad > rp" -apply(case_tac f, auto) -done - -lemma [elim]: "\a [+] b = []; a \ [] \ b \ []\ \ RR" -apply(auto simp: abc_append.simps abc_shift.simps) -done - -lemma [intro]: "rec_ci z = ([], aa, ba) \ False" -by(simp add: rec_ci.simps rec_ci_z_def) +inductive_cases terminate_z_reverse[elim!]: "terminate z xs" -lemma [intro]: "rec_ci s = ([], aa, ba) \ False" -by(auto simp: rec_ci.simps rec_ci_s_def addition.simps) - -lemma [intro]: "rec_ci (id m n) = ([], aa, ba) \ False" -by(auto simp: rec_ci.simps rec_ci_id.simps addition.simps) - -lemma [intro]: "rec_ci (Cn n f gs) = ([], aa, ba) \ False" -apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_append.simps) -apply(simp add: abc_shift.simps mv_box.simps) -done - -lemma [intro]: "rec_ci (Pr n f g) = ([], aa, ba) \ False" -apply(simp add: rec_ci.simps) -apply(case_tac "rec_ci f", case_tac "rec_ci g") -by(auto) - -lemma [intro]: "rec_ci (Mn n f) = ([], aa, ba) \ False" -apply(case_tac "rec_ci f", auto simp: rec_ci.simps) -done - -lemma rec_ci_not_null: "rec_ci g = (a, aa, ba) \ a \ []" -by(case_tac g, auto) - -lemma calc_pr_g_def: - "\rec_calc_rel (Pr rs_pos f g) (lm @ [Suc x]) rsa; - rec_calc_rel (Pr rs_pos f g) (lm @ [x]) rsxa\ - \ rec_calc_rel g (lm @ [x, rsxa]) rsa" -apply(erule_tac calc_pr_reverse, simp, simp) -apply(subgoal_tac "rsxa = rk", simp) -apply(erule_tac rec_calc_inj, auto) -done - -lemma ci_pr_md_def: - "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\ - \ a_md = Suc (max (n + 3) (max bc ba))" -by(simp add: rec_ci.simps) - -lemma ci_pr_f_paras: - "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_calc_rel (Pr n f g) lm rs; - rec_ci f = (ab, ac, bc)\ \ ac = rs_pos - Suc 0" -apply(subgoal_tac "\rs. rec_calc_rel f (butlast lm) rs", - erule_tac exE) -apply(drule_tac f = f and lm = "butlast lm" in para_pattern, - simp, simp) -apply(drule_tac para_pattern, simp) -apply(subgoal_tac "lm \ []", simp) -apply(erule_tac calc_pr_reverse, simp, simp) -apply(erule calc_pr_zero_ex) -done - -lemma ci_pr_md_ge_f: "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_ci f = (ab, ac, bc)\ \ Suc bc \ a_md" -apply(case_tac "rec_ci g") -apply(simp add: rec_ci.simps, auto) -done - -lemma ci_pr_md_ge_g: "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_ci g = (ab, ac, bc)\ \ bc < a_md" -apply(case_tac "rec_ci f") -apply(simp add: rec_ci.simps, auto) -done - -lemma rec_calc_rel_def0: - "\rec_calc_rel (Pr n f g) lm rs; rec_calc_rel f (butlast lm) rsa\ - \ rec_calc_rel (Pr n f g) (butlast lm @ [0]) rsa" - apply(rule_tac calc_pr_zero, simp) -apply(erule_tac calc_pr_reverse, simp, simp, simp) -done - -lemma [simp]: "length (mv_box m n) = 3" -by (auto simp: mv_box.simps) - - -lemma [simp]: "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\ - \ rs_pos = Suc n" -apply(simp add: ci_pr_para_eq) -done - - -lemma [simp]: "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\ - \ length lm = Suc n" -apply(subgoal_tac "rs_pos = Suc n", rule_tac para_pattern, simp, simp) -apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps) -done - -lemma [simp]: "rec_ci (Pr n f g) = (a, rs_pos, a_md) \ Suc (Suc n) < a_md" -apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps) -apply arith -done - -lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md) \ 0 < rs_pos" -apply(case_tac "rec_ci f", case_tac "rec_ci g") -apply(simp add: rec_ci.simps) -done - -lemma [simp]: "Suc (Suc rs_pos) < a_md \ - butlast lm @ (last lm - xa) # (rsa::nat) # 0 # 0\(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm = - butlast lm @ (last lm - xa) # rsa # 0\(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: - "\lm rs suf_lm. rec_calc_rel g lm rs \ - \stp. abc_steps_l (0, lm @ 0\(ba - aa) @ suf_lm) a stp = - (length a, lm @ rs # 0\(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 \ last lm" - "rec_ci g = (a, aa, ba)" - "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsa" - "lm \ []" - shows - "\stp. abc_steps_l - (0, butlast lm @ (last lm - Suc xa) # rsxa # - 0\(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) ap stp = - (0, butlast lm @ (last lm - xa) # rsa - # 0\(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)" -proof - - have k1: "\stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # - rsxa # 0\(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) ap stp = - (length a + 4, butlast lm @ (last lm - xa) # 0 # rsa # - 0\(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\(a_md - Suc (Suc rs_pos)) @ xa # suf_lm" in abc_append_exc2, - auto) - proof - - show - "\astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa - # 0\(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\(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 \ rs_pos = Suc n \ - a_md > Suc (Suc rs_pos)") - apply(simp add: abc_lm_v.simps nth_append abc_lm_s.simps) - apply(insert nth_append[of - "(last lm - Suc xa) # rsxa # 0\(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\(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 "\bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # - rsxa # 0\(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\(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\(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 \ a_md > ba \ ba > Suc rs_pos" - apply(insert h) - apply(insert ci_pr_g_paras[of n f g aprog rs_pos - a_md a aa ba "butlast lm" "last lm - xa" rsa], simp) - apply(drule_tac ci_pr_md_ge_g, auto) - apply(erule_tac ci_ad_ge_paras) - done - from h have j2: "rec_calc_rel g (butlast lm @ - [last lm - Suc xa, rsxa]) rsa" - apply(rule_tac calc_pr_g_def, simp, simp) - done - from j1 and j2 show - "\astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # - rsxa # 0\(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) a astp = - (length a, butlast lm @ (last lm - Suc xa) # rsxa # rsa - # 0\(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)" - apply(insert g_ind[of - "butlast lm @ (last lm - Suc xa) # [rsxa]" rsa - "0\(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 \ rs_pos > 0" - apply(rule_tac conjI) - apply(drule_tac lm = "(butlast lm @ [last lm - Suc xa])" - and xs = rsxa in para_pattern, simp, simp, simp) - done - from h have j4: "Suc (last lm - Suc xa) = last lm - xa" - apply(case_tac "last lm", simp, simp) - done - from j3 and j4 show - "\bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa # - rsa # 0\(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\(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\(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm"], simp) - done - qed - qed - from h have k2: - "\stp. abc_steps_l (length a + 4, butlast lm @ (last lm - xa) # 0 - # rsa # 0\(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm) ap stp = - (0, butlast lm @ (last lm - xa) # rsa # 0\(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\(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm"]) - apply(simp add: h ap_def) - apply(subgoal_tac "length lm = Suc n \ Suc (Suc rs_pos) < a_md", - simp) - apply(insert h, simp) - apply(frule_tac lm = "(butlast lm @ [last lm - Suc xa])" - and xs = rsxa in para_pattern, simp, simp) - done - from k1 and k2 show "?thesis" - apply(auto) - apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add) - done -qed - -lemma ci_pr_ex1: - "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_ci g = (a, aa, ba); - rec_ci f = (ab, ac, bc)\ -\ \ap bp. length ap = 6 + length ab \ - aprog = ap [+] bp \ - bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] (a [+] - [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)])) @ - [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" -apply(simp add: rec_ci.simps) -apply(rule_tac x = "mv_box n (max (Suc (Suc (Suc n))) - (max bc ba)) [+] ab [+] mv_box n (Suc n)" in exI, - simp) -apply(auto simp add: abc_append_commute numeral_3_eq_3) -done +inductive_cases terminate_s_reverse[elim!]: "terminate s xs" -lemma pr_cycle_part: - "\\lm rs suf_lm. rec_calc_rel g lm rs \ - \stp. abc_steps_l (0, lm @ 0\(ba - aa) @ suf_lm) a stp = - (length a, lm @ rs # 0\(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 \ []; - x \ last lm\ \ - \stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) # - rsx # 0\(a_md - Suc (Suc rs_pos)) @ x # suf_lm) aprog stp = - (6 + length ab, butlast lm @ last lm # rs # - 0\(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)" -proof - - assume g_ind: - "\lm rs suf_lm. rec_calc_rel g lm rs \ - \stp. abc_steps_l (0, lm @ 0\(ba - aa) @ suf_lm) a stp = - (length a, lm @ rs # 0\(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 \ []" - "x \ last lm" - "rec_ci f = (ab, ac, bc)" - from h show - "\stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) # - rsx # 0\(a_md - Suc (Suc rs_pos)) @ x # suf_lm) aprog stp = - (6 + length ab, butlast lm @ last lm # rs # - 0\(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 \ [] \ rs_pos = Suc n", simp) - apply(rule_tac rec_calc_inj, simp, simp) - apply(simp) - done - thus "\stp. abc_steps_l (6 + length ab, butlast lm @ last lm # - rsxa # 0\(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm) aprog stp = - (6 + length ab, butlast lm @ last lm # rs # - 0\(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: - "\rsx. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsx - \ \stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - xa) # - rsx # 0\(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) aprog stp = - (6 + length ab, butlast lm @ last lm # rs # - 0\(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 \ last lm" - "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" - "rec_calc_rel (Pr n f g) lm rs" - "rec_ci g = (a, aa, ba)" - "rec_ci f = (ab, ac, bc)" "lm \ []" - from g have k1: - "\ rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rs" - apply(rule_tac rs = rs in calc_pr_less_ex, simp, simp) - done - from g and this show - "\stp. abc_steps_l (6 + length ab, - butlast lm @ (last lm - Suc xa) # rsxa # - 0\(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) aprog stp = - (6 + length ab, butlast lm @ last lm # rs # - 0\(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 - "\stp. abc_steps_l (6 + length ab, butlast lm @ - (last lm - Suc xa) # rsxa # - 0\(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) aprog stp - = (6 + length ab, butlast lm @ (last lm - xa) # rsa # - 0\(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)" - proof - - from g have k2_1: - "\ ap bp. length ap = 6 + length ab \ - aprog = ap [+] bp \ - bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] - (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, - Goto (Suc 0)])) @ - [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" - apply(rule_tac ci_pr_ex1, auto) - done - from k2_1 and k2 and g show "?thesis" - proof(erule_tac exE, erule_tac exE) - fix ap bp - assume - "length ap = 6 + length ab \ - aprog = ap [+] bp \ bp = - ([Dec (a_md - Suc 0) (length a + 7)] [+] - (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, - Goto (Suc 0)])) @ - [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" - from g and this and k2 and g_ind show "?thesis" - apply(insert abc_append_exc3[of - "butlast lm @ (last lm - Suc xa) # rsxa # - 0\(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm" bp 0 - "butlast lm @ (last lm - xa) # rsa # - 0\(a_md - Suc (Suc rs_pos)) @ xa # suf_lm" "length ap" ap], - simp) - apply(subgoal_tac - "\stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) - # rsxa # 0\(a_md - Suc (Suc rs_pos)) @ Suc xa # - suf_lm) bp stp = - (0, butlast lm @ (last lm - xa) # rsa # - 0\(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: - "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_ci g = (a, aa, ba); - rec_ci f = (ab, ac, bc)\ - \ length aprog = 13 + length ab + length a" -apply(auto simp: rec_ci.simps) -done - -fun mv_box_inv :: "nat \ nat list \ nat \ nat \ nat list \ bool" - where - "mv_box_inv (as, lm) m n initlm = - (let plus = initlm ! m + initlm ! n in - length initlm > max m n \ m \ n \ - (if as = 0 then \ k l. lm = initlm[m := k, n := l] \ - k + l = plus \ k \ initlm ! m - else if as = 1 then \ k l. lm = initlm[m := k, n := l] - \ k + l + 1 = plus \ k < initlm ! m - else if as = 2 then \ k l. lm = initlm[m := k, n := l] - \ k + l = plus \ k \ initlm ! m - else if as = 3 then lm = initlm[m := 0, n := plus] - else False))" - -fun mv_box_stage1 :: "nat \ nat list \ nat \ nat" - where - "mv_box_stage1 (as, lm) m = - (if as = 3 then 0 - else 1)" - -fun mv_box_stage2 :: "nat \ nat list \ nat \ nat" - where - "mv_box_stage2 (as, lm) m = (lm ! m)" - -fun mv_box_stage3 :: "nat \ nat list \ nat \ 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 \ nat list) \ nat) \ (nat \ nat \ 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 \ nat) \ nat \ nat) set" - where - "lex_pair = less_than <*lex*> less_than" - -definition lex_triple :: - "((nat \ (nat \ nat)) \ (nat \ (nat \ nat))) set" - where - "lex_triple \ less_than <*lex*> lex_pair" - -definition mv_box_LE :: - "(((nat \ nat list) \ nat) \ ((nat \ nat list) \ nat)) set" - where - "mv_box_LE \ (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: -"\m < length initlm; n < length initlm; m \ n\ \ - 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 (mv_box m n) = Some (Dec m 3)" -apply(simp add: mv_box.simps abc_fetch.simps) -done - -lemma [simp]: "abc_fetch (Suc 0) (mv_box m n) = - Some (Inc n)" -apply(simp add: mv_box.simps abc_fetch.simps) -done - -lemma [simp]: "abc_fetch 2 (mv_box m n) = Some (Goto 0)" -apply(simp add: mv_box.simps abc_fetch.simps) -done - -lemma [simp]: "abc_fetch 3 (mv_box m n) = None" -apply(simp add: mv_box.simps abc_fetch.simps) -done - -lemma [simp]: - "\m \ n; m < length initlm; n < length initlm; - k + l = initlm ! m + initlm ! n; k \ initlm ! m; 0 < k\ - \ \ka la. initlm[m := k, n := l, m := k - Suc 0] = - initlm[m := ka, n := la] \ - Suc (ka + la) = initlm ! m + initlm ! n \ - ka < initlm ! m" -apply(rule_tac x = "k - Suc 0" in exI, rule_tac x = l in exI, - simp, auto) -apply(subgoal_tac - "initlm[m := k, n := l, m := k - Suc 0] = - initlm[n := l, m := k, m := k - Suc 0]") -apply(simp add: list_update_overwrite ) -apply(simp add: list_update_swap) -apply(simp add: list_update_swap) -done - -lemma [simp]: - "\m \ n; m < length initlm; n < length initlm; - Suc (k + l) = initlm ! m + initlm ! n; - k < initlm ! m\ - \ \ka la. initlm[m := k, n := l, n := Suc l] = - initlm[m := ka, n := la] \ - ka + la = initlm ! m + initlm ! n \ - ka \ initlm ! m" -apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto) -done - -lemma [simp]: - "\length initlm > max m n; m \ n\ \ - \na. \ (\(as, lm) m. as = 3) - (abc_steps_l (0, initlm) (mv_box m n) na) m \ - mv_box_inv (abc_steps_l (0, initlm) - (mv_box m n) na) m n initlm \ - mv_box_inv (abc_steps_l (0, initlm) - (mv_box m n) (Suc na)) m n initlm \ - ((abc_steps_l (0, initlm) (mv_box m n) (Suc na), m), - abc_steps_l (0, initlm) (mv_box m n) na, m) \ mv_box_LE" -apply(rule allI, rule impI, simp add: abc_steps_ind) -apply(case_tac "(abc_steps_l (0, initlm) (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: - "\length initlm > max m n; m \ n\ \ - \ stp. (\ (as, lm). as = 3 \ - mv_box_inv (as, lm) m n initlm) - (abc_steps_l (0::nat, initlm) (mv_box m n) stp)" -apply(insert halt_lemma2[of mv_box_LE - "\ ((as, lm), m). mv_box_inv (as, lm) m n initlm" - "\ stp. (abc_steps_l (0, initlm) (mv_box m n) stp, m)" - "\ ((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) (mv_box m n) na)", - simp, auto) -done - -lemma mv_box_halt_cond: - "\m \ n; mv_box_inv (a, b) m n lm; a = 3\ \ - 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: - "\length lm > max m n; m \ n\ \ - \ 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) (mv_box m n) stp", - simp) -apply(erule_tac mv_box_halt_cond, auto) -done - -lemma [simp]: - "\a_md = Suc (max (Suc (Suc n)) (max bc ba)); - length lm = rs_pos \ rs_pos = n \ n > 0\ - \ n - Suc 0 < length lm + - (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm) \ - Suc (Suc n) < length lm + (Suc (max (Suc (Suc n)) (max bc ba)) - - rs_pos + length suf_lm) \ bc < length lm + (Suc (max (Suc (Suc n)) - (max bc ba)) - rs_pos + length suf_lm) \ ba < length lm + - (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm)" -apply(arith) -done - -lemma [simp]: - "\a_md = Suc (max (Suc (Suc n)) (max bc ba)); - length lm = rs_pos \ rs_pos = n \ n > 0\ - \ n - Suc 0 < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \ - Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \ - bc < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \ - ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))" -apply(arith) -done - -lemma [simp]: "n - Suc 0 \ max (Suc (Suc n)) (max bc ba)" -apply(arith) -done - -lemma [simp]: - "a_md \ Suc bc \ rs_pos > 0 \ bc \ rs_pos \ - bc - (rs_pos - Suc 0) + a_md - Suc bc = Suc (a_md - rs_pos - Suc 0)" -apply(arith) -done - -lemma [simp]: "length lm = n \ rs_pos = n \ 0 < rs_pos \ - Suc rs_pos < a_md - \ n - Suc 0 < Suc (Suc (a_md + length suf_lm - Suc (Suc 0))) - \ n < Suc (Suc (a_md + length suf_lm - Suc (Suc 0)))" -apply(arith) -done - -lemma [simp]: "length lm = n \ rs_pos = n \ 0 < rs_pos \ - Suc rs_pos < a_md \ n - Suc 0 \ n" -by arith - -lemma ci_pr_ex2: - "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_calc_rel (Pr n f g) lm rs; - rec_ci g = (a, aa, ba); - rec_ci f = (ab, ac, bc)\ - \ \ap bp. aprog = ap [+] bp \ - ap = mv_box n (max (Suc (Suc (Suc n))) (max bc ba))" -apply(simp add: rec_ci.simps) -apply(rule_tac x = "(ab [+] (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 - -lemma [simp]: "length lm = n \ rs_pos = n \ 0 < n \ - lm[n - Suc 0 := 0::nat] = butlast lm @ [0]" -apply(auto) -apply(insert list_update_append[of "butlast lm" "[last lm]" - "length lm - Suc 0" "0"], simp) -done - -lemma [simp]: "\length lm = n; 0 < n\ \ lm ! (n - Suc 0) = last lm" -apply(insert nth_append[of "butlast lm" "[last lm]" "n - Suc 0"], - simp) -apply(insert butlast_append_last[of lm], auto) -done -lemma exp_suc_iff: "a\b @ [a] = a\(b + Suc 0)" -apply(simp add: exp_ind del: replicate.simps) -done - -lemma less_not_less[simp]: "n > 0 \ \ n < n - Suc 0" -by auto - -lemma [simp]: - "Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \ - bc < Suc (length suf_lm + max (Suc (Suc n)) - (max bc ba)) \ - ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))" - by arith - -lemma [simp]: "length lm = n \ rs_pos = n \ n > 0 \ -(lm @ 0\(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) - [max (Suc (Suc n)) (max bc ba) := - (lm @ 0\(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) ! (n - Suc 0) + - (lm @ 0\(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\(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\((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\a = c\b)" -apply(auto) -done - -lemma [simp]: - "\length lm = n; 0 < n; Suc n < a_md\ \ - (butlast lm @ rsa # 0\(a_md - Suc n) @ last lm # suf_lm) - [n := (butlast lm @ rsa # 0\(a_md - Suc n) @ last lm # suf_lm) ! - (n - Suc 0) + (butlast lm @ rsa # (0::nat)\(a_md - Suc n) @ - last lm # suf_lm) ! n, n - Suc 0 := 0] - = butlast lm @ 0 # rsa # 0\(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) \ a_md \ length lm = rs_pos \ 0 < rs_pos - \ a_md - Suc 0 < - Suc (Suc (Suc (a_md + length suf_lm - Suc (Suc (Suc 0)))))" -by arith - -lemma [simp]: - "Suc (Suc rs_pos) \ a_md \ length lm = rs_pos \ 0 < rs_pos \ - \ a_md - Suc 0 < rs_pos - Suc 0" -by arith - -lemma [simp]: "Suc (Suc rs_pos) \ a_md \ - \ a_md - Suc 0 < rs_pos - Suc 0" -by arith - -lemma [simp]: "\Suc (Suc rs_pos) \ a_md\ \ - \ a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))" -by arith - -lemma [simp]: - "Suc (Suc rs_pos) \ a_md \ length lm = rs_pos \ 0 < rs_pos - \ (abc_lm_v (butlast lm @ last lm # rs # 0\(a_md - Suc (Suc rs_pos)) @ - 0 # suf_lm) (a_md - Suc 0) = 0 \ - abc_lm_s (butlast lm @ last lm # rs # 0\(a_md - Suc (Suc rs_pos)) @ - 0 # suf_lm) (a_md - Suc 0) 0 = - lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm) \ - abc_lm_v (butlast lm @ last lm # rs # 0\(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\(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]: "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\ - \ \cp. aprog = mv_box n (max (n + 3) - (max bc ba)) [+] cp" -apply(simp add: rec_ci.simps) -apply(rule_tac x = "(ab [+] (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 \ []" -by (simp add: mv_box.simps) -(* -lemma [simp]: "\rs_pos = n; 0 < rs_pos ; Suc rs_pos < a_md\ \ - n - Suc 0 < a_md + length suf_lm" -by arith -*) -lemma [intro]: - "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_ci f = (ab, ac, bc)\ \ - \ap. (\cp. aprog = ap [+] ab [+] cp) \ length ap = 3" -apply(case_tac "rec_ci g", simp add: rec_ci.simps) -apply(rule_tac x = "mv_box n - (max (n + 3) (max bc c))" in exI, simp) -apply(rule_tac x = "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]: - "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_ci g = (a, aa, ba); - rec_ci f = (ab, ac, bc)\ \ - \ap. (\cp. aprog = ap [+] mv_box n (Suc n) [+] cp) - \ length ap = 3 + length ab" -apply(simp add: rec_ci.simps) -apply(rule_tac x = "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]: - "\rec_ci (Pr n f g) = (aprog, rs_pos, a_md); - rec_ci g = (a, aa, ba); - rec_ci f = (ab, ac, bc)\ - \ \ap. (\cp. aprog = ap [+] ([Dec (a_md - Suc 0) (length a + 7)] - [+] (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, - Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), - Goto (length a + 4)] [+] cp) \ - length ap = 6 + length ab" -apply(simp add: rec_ci.simps) -apply(rule_tac x = "mv_box n - (max (n + 3) (max bc ba)) [+] ab [+] - 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) \ - Suc (Suc n) < max (n + 3) (max bc ba) + length suf_lm \ - bc < Suc (max (n + 3) (max bc ba) + length suf_lm) \ - ba < Suc (max (n + 3) (max bc ba) + length suf_lm)" -by arith - -lemma [simp]: "n \ max (n + (3::nat)) (max bc ba)" -by arith - -lemma [simp]:"length lm = Suc n \ lm[n := (0::nat)] = butlast lm @ [0]" -apply(subgoal_tac "\ xs x. lm = xs @ [x]", auto simp: list_update_append) -apply(rule_tac x = "butlast lm" in exI, rule_tac x = "last lm" in exI) -apply(case_tac lm, auto) -done - -lemma [simp]: "length lm = Suc n \ lm ! n =last lm" -apply(subgoal_tac "lm \ []") -apply(simp add: last_conv_nth, case_tac lm, simp_all) -done - -lemma [simp]: "length lm = Suc n \ - (lm @ (0::nat)\(max (n + 3) (max bc ba) - n) @ suf_lm) - [max (n + 3) (max bc ba) := (lm @ 0\(max (n + 3) (max bc ba) - n) @ suf_lm) ! n + - (lm @ 0\(max (n + 3) (max bc ba) - n) @ suf_lm) ! max (n + 3) (max bc ba), n := 0] - = butlast lm @ 0 # 0\(max (n + 3) (max bc ba) - Suc n) @ last lm # suf_lm" -apply(auto simp: list_update_append nth_append) -apply(subgoal_tac "(0\(max (n + 3) (max bc ba) - n)) = 0\(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 \ - n < Suc (Suc (a_md + length suf_lm - 2)) \ - n < Suc (a_md + length suf_lm - 2)" -by(arith) - -lemma [simp]: "\length lm = Suc n; Suc (Suc n) < a_md\ - \(butlast lm @ (rsa::nat) # 0\(a_md - Suc (Suc n)) @ last lm # suf_lm) - [Suc n := (butlast lm @ rsa # 0\(a_md - Suc (Suc n)) @ last lm # suf_lm) ! n + - (butlast lm @ rsa # 0\(a_md - Suc (Suc n)) @ last lm # suf_lm) ! Suc n, n := 0] - = butlast lm @ 0 # rsa # 0\(a_md - Suc (Suc (Suc n))) @ last lm # suf_lm" -apply(auto simp: list_update_append) -apply(subgoal_tac "(0\(a_md - Suc (Suc n))) = (0::nat) # (0\(a_md - Suc (Suc (Suc n))))", simp add: nth_append) -apply(simp add: replicate_Suc[THEN sym]) -done +inductive_cases terminate_id_reverse[elim!]: "terminate (id m n) xs" -lemma pr_case: - assumes nf_ind: - "\ lm rs suf_lm. rec_calc_rel f lm rs \ - \stp. abc_steps_l (0, lm @ 0\(bc - ac) @ suf_lm) ab stp = - (length ab, lm @ rs # 0\(bc - Suc ac) @ suf_lm)" - and ng_ind: "\ lm rs suf_lm. rec_calc_rel g lm rs \ - \stp. abc_steps_l (0, lm @ 0\(ba - aa) @ suf_lm) a stp = - (length a, lm @ rs # 0\(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 "\stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = (length aprog, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm)" -proof - - from h have k1: "\ stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp - = (3, butlast lm @ 0 # 0\(a_md - rs_pos - 1) @ last lm # suf_lm)" - proof - - have "\bp cp. aprog = bp [+] cp \ 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 - "\stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) - ([] [+] mv_box n - (max (n + 3) (max bc ba)) [+] cp) stp = - (0 + 3, butlast lm @ 0 # 0\(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\(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 \ rs_pos = Suc n", simp) - apply(insert h) - apply(simp add: para_pattern ci_pr_para_eq) - apply(rule ci_pr_md_def, auto) - done - qed - from h have k2: - "\ stp. abc_steps_l (3, butlast lm @ 0 # 0\(a_md - rs_pos - 1) @ - last lm # suf_lm) aprog stp - = (length aprog, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm)" - proof - - from h have k2_1: "\ rs. rec_calc_rel f (butlast lm) rs" - apply(erule_tac calc_pr_zero_ex) - done - thus "?thesis" - proof(erule_tac exE) - fix rsa - assume k2_2: "rec_calc_rel f (butlast lm) rsa" - from h and k2_2 have k2_2_1: - "\ stp. abc_steps_l (3, butlast lm @ 0 # 0\(a_md - rs_pos - 1) - @ last lm # suf_lm) aprog stp - = (3 + length ab, butlast lm @ rsa # 0\(a_md - rs_pos - 1) @ - last lm # suf_lm)" - proof - - from h have j1: " - \ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = 3 \ - bp = ab" - apply(auto) - done - from h have j2: "ac = rs_pos - 1" - apply(drule_tac ci_pr_f_paras, simp, auto) - done - from h and j2 have j3: "a_md \ Suc bc \ rs_pos > 0 \ bc \ rs_pos" - apply(rule_tac conjI) - apply(erule_tac ab = ab and ac = ac in ci_pr_md_ge_f, simp) - apply(rule_tac context_conjI) - apply(simp_all add: rec_ci.simps) - apply(drule_tac ci_ad_ge_paras, drule_tac ci_ad_ge_paras) - apply(arith) - done - from j1 and j2 show "?thesis" - apply(auto simp del: abc_append_commute) - apply(rule_tac abc_append_exc1, simp_all) - apply(insert nf_ind[of "butlast lm" "rsa" - "0\(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: - "\ stp. abc_steps_l (3 + length ab, butlast lm @ rsa # - 0\(a_md - rs_pos - 1) @ last lm # suf_lm) aprog stp - = (6 + length ab, butlast lm @ 0 # rsa # - 0\(a_md - rs_pos - 2) @ last lm # suf_lm)" - proof - - from h have "\ ap bp cp. aprog = ap [+] bp [+] cp \ - length ap = 3 + length ab \ bp = 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 \ length ap = 3 + - length ab \ bp = mv_box n (Suc n)" - thus "?thesis" - apply(simp del: abc_append_commute) - apply(subgoal_tac - "\stp. abc_steps_l (3 + length ab, - butlast lm @ rsa # 0\(a_md - Suc rs_pos) @ - last lm # suf_lm) (ap [+] - mv_box n (Suc n) [+] cp) stp = - ((3 + length ab) + 3, butlast lm @ 0 # rsa # - 0\(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\(a_md - Suc rs_pos) @ - last lm # suf_lm"], simp) - apply(subgoal_tac "length lm = Suc n \ rs_pos = Suc n \ a_md > Suc (Suc n)", simp) - apply(insert h, simp) - done - qed - qed - from h have k2_3: "lm \ []" - apply(rule_tac calc_pr_para_not_null, simp) - done - from h and k2_2 and k2_3 have k2_2_3: - "\ stp. abc_steps_l (6 + length ab, butlast lm @ - (last lm - last lm) # rsa # - 0\(a_md - (Suc (Suc rs_pos))) @ last lm # suf_lm) aprog stp - = (6 + length ab, butlast lm @ last lm # rs # - 0\(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: - "\ stp. abc_steps_l (6 + length ab, - butlast lm @ last lm # rs # 0\(a_md - rs_pos - 2) @ - 0 # suf_lm) aprog stp - = (13 + length ab + length a, - lm @ rs # 0\(a_md - rs_pos - 1) @ suf_lm)" - proof - - from h have - "\ ap bp cp. aprog = ap [+] bp [+] cp \ - length ap = 6 + length ab \ - bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] - (a [+] [Inc (rs_pos - Suc 0), - Dec rs_pos 3, Goto (Suc 0)])) @ - [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" - by auto - thus "?thesis" - apply(auto) - apply(subgoal_tac - "\stp. abc_steps_l (6 + length ab, butlast lm @ - last lm # rs # 0\(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\(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) \ - length lm = rs_pos \ rs_pos > 0", simp) - apply(insert h, simp) - apply(subgoal_tac "rs_pos = Suc n", simp, simp) - done - qed - from h have k2_2_5: "length aprog = 13 + length ab + length a" - apply(rule_tac ci_pr_length, simp_all) - done - from k2_2_1 and k2_2_2 and k2_2_3 and k2_2_4 and k2_2_5 - show "?thesis" - apply(auto) - apply(rule_tac x = "stp + stpa + stpb + stpc" in exI, - simp add: abc_steps_add) - done - qed - qed - from k1 and k2 show - "\stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp - = (length aprog, lm @ rs # 0\(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 - -lemma eq_switch: "x = y \ y = x" -by simp - -lemma [simp]: - "\rec_ci f = (a, aa, ba); - rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ \ \bp. aprog = a @ bp" -apply(simp add: rec_ci.simps) -apply(rule_tac x = "[Dec (Suc n) (length a + 5), - Dec (Suc n) (length a + 3), Goto (Suc (length a)), - Inc n, Goto 0]" in exI, auto) -done - -lemma ci_mn_para_eq[simp]: - "rec_ci (Mn n f) = (aprog, rs_pos, a_md) \ rs_pos = n" -apply(case_tac "rec_ci f", simp add: rec_ci.simps) -done - -lemma [simp]: "rec_ci f = (a, aa, ba) \ aa < ba" -apply(simp add: ci_ad_ge_paras) -done - -lemma [simp]: "\rec_ci f = (a, aa, ba); - rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ - \ ba \ a_md" -apply(simp add: rec_ci.simps) -by arith - -lemma mn_calc_f: - assumes ind: - "\aprog a_md rs_pos rs suf_lm lm. - \rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\ - \ \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp - = (length aprog, lm @ [rs] @ 0\(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 "\stp. abc_steps_l (0, lm @ x # 0\(a_md - Suc rs_pos) @ suf_lm) - aprog stp = (length a, - lm @ x # rsx # 0\(a_md - Suc (Suc rs_pos)) @ suf_lm)" -proof - - from h have k1: "\ ap bp. aprog = ap @ bp \ ap = a" - by simp - from h have k2: "rs_pos = n" - apply(erule_tac ci_mn_para_eq) - done - from h and k1 and k2 show "?thesis" - - proof(erule_tac exE, erule_tac exE, simp, - rule_tac abc_add_exc1, auto) - fix bp - show - "\astp. abc_steps_l (0, lm @ x # 0\(a_md - Suc n) @ suf_lm) a astp - = (length a, lm @ x # rsx # 0\(a_md - Suc (Suc n)) @ suf_lm)" - apply(insert ind[of a "Suc n" ba "lm @ [x]" rsx - "0\(a_md - ba) @ suf_lm"], simp add: exponent_add_iff h k2) - apply(subgoal_tac "ba > aa \ a_md \ ba \ aa = Suc n", - insert h, auto) - done - qed -qed - -fun mn_ind_inv :: - "nat \ nat list \ nat \ nat \ nat \ nat list \ nat list \ bool" - where - "mn_ind_inv (as, lm') ss x rsx suf_lm lm = - (if as = ss then lm' = lm @ x # rsx # suf_lm - else if as = ss + 1 then - \y. (lm' = lm @ x # y # suf_lm) \ y \ rsx - else if as = ss + 2 then - \y. (lm' = lm @ x # y # suf_lm) \ y \ rsx - else if as = ss + 3 then lm' = lm @ x # 0 # suf_lm - else if as = ss + 4 then lm' = lm @ Suc x # 0 # suf_lm - else if as = 0 then lm' = lm @ Suc x # 0 # suf_lm - else False -)" - -fun mn_stage1 :: "nat \ nat list \ nat \ nat \ nat" - where - "mn_stage1 (as, lm) ss n = - (if as = 0 then 0 - else if as = ss + 4 then 1 - else if as = ss + 3 then 2 - else if as = ss + 2 \ as = ss + 1 then 3 - else if as = ss then 4 - else 0 -)" - -fun mn_stage2 :: "nat \ nat list \ nat \ nat \ nat" - where - "mn_stage2 (as, lm) ss n = - (if as = ss + 1 \ as = ss + 2 then (lm ! (Suc n)) - else 0)" - -fun mn_stage3 :: "nat \ nat list \ nat \ nat \ nat" - where - "mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)" - - -fun mn_measure :: "((nat \ nat list) \ nat \ nat) \ - (nat \ nat \ nat)" - where - "mn_measure ((as, lm), ss, n) = - (mn_stage1 (as, lm) ss n, mn_stage2 (as, lm) ss n, - mn_stage3 (as, lm) ss n)" - -definition mn_LE :: "(((nat \ nat list) \ nat \ nat) \ - ((nat \ nat list) \ nat \ nat)) set" - where "mn_LE \ (inv_image lex_triple mn_measure)" - -lemma wf_mn_le[intro]: "wf mn_LE" -by(auto intro:wf_inv_image wf_lex_triple simp: mn_LE_def) - -declare mn_ind_inv.simps[simp del] - -lemma mn_inv_init: - "mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0) - (length a) x rsx suf_lm lm" -apply(simp add: mn_ind_inv.simps abc_steps_zero) -done - -lemma mn_halt_init: - "rec_ci f = (a, aa, ba) \ - \ (\(as, lm') (ss, n). as = 0) - (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0) - (length a, n)" -apply(simp add: abc_steps_zero) -apply(erule_tac rec_ci_not_null) -done +inductive_cases terminate_cn_reverse[elim!]: "terminate (Cn n f gs) xs" -lemma [simp]: - "\rec_ci f = (a, aa, ba); - rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ - \ abc_fetch (length a) aprog = - Some (Dec (Suc n) (length a + 5))" -apply(simp add: rec_ci.simps abc_fetch.simps, - erule_tac conjE, erule_tac conjE, simp) -apply(drule_tac eq_switch, drule_tac eq_switch, simp) -done - -lemma [simp]: "\rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ - \ abc_fetch (Suc (length a)) aprog = Some (Dec (Suc n) (length a + 3))" -apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp) -apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append) -done - -lemma [simp]: - "\rec_ci f = (a, aa, ba); - rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ - \ abc_fetch (Suc (Suc (length a))) aprog = - Some (Goto (length a + 1))" -apply(simp add: rec_ci.simps abc_fetch.simps, - erule_tac conjE, erule_tac conjE, simp) -apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append) -done - -lemma [simp]: - "\rec_ci f = (a, aa, ba); - rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ - \ abc_fetch (length a + 3) aprog = Some (Inc n)" -apply(simp add: rec_ci.simps abc_fetch.simps, - erule_tac conjE, erule_tac conjE, simp) -apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append) -done - -lemma [simp]: "\rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ - \ abc_fetch (length a + 4) aprog = Some (Goto 0)" -apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp) -apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append) -done - -lemma [simp]: - "0 < rsx - \ \y. (lm @ x # rsx # suf_lm)[Suc (length lm) := rsx - Suc 0] - = lm @ x # y # suf_lm \ y \ rsx" -apply(case_tac rsx, simp, simp) -apply(rule_tac x = nat in exI, simp add: list_update_append) -done - -lemma [simp]: - "\y \ rsx; 0 < y\ - \ \ya. (lm @ x # y # suf_lm)[Suc (length lm) := y - Suc 0] - = lm @ x # ya # suf_lm \ ya \ rsx" -apply(case_tac y, simp, simp) -apply(rule_tac x = nat in exI, simp add: list_update_append) -done - -lemma mn_halt_lemma: - "\rec_ci f = (a, aa, ba); - rec_ci (Mn n f) = (aprog, rs_pos, a_md); - 0 < rsx; length lm = n\ - \ - \na. \ (\(as, lm') (ss, n). as = 0) - (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na) - (length a, n) - \ mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) - aprog na) (length a) x rsx suf_lm lm -\ mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog - (Suc na)) (length a) x rsx suf_lm lm - \ ((abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog (Suc na), - length a, n), - abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na, - length a, n) \ mn_LE" -apply(rule allI, rule impI, simp add: abc_steps_ind) -apply(case_tac "(abc_steps_l (length a, lm @ x # rsx # suf_lm) - aprog na)", simp) -apply(auto split:if_splits simp add:abc_steps_l.simps - mn_ind_inv.simps abc_steps_zero) -apply(auto simp add: mn_LE_def lex_triple_def lex_pair_def - abc_step_l.simps abc_steps_l.simps mn_ind_inv.simps - abc_lm_v.simps abc_lm_s.simps nth_append - split: if_splits) -apply(drule_tac rec_ci_not_null, simp) -done - -lemma mn_halt: - "\rec_ci f = (a, aa, ba); - rec_ci (Mn n f) = (aprog, rs_pos, a_md); - 0 < rsx; length lm = n\ - \ \ stp. (\ (as, lm'). (as = 0 \ - mn_ind_inv (as, lm') (length a) x rsx suf_lm lm)) - (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp)" -apply(insert wf_mn_le) -apply(insert halt_lemma2[of mn_LE - "\ ((as, lm'), ss, n). mn_ind_inv (as, lm') ss x rsx suf_lm lm" - "\ stp. (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp, - length a, n)" - "\ ((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 \ - Suc (a_md - Suc (Suc rs_pos)) = a_md - Suc rs_pos" -by arith - -lemma mn_ind_step: - assumes ind: - "\aprog a_md rs_pos rs suf_lm lm. - \rec_ci f = (aprog, rs_pos, a_md); - rec_calc_rel f lm rs\ \ - \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp - = (length aprog, lm @ [rs] @ 0\(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 "\stp. abc_steps_l (0, lm @ x # 0\(a_md - Suc rs_pos) @ suf_lm) - aprog stp = (0, lm @ Suc x # 0\(a_md - Suc rs_pos) @ suf_lm)" -proof - - have k1: - "\ stp. abc_steps_l (0, lm @ x # 0\(a_md - Suc (rs_pos)) @ suf_lm) - aprog stp = - (length a, lm @ x # rsx # 0\(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: - "\ stp. abc_steps_l (length a, - lm @ x # rsx # 0\(a_md - Suc (Suc rs_pos)) @ suf_lm) aprog stp = - (0, lm @ Suc x # 0\(a_md - rs_pos - 1) @ suf_lm)" - apply(frule_tac x = x and - suf_lm = "0\(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]: - "\rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md); - rec_calc_rel (Mn n f) lm rs\ \ aa = Suc rs_pos" -apply(rule_tac calc_mn_reverse, simp) -apply(insert para_pattern [of f a aa ba "lm @ [rs]" 0], simp) -apply(subgoal_tac "rs_pos = length lm", simp) -apply(drule_tac ci_mn_para_eq, simp) -done - -lemma [simp]: "\rec_ci (Mn n f) = (aprog, rs_pos, a_md); - rec_calc_rel (Mn n f) lm rs\ \ Suc rs_pos < a_md" -apply(case_tac "rec_ci f") -apply(subgoal_tac "c > b \ b = Suc rs_pos \ a_md \ c") -apply(arith, auto) -done - -lemma mn_ind_steps: - assumes ind: - "\aprog a_md rs_pos rs suf_lm lm. - \rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\ \ - \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (length aprog, lm @ [rs] @ 0\(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" - "\x v. rec_calc_rel f (lm @ [x]) v \ 0 < v)" - "n = length lm" - "x \ rs" - shows "\stp. abc_steps_l (0, lm @ 0 # 0\(a_md - Suc rs_pos) @ suf_lm) - aprog stp = (0, lm @ x # 0\(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: - "\stp. abc_steps_l (0, lm @ 0 # 0\(a_md - Suc rs_pos) @ suf_lm) - aprog stp = (0, lm @ x # 0\(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" - "\x v. rec_calc_rel f (lm @ [x]) v \ v > 0)" - "n = length lm" - "Suc x \ rs" - "rec_ci f = (a, aa, ba)" - hence k2: - "\stp. abc_steps_l (0, lm @ x # 0\(a_md - rs_pos - 1) @ suf_lm) aprog - stp = (0, lm @ Suc x # 0\(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 - "\stp. abc_steps_l (0, lm @ 0 # 0\(a_md - Suc rs_pos) @ suf_lm) - aprog stp = (0, lm @ Suc x # 0\(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]: -"\rec_ci f = (a, aa, ba); - rec_ci (Mn n f) = (aprog, rs_pos, a_md); - rec_calc_rel (Mn n f) lm rs; - length lm = n\ - \ abc_lm_v (lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm) (Suc n) = 0" -apply(auto simp: abc_lm_v.simps nth_append) -done - -lemma [simp]: - "\rec_ci f = (a, aa, ba); - rec_ci (Mn n f) = (aprog, rs_pos, a_md); - rec_calc_rel (Mn n f) lm rs; - length lm = n\ - \ abc_lm_s (lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm) (Suc n) 0 = - lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm" -apply(auto simp: abc_lm_s.simps list_update_append) -done - -lemma mn_length: - "\rec_ci f = (a, aa, ba); - rec_ci (Mn n f) = (aprog, rs_pos, a_md)\ - \ length aprog = length a + 5" -apply(simp add: rec_ci.simps, erule_tac conjE) -apply(drule_tac eq_switch, drule_tac eq_switch, simp) -done - -lemma mn_final_step: - assumes ind: - "\aprog a_md rs_pos rs suf_lm lm. - \rec_ci f = (aprog, rs_pos, a_md); - rec_calc_rel f lm rs\ \ - \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (length aprog, lm @ [rs] @ 0\(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 "\stp. abc_steps_l (0, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm) - aprog stp = (length aprog, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm)" -proof - - from h and ind have k1: - "\stp. abc_steps_l (0, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm) - aprog stp = (length a, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm)" - apply(insert mn_calc_f[of f a aa ba n aprog - rs_pos a_md lm rs 0 suf_lm], simp) - apply(subgoal_tac "aa = Suc n", simp add: exponent_cons_iff) - apply(subgoal_tac "rs_pos = n", simp, simp) - done - from h have k2: "length lm = n" - apply(subgoal_tac "rs_pos = n") - apply(drule_tac f = "Mn n f" in para_pattern, simp, simp, simp) - done - from h and k2 have k3: - "\stp. abc_steps_l (length a, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm) - aprog stp = (length a + 5, lm @ rs # 0\(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: - "\aprog a_md rs_pos rs suf_lm lm. - \rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\ \ - \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (length aprog, lm @ [rs] @ 0\(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 "\stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp - = (length aprog, lm @ [rs] @ 0\(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" - "\xv. rec_calc_rel f (lm @ [x]) v \ 0 < v" - "n = length lm" - hence k1: - "\stp. abc_steps_l (0, lm @ 0 # 0\(a_md - Suc rs_pos) @ suf_lm) aprog - stp = (0, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm)" - apply(auto intro: mn_ind_steps ind) - done - from h have k2: - "\stp. abc_steps_l (0, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm) aprog - stp = (length aprog, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm)" - apply(auto intro: mn_final_step ind) - done - from k1 and k2 show - "\stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (length aprog, lm @ rs # 0\(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 \ rs = 0" -apply(rule_tac calc_z_reverse, auto) -done - -lemma z_case: - "\rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm rs\ - \ \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (length aprog, lm @ [rs] @ 0\(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 +inductive_cases terminate_mn_reverse[elim!]:"terminate (Mn n f) xs" fun addition_inv :: "nat \ nat list \ nat \ nat \ nat \ nat list \ bool" @@ -2427,14 +166,15 @@ where "addition_LE \ (inv_image lex_triple addition_measure)" lemma [simp]: "wf addition_LE" -by(simp add: wf_inv_image wf_lex_triple addition_LE_def) +by(auto simp: wf_inv_image addition_LE_def lex_triple_def + lex_pair_def) declare addition_inv.simps[simp del] lemma addition_inv_init: "\m \ n; max m n < p; length lm > p; lm ! p = 0\ \ addition_inv (0, lm) m n p lm" -apply(simp add: addition_inv.simps) +apply(simp add: addition_inv.simps Let_def) apply(rule_tac x = "lm ! m" in exI, simp) done @@ -2526,6 +266,11 @@ apply(rule_tac x = "Suc x" in exI, simp) done +lemma abc_steps_zero: "abc_steps_l asm ap 0 = asm" +apply(case_tac asm, simp add: abc_steps_l.simps) +done + +declare Let_def[simp] lemma addition_halt_lemma: "\m \ n; max m n < p; length lm > p; lm ! p = 0\ \ \na. \ (\(as, lm') (m, p). as = 7) @@ -2535,19 +280,19 @@ (Suc na)) m n p lm \ ((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p), abc_steps_l (0, lm) (addition m n p) na, m, p) \ addition_LE" -apply(rule allI, rule impI, simp add: abc_steps_ind) +apply(rule allI, rule impI, simp add: abc_step_red2) 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(simp_all add: addition.simps abc_steps_l.simps) 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 +apply(rule_tac [!] x = x in exI, simp_all add: list_update_overwrite Suc_diff_Suc) +by (metis list_update_overwrite list_update_swap nat_neq_iff) -lemma addition_ex: +lemma addition_correct': "\m \ n; max m n < p; length lm > p; lm ! p = 0\ \ \ stp. (\ (as, lm'). as = 7 \ addition_inv (as, lm') m n p lm) (abc_steps_l (0, lm) (addition m n p) stp)" @@ -2562,315 +307,474 @@ case_tac "(abc_steps_l (0, lm) (addition m n p) na)", auto) done -lemma [simp]: "length (addition m n p) = 7" -by (simp add: addition.simps) - -lemma [elim]: "addition 0 (Suc 0) 2 = [] \ RR" -by(simp add: addition.simps) - -lemma [simp]: "(0\2)[0 := n] = [n, 0::nat]" -apply(subgoal_tac "2 = Suc 1", - simp only: replicate.simps) -apply(auto) -done +lemma length_addition[simp]: "length (addition a b c) = 7" +by(auto simp: addition.simps) -lemma [simp]: - "\stp. abc_steps_l (0, n # 0\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\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\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 addition_correct: + "\m \ n; max m n < p; length lm > p; lm ! p = 0\ + \ {\ a. a = lm} (addition m n p) {\ nl. addition_inv (7, nl) m n p lm}" +using assms +proof(rule_tac abc_Hoare_haltI, simp) + fix lma + assume "m \ n" "m < p \ n < p" "p < length lm" "lm ! p = 0" + then have "\ stp. (\ (as, lm'). as = 7 \ addition_inv (as, lm') m n p lm) + (abc_steps_l (0, lm) (addition m n p) stp)" + by(rule_tac addition_correct', auto simp: addition_inv.simps) + thus "\na. abc_final (abc_steps_l (0, lm) (addition m n p) na) (addition m n p) \ + (\nl. addition_inv (7, nl) m n p lm) abc_holds_for abc_steps_l (0, lm) (addition m n p) na" + apply(erule_tac exE) + apply(rule_tac x = stp in exI) + apply(auto) + done +qed -lemma s_case: - "\rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm rs\ - \ \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (length aprog, lm @ [rs] @ 0\(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 compile_s_correct': + "{\nl. nl = n # 0 \ 2 @ anything} addition 0 (Suc 0) 2 [+] [Inc (Suc 0)] {\nl. nl = n # Suc n # 0 # anything}" +proof(rule_tac abc_Hoare_plus_halt) + show "{\nl. nl = n # 0 \ 2 @ anything} addition 0 (Suc 0) 2 {\ nl. addition_inv (7, nl) 0 (Suc 0) 2 (n # 0 \ 2 @ anything)}" + by(rule_tac addition_correct, auto simp: numeral_2_eq_2) +next + show "{\nl. addition_inv (7, nl) 0 (Suc 0) 2 (n # 0 \ 2 @ anything)} [Inc (Suc 0)] {\nl. nl = n # Suc n # 0 # anything}" + by(rule_tac abc_Hoare_haltI, rule_tac x = 1 in exI, auto simp: addition_inv.simps + abc_steps_l.simps abc_step_l.simps abc_fetch.simps numeral_2_eq_2 abc_lm_s.simps abc_lm_v.simps) +qed + +declare rec_exec.simps[simp del] -lemma [simp]: - "\n < length lm; lm ! n = rs\ - \ \stp. abc_steps_l (0, lm @ 0 # 0 #suf_lm) - (addition n (length lm) (Suc (length lm))) stp - = (7, lm @ rs # 0 # suf_lm)" -apply(insert addition_ex[of n "length lm" - "Suc (length lm)" "lm @ 0 # 0 # suf_lm"]) -apply(simp add: nth_append, erule_tac exE) -apply(rule_tac x = stp in exI) -apply(case_tac "abc_steps_l (0, lm @ 0 # 0 # suf_lm) (addition n (length lm) - (Suc (length lm))) stp", simp) -apply(simp add: addition_inv.simps) -apply(insert nth_append[of lm "0 # 0 # suf_lm" "n"], simp) +lemma abc_comp_commute: "(A [+] B) [+] C = A [+] (B [+] C)" +apply(auto simp: abc_comp.simps abc_shift.simps) +apply(case_tac x, auto) done -lemma [simp]: "0\2 = [0, 0::nat]" -apply(auto simp:numeral_2_eq_2) + + +lemma compile_z_correct: + "\rec_ci z = (ap, arity, fp); rec_exec z [n] = r\ \ + {\nl. nl = n # 0 \ (fp - arity) @ anything} ap {\nl. nl = n # r # 0 \ (fp - Suc arity) @ anything}" +apply(rule_tac abc_Hoare_haltI) +apply(rule_tac x = 1 in exI) +apply(auto simp: abc_steps_l.simps rec_ci.simps rec_ci_z_def + numeral_2_eq_2 abc_fetch.simps abc_step_l.simps rec_exec.simps) +done + +lemma compile_s_correct: + "\rec_ci s = (ap, arity, fp); rec_exec s [n] = r\ \ + {\nl. nl = n # 0 \ (fp - arity) @ anything} ap {\nl. nl = n # r # 0 \ (fp - Suc arity) @ anything}" +apply(auto simp: rec_ci.simps rec_ci_s_def compile_s_correct' rec_exec.simps) done -lemma id_case: - "\rec_ci (id m n) = (aprog, rs_pos, a_md); - rec_calc_rel (id m n) lm rs\ - \ \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (length aprog, lm @ [rs] @ 0\(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: - "\P []; \a list. P list \ P (list @ [a::'a])\ \ - P ((list::'a list))" -apply(case_tac "length list", simp) +lemma compile_id_correct': + assumes "n < length args" + shows "{\nl. nl = args @ 0 \ 2 @ anything} addition n (length args) (Suc (length args)) + {\nl. nl = args @ rec_exec (recf.id (length args) n) args # 0 # anything}" proof - - fix nat - assume ind: "\a list. P list \ P (list @ [a])" - and h: "length list = Suc nat" "P []" - from h show "P list" - proof(induct nat arbitrary: list, case_tac lista, simp, simp) - fix lista a listaa - from h show "P [a]" - by(insert ind[of "[]"], simp add: h) - next - fix nat list - assume nind: "\list. \length list = Suc nat; P []\ \ P list" - and g: "length (list:: 'a list) = Suc (Suc nat)" - from g show "P (list::'a list)" - apply(insert nind[of "butlast list"], simp add: h) - apply(insert ind[of "butlast list" "last list"], simp) - apply(subgoal_tac "butlast list @ [last list] = list", simp) - apply(case_tac "list::'a list", simp, simp) - done - qed -qed - -lemma nth_eq_butlast_nth: "\length ys > Suc k\ \ - ys ! k = butlast ys ! k" -apply(subgoal_tac "\ xs y. ys = xs @ [y]", auto simp: nth_append) -apply(rule_tac x = "butlast ys" in exI, rule_tac x = "last ys" in exI) -apply(case_tac "ys = []", simp, simp) -done + have "{\nl. nl = args @ 0 \ 2 @ anything} addition n (length args) (Suc (length args)) + {\nl. addition_inv (7, nl) n (length args) (Suc (length args)) (args @ 0 \ 2 @ anything)}" + using assms + by(rule_tac addition_correct, auto simp: numeral_2_eq_2 nth_append) + thus "?thesis" + using assms + by(simp add: addition_inv.simps rec_exec.simps + nth_append numeral_2_eq_2 list_update_append) +qed -lemma [simp]: -"\\k - \ \kn < m; length xs = m; rec_ci (recf.id m n) = (ap, arity, fp); rec_exec (recf.id m n) xs = r\ + \ {\nl. nl = xs @ 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ r # 0 \ (fp - Suc arity) @ anything}" +apply(auto simp: rec_ci.simps rec_ci_id.simps compile_id_correct') 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) +apply(induct gs arbitrary: pstr, simp add: cn_merge_gs.simps, auto) +apply(simp add: abc_comp_commute) done -lemma cn_merge_gs_length: - "length (cn_merge_gs (map rec_ci list) pstr) = - (\(ap, pos, n)\map rec_ci list. length ap) + 3 * length list " -apply(induct list arbitrary: pstr, simp, simp) +lemma footprint_ge: + "rec_ci a = (p, arity, fp) \ arity < fp" +apply(induct a) +apply(auto simp: rec_ci.simps) apply(case_tac "rec_ci a", simp) +apply(case_tac "rec_ci a1", case_tac "rec_ci a2", auto) +apply(case_tac "rec_ci a", auto) +done + +lemma param_pattern: + "\terminate f xs; rec_ci f = (p, arity, fp)\ \ length xs = arity" +apply(induct arbitrary: p arity fp rule: terminate.induct) +apply(auto simp: rec_ci.simps) +apply(case_tac "rec_ci f", simp) +apply(case_tac "rec_ci f", case_tac "rec_ci g", simp) +apply(case_tac "rec_ci f", case_tac "rec_ci gs", simp) done -lemma [simp]: "Suc n \ pstr \ pstr + x - n > 0" -by arith +lemma replicate_merge_anywhere: + "x\a @ x\b @ ys = x\(a+b) @ ys" +by(simp add:replicate_add) + +fun mv_box_inv :: "nat \ nat list \ nat \ nat \ nat list \ bool" + where + "mv_box_inv (as, lm) m n initlm = + (let plus = initlm ! m + initlm ! n in + length initlm > max m n \ m \ n \ + (if as = 0 then \ k l. lm = initlm[m := k, n := l] \ + k + l = plus \ k \ initlm ! m + else if as = 1 then \ k l. lm = initlm[m := k, n := l] + \ k + l + 1 = plus \ k < initlm ! m + else if as = 2 then \ k l. lm = initlm[m := k, n := l] + \ k + l = plus \ k \ initlm ! m + else if as = 3 then lm = initlm[m := 0, n := plus] + else False))" + +fun mv_box_stage1 :: "nat \ nat list \ nat \ nat" + where + "mv_box_stage1 (as, lm) m = + (if as = 3 then 0 + else 1)" + +fun mv_box_stage2 :: "nat \ nat list \ nat \ nat" + where + "mv_box_stage2 (as, lm) m = (lm ! m)" + +fun mv_box_stage3 :: "nat \ nat list \ nat \ 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 \ nat list) \ nat) \ (nat \ nat \ 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 \ nat) \ nat \ nat) set" + where + "lex_pair = less_than <*lex*> less_than" + +definition lex_triple :: + "((nat \ (nat \ nat)) \ (nat \ (nat \ nat))) set" + where + "lex_triple \ less_than <*lex*> lex_pair" -lemma [simp]: - "\Suc (pstr + length list) \ a_md; - length ys = Suc (length list); - length lm = n; - Suc n \ pstr\ - \ (ys ! length list # 0\(pstr - Suc n) @ butlast ys @ - 0\(a_md - (pstr + length list)) @ suf_lm) ! - (pstr + length list - n) = (0 :: nat)" -apply(insert nth_append[of "ys ! length list # 0\(pstr - Suc n) @ - butlast ys" "0\(a_md - (pstr + length list)) @ suf_lm" - "(pstr + length list - n)"], simp add: nth_append) +definition mv_box_LE :: + "(((nat \ nat list) \ nat) \ ((nat \ nat list) \ nat)) set" + where + "mv_box_LE \ (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: +"\m < length initlm; n < length initlm; m \ n\ \ + 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 (mv_box m n) = Some (Dec m 3)" +apply(simp add: mv_box.simps abc_fetch.simps) +done + +lemma [simp]: "abc_fetch (Suc 0) (mv_box m n) = + Some (Inc n)" +apply(simp add: mv_box.simps abc_fetch.simps) +done + +lemma [simp]: "abc_fetch 2 (mv_box m n) = Some (Goto 0)" +apply(simp add: mv_box.simps abc_fetch.simps) +done + +lemma [simp]: "abc_fetch 3 (mv_box m n) = None" +apply(simp add: mv_box.simps abc_fetch.simps) +done + +lemma replicate_Suc_iff_anywhere: "x # x\b @ ys = x\(Suc b) @ ys" +by simp + +lemma [simp]: + "\m \ n; m < length initlm; n < length initlm; + k + l = initlm ! m + initlm ! n; k \ initlm ! m; 0 < k\ + \ \ka la. initlm[m := k, n := l, m := k - Suc 0] = + initlm[m := ka, n := la] \ + Suc (ka + la) = initlm ! m + initlm ! n \ + ka < initlm ! m" +apply(rule_tac x = "k - Suc 0" in exI, rule_tac x = l in exI, + simp, auto) +apply(subgoal_tac + "initlm[m := k, n := l, m := k - Suc 0] = + initlm[n := l, m := k, m := k - Suc 0]") +apply(simp add: list_update_overwrite ) +apply(simp add: list_update_swap) +apply(simp add: list_update_swap) done lemma [simp]: - "\Suc (pstr + length list) \ a_md; - length ys = Suc (length list); - length lm = n; - Suc n \ pstr\ - \ (lm @ last ys # 0\(pstr - Suc n) @ butlast ys @ - 0\(a_md - (pstr + length list)) @ suf_lm)[pstr + length list := - last ys, n := 0] = - lm @ (0::nat)\(pstr - n) @ ys @ 0\(a_md - Suc (pstr + length list)) @ suf_lm" -apply(insert list_update_length[of - "lm @ last ys # 0\(pstr - Suc n) @ butlast ys" 0 - "0\(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\(pstr - Suc n) @ butlast ys @ - last ys # 0\(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) + "\m \ n; m < length initlm; n < length initlm; + Suc (k + l) = initlm ! m + initlm ! n; + k < initlm ! m\ + \ \ka la. initlm[m := k, n := l, n := Suc l] = + initlm[m := ka, n := la] \ + ka + la = initlm ! m + initlm ! n \ + ka \ initlm ! m" +apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto) +done + +lemma [simp]: + "\length initlm > max m n; m \ n\ \ + \na. \ (\(as, lm) m. as = 3) + (abc_steps_l (0, initlm) (mv_box m n) na) m \ + mv_box_inv (abc_steps_l (0, initlm) + (mv_box m n) na) m n initlm \ + mv_box_inv (abc_steps_l (0, initlm) + (mv_box m n) (Suc na)) m n initlm \ + ((abc_steps_l (0, initlm) (mv_box m n) (Suc na), m), + abc_steps_l (0, initlm) (mv_box m n) na, m) \ mv_box_LE" +apply(rule allI, rule impI, simp add: abc_step_red2) +apply(case_tac "(abc_steps_l (0, initlm) (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: + "\length initlm > max m n; m \ n\ \ + \ stp. (\ (as, lm). as = 3 \ + mv_box_inv (as, lm) m n initlm) + (abc_steps_l (0::nat, initlm) (mv_box m n) stp)" +apply(insert halt_lemma2[of mv_box_LE + "\ ((as, lm), m). mv_box_inv (as, lm) m n initlm" + "\ stp. (abc_steps_l (0, initlm) (mv_box m n) stp, m)" + "\ ((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) (mv_box m n) na)", + simp, auto) +done + +lemma mv_box_halt_cond: + "\m \ n; mv_box_inv (a, b) m n lm; a = 3\ \ + 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_correct': + "\length lm > max m n; m \ n\ \ + \ 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) (mv_box m n) stp", + simp) +apply(erule_tac mv_box_halt_cond, auto) +done + +lemma length_mvbox[simp]: "length (mv_box m n) = 3" +by(simp add: mv_box.simps) + +lemma mv_box_correct: + "\length lm > max m n; m\n\ + \ {\ nl. nl = lm} mv_box m n {\ nl. nl = lm[n := (lm ! m + lm ! n), m:=0]}" +apply(drule_tac mv_box_correct', simp) +apply(auto simp: abc_Hoare_halt_def) +apply(rule_tac x = stp in exI, simp) +done + +declare list_update.simps(2)[simp del] + +lemma [simp]: + "\length xs < gf; gf \ ft; n < length gs\ + \ (rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) + [ft + n - length xs := rec_exec (gs ! n) xs, 0 := 0] = + 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything" +using list_update_append[of "rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs)" + "0 \ (length gs - n) @ 0 # 0 \ length xs @ anything" "ft + n - length xs" "rec_exec (gs ! n) xs"] +apply(auto) +apply(case_tac "length gs - n", simp, simp add: list_update.simps replicate_Suc_iff_anywhere Suc_diff_Suc del: replicate_Suc) +apply(simp add: list_update.simps) done -lemma cn_merge_gs_ex: - "\\x aprog a_md rs_pos rs suf_lm lm. - \x \ set gs; rec_ci x = (aprog, rs_pos, a_md); - rec_calc_rel x lm rs\ - \ \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp - = (length aprog, lm @ [rs] @ 0\(a_md - rs_pos - 1) @ suf_lm); - pstr + length gs\ a_md; - \k Max (set (Suc n # map (\(aprog, p, n). n) (map rec_ci gs)))\ - \ \ stp. abc_steps_l (0, lm @ 0\(a_md - n) @ suf_lm) - (cn_merge_gs (map rec_ci gs) pstr) stp - = (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) gs) + - 3 * length gs, lm @ 0\(pstr - n) @ ys @ 0\(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: "\x aprog a_md rs_pos rs suf_lm lm. - \x = a \ x \ set list; rec_ci x = (aprog, rs_pos, a_md); - rec_calc_rel x lm rs\ - \ \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (length aprog, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm)" - and ind2: - "\ys. \\x aprog a_md rs_pos rs suf_lm lm. - \x \ set list; rec_ci x = (aprog, rs_pos, a_md); - rec_calc_rel x lm rs\ - \ \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp - = (length aprog, lm @ rs # 0\(a_md - Suc rs_pos) @ suf_lm); - \k - \ \stp. abc_steps_l (0, lm @ 0\(a_md - n) @ suf_lm) - (cn_merge_gs (map rec_ci list) pstr) stp = - (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) list) + - 3 * length list, - lm @ 0\(pstr - n) @ ys @ 0\(a_md - (pstr + length list)) @ suf_lm)" - and h: "Suc (pstr + length list) \ a_md" - "\k pstr \ (\(aprog, p, n). n) (rec_ci a) \ pstr \ - (\a\set list. (\(aprog, p, n). n) (rec_ci a) \ pstr)" - from h have k1: - "\stp. abc_steps_l (0, lm @ 0\(a_md - n) @ suf_lm) - (cn_merge_gs (map rec_ci list) pstr) stp = - (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) list) + - 3 * length list, lm @ 0\(pstr - n) @ butlast ys @ - 0\(a_md - (pstr + length list)) @ suf_lm) " - apply(rule_tac ind2) - apply(rule_tac ind, auto) +lemma compile_cn_gs_correct': + assumes + g_cond: "\g\set (take n gs). terminate g xs \ + (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc}))" + and ft: "ft = max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" + shows + "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} + cn_merge_gs (map rec_ci (take n gs)) ft + {\nl. nl = xs @ 0 \ (ft - length xs) @ + map (\i. rec_exec i xs) (take n gs) @ 0\(length gs - n) @ 0 \ Suc (length xs) @ anything}" + using g_cond +proof(induct n) + case 0 + have "ft > length xs" + using ft + by simp + thus "?case" + apply(rule_tac abc_Hoare_haltI) + apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps replicate_add[THEN sym] + replicate_Suc[THEN sym] del: replicate_Suc) done - from k1 and h show - "\stp. abc_steps_l (0, lm @ 0\(a_md - n) @ suf_lm) - (cn_merge_gs (map rec_ci list @ [rec_ci a]) pstr) stp = - (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) list) + - (\(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list), - lm @ 0\(pstr - n) @ ys @ 0\(a_md - Suc (pstr + length list)) @ suf_lm)" - apply(simp add: cn_merge_gs_tl_app) - apply(rule_tac as = - "(\(ap, pos, n)\map rec_ci list. length ap) + 3 * length list" - and bm = "lm @ 0\(pstr - n) @ butlast ys @ - 0\(a_md - (pstr + length list)) @ suf_lm" - and bs = "(\(ap, pos, n). length ap) (rec_ci a) + 3" - and bm' = "lm @ 0\(pstr - n) @ ys @ 0\(a_md - Suc (pstr + length list)) @ - suf_lm" in abc_append_exc2, simp) - apply(simp add: cn_merge_gs_length) - proof - - from h show - "\bstp. abc_steps_l (0, lm @ 0\(pstr - n) @ butlast ys @ - 0\(a_md - (pstr + length list)) @ suf_lm) - ((\(gprog, gpara, gn). gprog [+] mv_box gpara - (pstr + length list)) (rec_ci a)) bstp = - ((\(ap, pos, n). length ap) (rec_ci a) + 3, - lm @ 0\(pstr - n) @ ys @ 0\(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\(pstr - Suc n) @ butlast ys @ 0\(a_md - (pstr + length list)) @ suf_lm" - and bs = "3" and bm' = "lm @ 0\(pstr - n) @ ys @ - 0\(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 - "\astp. abc_steps_l (0, lm @ 0\(pstr - n) @ butlast ys @ - 0\(a_md - (pstr + length list)) @ suf_lm) aa astp = - (length aa, lm @ ys ! length list # 0\(pstr - Suc n) @ - butlast ys @ 0\(a_md - (pstr + length list)) @ suf_lm)" - apply(subgoal_tac "c \ Suc n") - apply(insert ind[of a aa b c lm "ys ! length list" - "0\(pstr - c) @ butlast ys @ 0\(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 - "\bstp. abc_steps_l (0, lm @ ys ! length list # - 0\(pstr - Suc n) @ butlast ys @ 0\(a_md - (pstr + length list)) @ suf_lm) - (mv_box b (pstr + length list)) bstp = - (3, lm @ 0\(pstr - n) @ ys @ 0\(a_md - Suc (pstr + length list)) @ suf_lm)" - apply(insert mv_box_ex [of b "pstr + length list" - "lm @ ys ! length list # 0\(pstr - Suc n) @ butlast ys @ - 0\(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 (mv_box b (pstr + length list))" +next + case (Suc n) + have ind': "\g\set (take n gs). + terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ + (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc})) \ + {\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft + {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything}" + by fact + have g_newcond: "\g\set (take (Suc n) gs). + terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc}))" + by fact + from g_newcond have ind: + "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft + {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything}" + apply(rule_tac ind', rule_tac ballI, erule_tac x = g in ballE, simp_all add: take_Suc) + by(case_tac "n < length gs", simp add:take_Suc_conv_app_nth, simp) + show "?case" + proof(cases "n < length gs") + case True + have h: "n < length gs" by fact + thus "?thesis" + proof(simp add: take_Suc_conv_app_nth cn_merge_gs_tl_app) + obtain gp ga gf where a: "rec_ci (gs!n) = (gp, ga, gf)" + by (metis prod_cases3) + moreover have "min (length gs) n = n" + using h by simp + moreover have + "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} + cn_merge_gs (map rec_ci (take n gs)) ft [+] (gp [+] mv_box ga (ft + n)) + {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ + rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything}" + proof(rule_tac abc_Hoare_plus_halt) + show "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft + {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything}" + using ind by simp + next + have x: "gs!n \ set (take (Suc n) gs)" + using h + by(simp add: take_Suc_conv_app_nth) + have b: "terminate (gs!n) xs" + using a g_newcond h x + by(erule_tac x = "gs!n" in ballE, simp_all) + hence c: "length xs = ga" + using a param_pattern by metis + have d: "gf > ga" using footprint_ge a by simp + have e: "ft \ gf" using ft a h + by(simp, rule_tac min_max.le_supI2, rule_tac Max_ge, simp, + rule_tac insertI2, + rule_tac f = "(\(aprog, p, n). n)" and x = "rec_ci (gs!n)" in image_eqI, simp, + rule_tac x = "gs!n" in image_eqI, simp, simp) + show "{\nl. nl = xs @ 0 \ (ft - length xs) @ + map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything} gp [+] mv_box ga (ft + n) + {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) + (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything}" + proof(rule_tac abc_Hoare_plus_halt) + show "{\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything} gp + {\nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) + (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything}" + proof - + have + "({\nl. nl = xs @ 0 \ (gf - ga) @ 0\(ft - gf)@map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything} + gp {\nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \ (gf - Suc ga) @ + 0\(ft - gf)@map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything})" + using a g_newcond h x + apply(erule_tac x = "gs!n" in ballE) + apply(simp, simp) + done + thus "?thesis" + using a b c d e + by(simp add: replicate_merge_anywhere) + qed + next + show + "{\nl. nl = xs @ rec_exec (gs ! n) xs # + 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything} + mv_box ga (ft + n) + {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ + rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything}" + proof - + have "{\nl. nl = xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ + map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything} + mv_box ga (ft + n) {\nl. nl = (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ + map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) + [ft + n := (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ + 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! ga + + (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ + map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! + (ft + n), ga := 0]}" + using a c d e h + apply(rule_tac mv_box_correct) + apply(simp, arith, arith) + done + moreover have "(xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ + map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) + [ft + n := (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ + 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! ga + + (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ + map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! + (ft + n), ga := 0]= + xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything" + using a c d e h + by(simp add: list_update_append nth_append length_replicate split: if_splits del: list_update.simps(2), auto) + ultimately show "?thesis" + by(simp) + qed + qed + qed + ultimately show + "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} + cn_merge_gs (map rec_ci (take n gs)) ft [+] (case rec_ci (gs ! n) of (gprog, gpara, gn) \ + gprog [+] mv_box gpara (ft + min (length gs) n)) + {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything}" 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) \ []" - by(simp add: mv_box.simps) qed next - show "(\(ap, pos, n). length ap) (rec_ci a) + 3 = - length ((\(gprog, gpara, gn). gprog [+] - mv_box gpara (pstr + length list)) (rec_ci a))" - by(case_tac "rec_ci a", simp) - next - show "listsum (map ((\(ap, pos, n). length ap) \ rec_ci) list) + - (\(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list)= - (\(ap, pos, n)\map rec_ci list. length ap) + 3 * length list + - ((\(ap, pos, n). length ap) (rec_ci a) + 3)" by simp - next - show "(\(gprog, gpara, gn). gprog [+] - mv_box gpara (pstr + length list)) (rec_ci a) \ []" - by(case_tac "rec_ci a", - simp add: abc_append.simps abc_shift.simps) + case False + have h: "\ n < length gs" by fact + hence ind': + "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} cn_merge_gs (map rec_ci gs) ft + {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) gs @ 0 \ Suc (length xs) @ anything}" + using ind + by simp + thus "?thesis" + using h + by(simp) qed qed - -lemma [simp]: "length (mv_boxes aa ba n) = 3*n" + +lemma compile_cn_gs_correct: + assumes + g_cond: "\g\set gs. terminate g xs \ + (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc}))" + and ft: "ft = max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" + shows + "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} + cn_merge_gs (map rec_ci gs) ft + {\nl. nl = xs @ 0 \ (ft - length xs) @ + map (\i. rec_exec i xs) gs @ 0 \ Suc (length xs) @ anything}" +using assms +using compile_cn_gs_correct'[of "length gs" gs xs ft ffp anything ] +apply(auto) +done + +lemma length_mvboxes[simp]: "length (mv_boxes aa ba n) = 3*n" by(induct n, auto simp: mv_boxes.simps) lemma exp_suc: "a\Suc b = a\b @ [a]" @@ -2919,1527 +823,1391 @@ using list_update_append[of "lm1 @ 0\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(simp add: list_update_append list_update.simps replicate_Suc_iff_anywhere exp_suc + del: replicate_Suc) apply(case_tac lm2, simp, simp) done -lemma mv_boxes_ex: - "\n \ ba - aa; ba > aa; length lm1 = aa; - length (lm2::nat list) = n; length lm3 = ba - aa - n\ - \ \ stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\n @ lm4) - (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\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\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: - "\lm2 lm3 lm4. \length lm2 = n; length lm3 = ba - (aa + n)\ \ - \stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\n @ lm4) - (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\n @ lm3 @ lm2 @ lm4)" - and h: "Suc n \ ba - aa" "aa < ba" "length (lm1::nat list) = aa" - "length (lm2::nat list) = Suc n" - "length (lm3::nat list) = ba - Suc (aa + n)" - from h show - "\astp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\n @ 0 # lm4) - (mv_boxes aa ba n) astp = - (3 * n, lm1 @ 0\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\n @ - 0 # lm4 = lm1 @ lm2 @ lm3 @ 0\n @ 0 # lm4", simp, simp) - apply(case_tac "lm2 = []", simp, simp) - done -next - fix n lm2 lm3 lm4 - assume h: "Suc n \ ba - aa" - "aa < ba" - "length (lm1::nat list) = aa" - "length (lm2::nat list) = Suc n" - "length (lm3::nat list) = ba - Suc (aa + n)" - thus " \bstp. abc_steps_l (0, lm1 @ 0\n @ last lm2 # lm3 @ - butlast lm2 @ 0 # lm4) - (mv_box (aa + n) (ba + n)) bstp - = (3, lm1 @ 0 # 0\n @ lm3 @ lm2 @ lm4)" - apply(insert mv_box_ex[of "aa + n" "ba + n" - "lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4"], simp) - done -qed - -lemma [simp]: "\Suc n \ aa - ba; ba < aa; length lm1 = ba; - length lm2 = aa - Suc (ba + n); length lm3 = Suc n\ - \ (lm1 @ butlast lm3 @ 0 # lm2 @ 0\n @ last lm3 # lm4) ! (aa + n) = last lm3" -using nth_append[of "lm1 @ butlast lm3 @ 0 # lm2 @ 0\n" "last lm3 # lm4" "aa + n"] -apply(simp) +lemma [simp]: + "\Suc (length lm1 + n) \ ba; length lm2 = Suc n; length lm3 = ba - Suc (length lm1 + n); + \ ba - Suc (length lm1) < ba - Suc (length lm1 + n); \ ba + n - length lm1 < n\ + \ (0::nat) \ n @ (last lm2 # lm3 @ butlast lm2 @ 0 # lm4)[ba - length lm1 := last lm2, 0 := 0] = + 0 # 0 \ n @ lm3 @ lm2 @ lm4" +apply(subgoal_tac "ba - length lm1 = Suc n + length lm3", simp add: list_update.simps list_update_append) +apply(simp add: replicate_Suc_iff_anywhere exp_suc del: replicate_Suc) +apply(case_tac lm2, simp, simp) +apply(auto) done -lemma [simp]: "\Suc n \ aa - ba; ba < aa; length lm1 = ba; - length lm2 = aa - Suc (ba + n); length lm3 = Suc n\ - \ (lm1 @ butlast lm3 @ 0 # lm2 @ 0\n @ last lm3 # lm4) ! (ba + n) = 0" -apply(simp add: nth_append) -done +lemma mv_boxes_correct: + "\aa + n \ ba; ba > aa; length lm1 = aa; length lm2 = n; length lm3 = ba - aa - n\ + \ {\ nl. nl = lm1 @ lm2 @ lm3 @ 0\n @ lm4} (mv_boxes aa ba n) + {\ nl. nl = lm1 @ 0\n @ lm3 @ lm2 @ lm4}" +proof(induct n arbitrary: lm2 lm3 lm4) + case 0 + thus "?case" + by(simp add: mv_boxes.simps abc_Hoare_halt_def, rule_tac x = 0 in exI, simp add: abc_steps_l.simps) +next + case (Suc n) + have ind: + "\lm2 lm3 lm4. + \aa + n \ ba; aa < ba; length lm1 = aa; length lm2 = n; length lm3 = ba - aa - n\ + \ {\nl. nl = lm1 @ lm2 @ lm3 @ 0 \ n @ lm4} mv_boxes aa ba n {\nl. nl = lm1 @ 0 \ n @ lm3 @ lm2 @ lm4}" + by fact + have h1: "aa + Suc n \ ba" by fact + have h2: "aa < ba" by fact + have h3: "length lm1 = aa" by fact + have h4: "length lm2 = Suc n" by fact + have h5: "length lm3 = ba - aa - Suc n" by fact + have "{\nl. nl = lm1 @ lm2 @ lm3 @ 0 \ Suc n @ lm4} mv_boxes aa ba n [+] mv_box (aa + n) (ba + n) + {\nl. nl = lm1 @ 0 \ Suc n @ lm3 @ lm2 @ lm4}" + proof(rule_tac abc_Hoare_plus_halt) + have "{\nl. nl = lm1 @ butlast lm2 @ (last lm2 # lm3) @ 0 \ n @ (0 # lm4)} mv_boxes aa ba n + {\ nl. nl = lm1 @ 0\n @ (last lm2 # lm3) @ butlast lm2 @ (0 # lm4)}" + using h1 h2 h3 h4 h5 + by(rule_tac ind, simp_all) + moreover have " lm1 @ butlast lm2 @ (last lm2 # lm3) @ 0 \ n @ (0 # lm4) + = lm1 @ lm2 @ lm3 @ 0 \ Suc n @ lm4" + using h4 + by(simp add: replicate_Suc[THEN sym] exp_suc del: replicate_Suc, + case_tac lm2, simp_all) + ultimately show "{\nl. nl = lm1 @ lm2 @ lm3 @ 0 \ Suc n @ lm4} mv_boxes aa ba n + {\ nl. nl = lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4}" + by (metis append_Cons) + next + let ?lm = "lm1 @ 0 \ n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4" + have "{\nl. nl = ?lm} mv_box (aa + n) (ba + n) + {\ nl. nl = ?lm[(ba + n) := ?lm!(aa+n) + ?lm!(ba+n), (aa+n):=0]}" + using h1 h2 h3 h4 h5 + by(rule_tac mv_box_correct, simp_all) + moreover have "?lm[(ba + n) := ?lm!(aa+n) + ?lm!(ba+n), (aa+n):=0] + = lm1 @ 0 \ Suc n @ lm3 @ lm2 @ lm4" + using h1 h2 h3 h4 h5 + by(auto simp: nth_append list_update_append split: if_splits) + ultimately show "{\nl. nl = lm1 @ 0 \ n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4} mv_box (aa + n) (ba + n) + {\nl. nl = lm1 @ 0 \ Suc n @ lm3 @ lm2 @ lm4}" + by simp + qed + thus "?case" + by(simp add: mv_boxes.simps) +qed + +lemma [simp]: + "\Suc n \ aa - length lm1; length lm1 < aa; + length lm2 = aa - Suc (length lm1 + n); + length lm3 = Suc n; + \ aa - Suc (length lm1) < aa - Suc (length lm1 + n); + \ aa + n - length lm1 < n\ + \ butlast lm3 @ ((0::nat) # lm2 @ 0 \ n @ last lm3 # lm4)[0 := last lm3, aa - length lm1 := 0] = lm3 @ lm2 @ 0 # 0 \ n @ lm4" + apply(subgoal_tac "aa - length lm1 = length lm2 + Suc n") + apply(simp add: list_update.simps list_update_append) + apply(simp add: replicate_Suc[THEN sym] exp_suc del: replicate_Suc) + apply(case_tac lm3, simp, simp) + apply(auto) + done -lemma [simp]: "\Suc n \ aa - ba; ba < aa; length lm1 = ba; - length lm2 = aa - Suc (ba + n); length lm3 = Suc n\ - \ (lm1 @ butlast lm3 @ 0 # lm2 @ 0\n @ last lm3 # lm4)[ba + n := last lm3, aa + n := 0] - = lm1 @ lm3 @ lm2 @ 0 # 0\n @ lm4" -using list_update_append[of "lm1 @ butlast lm3" "(0\'a) # lm2 @ (0\'a)\n @ last lm3 # lm4"] -apply(simp) -using list_update_append[of "lm1 @ butlast lm3 @ last lm3 # lm2 @ (0\'a)\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: +lemma mv_boxes_correct2: "\n \ aa - ba; ba < aa; length (lm1::nat list) = ba; length (lm2::nat list) = aa - ba - n; length (lm3::nat list) = n\ - \ \ stp. abc_steps_l (0, lm1 @ 0\n @ lm2 @ lm3 @ lm4) - (mv_boxes aa ba n) stp = - (3 * n, lm1 @ lm3 @ lm2 @ 0\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\n @ last lm3 # lm4" in abc_append_exc2, simp_all) -apply(simp only: exponent_cons_iff, simp only: exp_suc, simp) + \{\ nl. nl = lm1 @ 0\n @ lm2 @ lm3 @ lm4} + (mv_boxes aa ba n) + {\ nl. nl = lm1 @ lm3 @ lm2 @ 0\n @ lm4}" +proof(induct n arbitrary: lm2 lm3 lm4) + case 0 + thus "?case" + by(simp add: mv_boxes.simps abc_Hoare_halt_def, rule_tac x = 0 in exI, simp add: abc_steps_l.simps) +next + case (Suc n) + have ind: + "\lm2 lm3 lm4. + \n \ aa - ba; ba < aa; length lm1 = ba; length lm2 = aa - ba - n; length lm3 = n\ + \ {\nl. nl = lm1 @ 0 \ n @ lm2 @ lm3 @ lm4} mv_boxes aa ba n {\nl. nl = lm1 @ lm3 @ lm2 @ 0 \ n @ lm4}" + by fact + have h1: "Suc n \ aa - ba" by fact + have h2: "ba < aa" by fact + have h3: "length lm1 = ba" by fact + have h4: "length lm2 = aa - ba - Suc n" by fact + have h5: "length lm3 = Suc n" by fact + have "{\nl. nl = lm1 @ 0 \ Suc n @ lm2 @ lm3 @ lm4} mv_boxes aa ba n [+] mv_box (aa + n) (ba + n) + {\nl. nl = lm1 @ lm3 @ lm2 @ 0 \ Suc n @ lm4}" + proof(rule_tac abc_Hoare_plus_halt) + have "{\ nl. nl = lm1 @ 0 \ n @ (0 # lm2) @ (butlast lm3) @ (last lm3 # lm4)} mv_boxes aa ba n + {\ nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0\n @ (last lm3 # lm4)}" + using h1 h2 h3 h4 h5 + by(rule_tac ind, simp_all) + moreover have "lm1 @ 0 \ n @ (0 # lm2) @ (butlast lm3) @ (last lm3 # lm4) + = lm1 @ 0 \ Suc n @ lm2 @ lm3 @ lm4" + using h5 + by(simp add: replicate_Suc_iff_anywhere exp_suc + del: replicate_Suc, case_tac lm3, simp_all) + ultimately show "{\nl. nl = lm1 @ 0 \ Suc n @ lm2 @ lm3 @ lm4} mv_boxes aa ba n + {\ nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0\n @ (last lm3 # lm4)}" + by metis + next + thm mv_box_correct + let ?lm = "lm1 @ butlast lm3 @ (0 # lm2) @ 0 \ n @ last lm3 # lm4" + have "{\nl. nl = ?lm} mv_box (aa + n) (ba + n) + {\nl. nl = ?lm[ba+n := ?lm!(aa+n)+?lm!(ba+n), (aa+n):=0]}" + using h1 h2 h3 h4 h5 + by(rule_tac mv_box_correct, simp_all) + moreover have "?lm[ba+n := ?lm!(aa+n)+?lm!(ba+n), (aa+n):=0] + = lm1 @ lm3 @ lm2 @ 0 \ Suc n @ lm4" + using h1 h2 h3 h4 h5 + by(auto simp: nth_append list_update_append split: if_splits) + ultimately show "{\nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0 \ n @ last lm3 # lm4} mv_box (aa + n) (ba + n) + {\nl. nl = lm1 @ lm3 @ lm2 @ 0 \ Suc n @ lm4}" + by simp + qed + thus "?case" + by(simp add: mv_boxes.simps) +qed + +lemma save_paras: + "{\nl. nl = xs @ 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length xs) @ + map (\i. rec_exec i xs) gs @ 0 \ Suc (length xs) @ anything} + mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs) + {\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ map (\i. rec_exec i xs) gs @ 0 # xs @ anything}" proof - - fix n lm2 lm3 lm4 - assume ind: -"\lm2 lm3 lm4. \length lm2 = aa - (ba + n); length lm3 = n\ \ - \stp. abc_steps_l (0, lm1 @ 0\n @ lm2 @ lm3 @ lm4) - (mv_boxes aa ba n) stp = - (3 * n, lm1 @ lm3 @ lm2 @ 0\n @ lm4)" - and h: "Suc n \ aa - ba" - "ba < aa" - "length (lm1::nat list) = ba" - "length (lm2::nat list) = aa - Suc (ba + n)" - "length (lm3::nat list) = Suc n" - from h show - "\astp. abc_steps_l (0, lm1 @ 0\n @ 0 # lm2 @ lm3 @ lm4) - (mv_boxes aa ba n) astp = - (3 * n, lm1 @ butlast lm3 @ 0 # lm2 @ 0\n @ last lm3 # lm4)" - apply(insert ind[of "0 # lm2" "butlast lm3" "last lm3 # lm4"], - simp) - apply(subgoal_tac - "lm1 @ 0\n @ 0 # lm2 @ butlast lm3 @ last lm3 # lm4 = - lm1 @ 0\n @ 0 # lm2 @ lm3 @ lm4", simp, simp) - apply(case_tac "lm3 = []", simp, simp) - done -next - fix n lm2 lm3 lm4 - assume h: - "Suc n \ aa - ba" - "ba < aa" - "length lm1 = ba" - "length (lm2::nat list) = aa - Suc (ba + n)" - "length (lm3::nat list) = Suc n" - thus - "\bstp. abc_steps_l (0, lm1 @ butlast lm3 @ 0 # lm2 @ 0\n @ - last lm3 # lm4) - (mv_box (aa + n) (ba + n)) bstp = - (3, lm1 @ lm3 @ lm2 @ 0 # 0\n @ lm4)" - apply(insert mv_box_ex[of "aa + n" "ba + n" "lm1 @ butlast lm3 @ - 0 # lm2 @ 0\n @ last lm3 # lm4"], simp) - done + let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" + have "{\nl. nl = [] @ xs @ (0\(?ft - length xs) @ map (\i. rec_exec i xs) gs @ [0]) @ + 0 \ (length xs) @ anything} mv_boxes 0 (Suc ?ft + length gs) (length xs) + {\nl. nl = [] @ 0 \ (length xs) @ (0\(?ft - length xs) @ map (\i. rec_exec i xs) gs @ [0]) @ xs @ anything}" + by(rule_tac mv_boxes_correct, auto) + thus "?thesis" + by(simp add: replicate_merge_anywhere) qed -lemma cn_merge_gs_len: - "length (cn_merge_gs (map rec_ci gs) pstr) = - (\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs" -apply(induct gs arbitrary: pstr, simp, simp) -apply(case_tac "rec_ci a", simp) -done - -lemma [simp]: "n < pstr \ - Suc (pstr + length ys - n) = Suc (pstr + length ys) - n" -by arith - -lemma save_paras': - "\length lm = n; pstr > n; a_md > pstr + length ys + n\ - \ \stp. abc_steps_l (0, lm @ 0\(pstr - n) @ ys @ - 0\(a_md - pstr - length ys) @ suf_lm) - (mv_boxes 0 (pstr + Suc (length ys)) n) stp - = (3 * n, 0\pstr @ ys @ 0 # lm @ 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm)" -apply(insert mv_boxes_ex[of n "pstr + Suc (length ys)" 0 "[]" "lm" - "0\(pstr - n) @ ys @ [0]" "0\(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 (((\(aprog, p, n). n) o rec_ci) ` set gs)))) - = (Max (insert ba (((\(aprog, p, n). n) o rec_ci) ` set gs)))" -apply(rule min_max.sup_absorb2, auto) -done - -lemma [simp]: - "((\(aprog, p, n). n) ` rec_ci ` set gs) = - (((\(aprog, p, n). n) o rec_ci) ` set gs)" -apply(induct gs) -apply(simp, simp) -done - -lemma ci_cn_md_def: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_ci f = (a, aa, ba)\ - \ a_md = max (Suc n) (Max (insert ba (((\(aprog, p, n). n) o - rec_ci) ` set gs))) + Suc (length gs) + n" -apply(simp add: rec_ci.simps, auto) -done - -lemma save_paras_prog_ex: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_ci f = (a, aa, ba); - pstr = Max (set (Suc n # ba # map (\(aprog, p, n). n) - (map rec_ci (gs))))\ - \ \ap bp cp. - aprog = ap [+] bp [+] cp \ - length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + - 3 * length gs \ bp = mv_boxes 0 (pstr + Suc (length gs)) n" -apply(simp add: rec_ci.simps) -apply(rule_tac x = - "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) o rec_ci) ` set gs))))" in exI, - simp add: cn_merge_gs_len) -apply(rule_tac x = - "mv_boxes (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) - 0 (length gs) [+] a [+]mv_box aa (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] - empty_boxes (length gs) [+] mv_box (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] - mv_boxes (Suc (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) - ` set gs))) + length gs)) 0 n" in exI, auto) -apply(simp add: abc_append_commute) -done - -lemma save_paras: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rs_pos = n; - \k(aprog, p, n). n) - (map rec_ci (gs))))\ - \ \stp. abc_steps_l ((\(ap, pos, n)\map rec_ci gs. length ap) + - 3 * length gs, lm @ 0\(pstr - n) @ ys @ - 0\(a_md - pstr - length ys) @ suf_lm) aprog stp = - ((\(ap, pos, n)\map rec_ci gs. length ap) + - 3 * length gs + 3 * n, - 0\pstr @ ys @ 0 # lm @ 0\(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" - "\k(aprog, p, n). n) - (map rec_ci (gs))))" - from h and g have k1: - "\ ap bp cp. aprog = ap [+] bp [+] cp \ - length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + - 3 *length gs \ bp = mv_boxes 0 (pstr + Suc (length ys)) n" - apply(drule_tac save_paras_prog_ex, auto) - done - from h have k2: - "\ stp. abc_steps_l (0, lm @ 0\(pstr - n) @ ys @ - 0\(a_md - pstr - length ys) @ suf_lm) - (mv_boxes 0 (pstr + Suc (length ys)) n) stp = - (3 * n, 0\pstr @ ys @ 0 # lm @ 0\(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 - "\stp. abc_steps_l ((\(ap, pos, n)\map rec_ci gs. length ap) + - 3 * length gs, lm @ 0\(pstr - n) @ ys @ - 0\(a_md - pstr - length ys) @ suf_lm) aprog stp = - ((\(ap, pos, n)\map rec_ci gs. length ap) + - 3 * length gs + 3 * n, - 0\ pstr @ ys @ 0 # lm @ 0\(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 \ length ap = - (\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs - \ bp = mv_boxes 0 (pstr + Suc (length ys)) n" - from this and k2 show "?thesis" - apply(simp) - apply(rule_tac abc_append_exc1, simp, simp, simp) - done - qed -qed - -lemma ci_cn_para_eq: - "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md) \ rs_pos = n" -apply(simp add: rec_ci.simps, case_tac "rec_ci f", simp) -done - -lemma calc_gs_prog_ex: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_ci f = (a, aa, ba); - Max (set (Suc n # ba # map (\(aprog, p, n). n) - (map rec_ci (gs)))) = pstr\ - \ \ap bp. aprog = ap [+] bp \ - ap = cn_merge_gs (map rec_ci gs) pstr" -apply(simp add: rec_ci.simps) -apply(rule_tac x = "mv_boxes 0 (Suc (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) n [+] - mv_boxes (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs) [+] - a [+] mv_box aa (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] - empty_boxes (length gs) [+] mv_box (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] - mv_boxes (Suc (max (Suc n) (Max - (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) 0 n" - in exI) -apply(auto simp: abc_append_commute) -done - -lemma cn_calc_gs: - assumes ind: - "\x aprog a_md rs_pos rs suf_lm lm. - \x \ set gs; - rec_ci x = (aprog, rs_pos, a_md); - rec_calc_rel x lm rs\ - \ \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (length aprog, lm @ [rs] @ 0\(a_md - rs_pos - 1) @ suf_lm)" - and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" - "\k(aprog, p, n). n) - (map rec_ci (gs)))) = pstr" - shows - "\stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - ((\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs, - lm @ 0\(pstr - n) @ ys @ 0\(a_md -pstr - length ys) @ suf_lm) " -proof - - from h have k1: - "\ ap bp. aprog = ap [+] bp \ ap = - cn_merge_gs (map rec_ci gs) pstr" - by(erule_tac calc_gs_prog_ex, auto) - from h have j1: "rs_pos = n" - by(simp add: ci_cn_para_eq) - from h have j2: "a_md \ pstr" - by(drule_tac a = a and aa = aa and ba = ba in - ci_cn_md_def, simp, simp) - from h have j3: "pstr > n" - by(auto) - from j1 and j2 and j3 and h have k2: - "\stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) - (cn_merge_gs (map rec_ci gs) pstr) stp - = ((\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs, - lm @ 0\(pstr - n) @ ys @ 0\(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 - "\stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) - (cn_merge_gs (map rec_ci gs) pstr [+] bp) stp = - (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) gs) + - 3 * length gs, - lm @ 0\(pstr - n) @ ys @ 0\(a_md - (pstr + length ys)) @ suf_lm)" - apply(insert abc_append_exc1[of - "lm @ 0\(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\(pstr - n) @ ys @ 0\(a_md - pstr - length ys) @ suf_lm" 0 - "[]" bp], simp add: cn_merge_gs_len) - done - qed -qed - -lemma reset_new_paras': - "\length lm = n; - pstr > 0; - a_md \ pstr + length ys + n; - pstr > length ys\ \ - \stp. abc_steps_l (0, 0\pstr @ ys @ 0 # lm @ 0\(a_md - Suc (pstr + length ys + n)) @ - suf_lm) (mv_boxes pstr 0 (length ys)) stp = - (3 * length ys, ys @ 0\pstr @ 0 # lm @ 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm)" -apply(insert mv_boxes_ex2[of "length ys" "pstr" 0 "[]" - "0\(pstr - length ys)" "ys" - "0 # lm @ 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm"], - simp add: exponent_add_iff) +lemma [intro]: + "length gs \ ffp \ length gs \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" + apply(rule_tac min_max.le_supI2) + apply(simp add: Max_ge_iff) done -lemma [simp]: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_calc_rel f ys rs; rec_ci f = (a, aa, ba); - pstr = Max (set (Suc n # ba # map (\(aprog, p, n). n) - (map rec_ci (gs))))\ - \ length ys < pstr" -apply(subgoal_tac "length ys = aa", simp) -apply(subgoal_tac "aa < ba \ ba \ pstr", - rule basic_trans_rules(22), auto) -apply(rule min_max.le_supI2) -apply(auto) -apply(erule_tac para_pattern, simp) -done - -lemma reset_new_paras_prog_ex: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_ci f = (a, aa, ba); - Max (set (Suc n # ba # map (\(aprog, p, n). n) - (map rec_ci (gs)))) = pstr\ - \ \ ap bp cp. aprog = ap [+] bp [+] cp \ - length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + - 3 *length gs + 3 * n \ bp = mv_boxes pstr 0 (length gs)" -apply(simp add: rec_ci.simps) -apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] - mv_boxes 0 (Suc (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) n" in exI, - simp add: cn_merge_gs_len) -apply(rule_tac x = "a [+] - mv_box aa (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] - empty_boxes (length gs) [+] mv_box - (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n - [+] mv_boxes (Suc (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) 0 n" in exI, - auto simp: abc_append_commute) -done - -lemma reset_new_paras: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rs_pos = n; - \k(aprog, p, n). n) - (map rec_ci (gs))))\ -\ \stp. abc_steps_l ((\(ap, pos, n)\map rec_ci gs. length ap) + - 3 * length gs + 3 * n, - 0\pstr @ ys @ 0 # lm @ 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp = - ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + 3 * n, - ys @ 0\pstr @ 0 # lm @ 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm)" +lemma restore_new_paras: + "ffp \ length gs + \ {\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ map (\i. rec_exec i xs) gs @ 0 # xs @ anything} + mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs) + {\nl. nl = map (\i. rec_exec i xs) gs @ 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ 0 # xs @ anything}" proof - - assume h: - "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" - "rs_pos = n" - "length ys = aa" - "\k(aprog, p, n). n) - (map rec_ci (gs))))" - from h and g have k1: - "\ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = - (\(ap, pos, n)\map rec_ci gs. length ap) + - 3 *length gs + 3 * n \ bp = mv_boxes pstr 0 (length ys)" - by(drule_tac reset_new_paras_prog_ex, auto) - from h have k2: - "\ stp. abc_steps_l (0, 0\pstr @ ys @ 0 # lm @ 0\(a_md - Suc (pstr + length ys + n)) @ - suf_lm) (mv_boxes pstr 0 (length ys)) stp = - (3 * (length ys), - ys @ 0\pstr @ 0 # lm @ 0\(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 \ aa < ba \ ba \ pstr", arith, - simp add: para_pattern) - apply(insert g, auto intro: min_max.le_supI2) - done - from k1 show - "\stp. abc_steps_l ((\(ap, pos, n)\map rec_ci gs. length ap) + 3 - * length gs + 3 * n, 0\pstr @ ys @ 0 # lm @ 0\(a_md - Suc (pstr + length ys + n)) @ - suf_lm) aprog stp = - ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + - 3 * n, ys @ 0\pstr @ 0 # lm @ 0\(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 \ length ap = - (\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs + - 3 * n \ bp = mv_boxes pstr 0 (length ys)" - from this and k2 show "?thesis" - apply(simp) - apply(drule_tac as = - "(\(ap, pos, n)\map rec_ci gs. length ap) + 3 * length gs + - 3 * n" and ap = ap and cp = cp in abc_append_exc1, auto) - apply(rule_tac x = stp in exI, simp add: h) - using h - apply(simp) - done - qed + let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" + assume j: "ffp \ length gs" + hence "{\ nl. nl = [] @ 0\length gs @ 0\(?ft - length gs) @ map (\i. rec_exec i xs) gs @ ((0 # xs) @ anything)} + mv_boxes ?ft 0 (length gs) + {\ nl. nl = [] @ map (\i. rec_exec i xs) gs @ 0\(?ft - length gs) @ 0\length gs @ ((0 # xs) @ anything)}" + by(rule_tac mv_boxes_correct2, auto) + moreover have "?ft \ length gs" + using j + by(auto) + ultimately show "?thesis" + using j + by(simp add: replicate_merge_anywhere le_add_diff_inverse) qed - -lemma calc_f_prog_ex: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_ci f = (a, aa, ba); - Max (set (Suc n # ba # map (\(aprog, p, n). n) - (map rec_ci (gs)))) = pstr\ - \ \ ap bp cp. aprog = ap [+] bp [+] cp \ - length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + - 6 *length gs + 3 * n \ bp = a" -apply(simp add: rec_ci.simps) -apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] - mv_boxes 0 (Suc (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) n [+] - mv_boxes (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs)" in exI, - simp add: cn_merge_gs_len) -apply(rule_tac x = "mv_box aa (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] - empty_boxes (length gs) [+] mv_box (max (Suc n) ( - Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] - mv_boxes (Suc (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) 0 n" in exI, - auto simp: abc_append_commute) + +lemma [intro]: "ffp \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" +apply(rule_tac min_max.le_supI2) +apply(rule_tac Max_ge, auto) done -lemma calc_cn_f: - assumes ind: - "\x ap fp arity r anything args. - \x = map rec_ci gs; rec_ci f = (ap, arity, fp); rec_calc_rel f args r\ - \ \stp. abc_steps_l (0, args @ 0 \ (fp - arity) @ anything) ap stp = - (length ap, args @ [r] @ 0 \ (fp - arity - 1) @ anything)" - and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" - "rec_calc_rel (Cn n f gs) lm rs" - "length ys = length gs" - "rec_calc_rel f ys rs" - "length lm = n" - "rec_ci f = (a, aa, ba)" - and p: "pstr = Max (set (Suc n # ba # map (\(aprog, p, n). n) - (map rec_ci (gs))))" - shows "\stp. abc_steps_l - ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + 3 * n, - ys @ 0\pstr @ 0 # lm @ 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp = - ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + - 3 * n + length a, - ys @ rs # 0\pstr @ lm @ 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm)" -proof - - from h have k1: - "\ ap bp cp. aprog = ap [+] bp [+] cp \ - length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + - 6 *length gs + 3 * n \ bp = a" - by(drule_tac calc_f_prog_ex, auto) - from h and k1 show "?thesis" - proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE) - fix ap bp apa cp - assume - "aprog = ap [+] bp [+] cp \ - length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + - 6 * length gs + 3 * n \ bp = a" - from h and this show "?thesis" - apply(simp, rule_tac abc_append_exc1, simp_all) - thm ind - apply(insert ind[of "map rec_ci gs" a aa ba ys rs - "0\(pstr - ba + length gs) @ 0 # lm @ - 0\(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp) - apply(subgoal_tac "ba > aa \ aa = length gs\ pstr \ ba", simp) - apply(simp add: exponent_add_iff) - apply(case_tac pstr, simp add: p) - apply(simp only: exp_suc, simp) - apply(rule conjI, rule ci_ad_ge_paras, simp, rule conjI) - apply(subgoal_tac "length ys = aa", simp, - rule para_pattern, simp, simp) - apply(insert p, simp) - apply(auto intro: min_max.le_supI2) - done - qed -qed - -lemma [simp]: - "\pstr + length ys + n \ a_md; ys \ []\ \ - pstr < a_md + length suf_lm" -apply(case_tac "length ys", simp) -apply(arith) -done - -lemma [simp]: - "pstr > length ys - \ (ys @ rs # 0\pstr @ lm @ - 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm) ! pstr = (0::nat)" -apply(simp add: nth_append) -done - -(* -lemma [simp]: "\length ys < pstr; pstr - length ys = Suc x\ - \ pstr - Suc (length ys) = x" -by arith -*) - -lemma [simp]: "pstr > length ys \ - (ys @ rs # 0\pstr @ lm @ 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm) - [pstr := rs, length ys := 0] = - ys @ 0\(pstr - length ys) @ (rs::nat) # 0\length ys @ lm @ 0\(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\(pstr - Suc (length ys))" "0" "0\length ys @ lm @ - 0\(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 +declare max_less_iff_conj[simp del] -lemma save_rs': - "\pstr > length ys\ - \ \stp. abc_steps_l (0, ys @ rs # 0\pstr @ lm @ - 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm) - (mv_box (length ys) pstr) stp = - (3, ys @ 0\(pstr - (length ys)) @ rs # - 0\length ys @ lm @ 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm)" -apply(insert mv_box_ex[of "length ys" pstr - "ys @ rs # 0\pstr @ lm @ 0\(a_md - Suc(pstr + length ys + n)) @ suf_lm"], - simp) -done - - -lemma save_rs_prog_ex: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_ci f = (a, aa, ba); - Max (set (Suc n # ba # map (\(aprog, p, n). n) - (map rec_ci (gs)))) = pstr\ - \ \ ap bp cp. aprog = ap [+] bp [+] cp \ - length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + - 6 *length gs + 3 * n + length a - \ bp = 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 - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) - [+] mv_boxes 0 (Suc (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) n [+] - mv_boxes (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) - 0 (length gs) [+] a" - in exI, simp add: cn_merge_gs_len) -apply(rule_tac x = - "empty_boxes (length gs) [+] - mv_box (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] - mv_boxes (Suc (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) - + length gs)) 0 n" in exI, - auto simp: abc_append_commute) -done - -lemma save_rs: - assumes h: - "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" - "rec_calc_rel (Cn n f gs) lm rs" - "\k(aprog, p, n). n) - (map rec_ci (gs))))" - shows "\stp. abc_steps_l - ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs - + 3 * n + length a, ys @ rs # 0\pstr @ lm @ - 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp = - ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs - + 3 * n + length a + 3, - ys @ 0\(pstr - length ys) @ rs # 0\length ys @ lm @ - 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm)" +lemma save_rs: + "\far = length gs; + ffp \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))); + far < ffp\ +\ {\nl. nl = map (\i. rec_exec i xs) gs @ + rec_exec (Cn (length xs) f gs) xs # 0 \ max (Suc (length xs)) + (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ xs @ anything} + mv_box far (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) + {\nl. nl = map (\i. rec_exec i xs) gs @ + 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @ + rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything}" proof - - from h and pdef have k1: - "\ ap bp cp. aprog = ap [+] bp [+] cp \ - length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + - 6 *length gs + 3 * n + length a \ bp = 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 - "\stp. abc_steps_l - ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + 3 * n - + length a, ys @ rs # 0\pstr @ lm @ 0\(a_md - Suc (pstr + length ys + n)) - @ suf_lm) aprog stp = - ((\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + 3 * n - + length a + 3, ys @ 0\(pstr - length ys) @ rs # - 0\length ys @ lm @ 0\(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 \ length ap = - (\(ap, pos, n)\map rec_ci gs. length ap) + 6 * length gs + - 3 * n + length a \ bp = 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 \ pstr \ ba \ ba > aa", - arith) - apply(simp add: para_pattern, insert pdef, auto) - apply(rule_tac min_max.le_supI2, simp) - done - qed + let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" + thm mv_box_correct + let ?lm= " map (\i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ ?ft @ xs @ anything" + assume h: "far = length gs" "ffp \ ?ft" "far < ffp" + hence "{\ nl. nl = ?lm} mv_box far ?ft {\ nl. nl = ?lm[?ft := ?lm!far + ?lm!?ft, far := 0]}" + apply(rule_tac mv_box_correct) + by(case_tac "rec_ci a", auto) + moreover have "?lm[?ft := ?lm!far + ?lm!?ft, far := 0] + = map (\i. rec_exec i xs) gs @ + 0 \ (?ft - length gs) @ + rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" + using h + apply(simp add: nth_append) + using list_update_length[of "map (\i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # + 0 \ (?ft - Suc (length gs))" 0 "0 \ length gs @ xs @ anything" "rec_exec (Cn (length xs) f gs) xs"] + apply(simp add: replicate_merge_anywhere replicate_Suc_iff_anywhere del: replicate_Suc) + by(simp add: list_update_append list_update.simps replicate_Suc_iff_anywhere del: replicate_Suc) + ultimately show "?thesis" + by(simp) qed lemma [simp]: "length (empty_boxes n) = 2*n" apply(induct n, simp, simp) done -lemma mv_box_step_ex: "length lm = n \ - \stp. abc_steps_l (0, lm @ Suc x # suf_lm) [Dec n 2, Goto 0] stp - = (0, lm @ x # suf_lm)" -apply(rule_tac x = "Suc (Suc 0)" in exI, - simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps - abc_lm_v.simps abc_lm_s.simps nth_append list_update_append) -done - -lemma mv_box_ex': - "\length lm = n\ \ - \ stp. abc_steps_l (0, lm @ x # suf_lm) [Dec n 2, Goto 0] stp = - (Suc (Suc 0), lm @ 0 # suf_lm)" -apply(induct x) -apply(rule_tac x = "Suc 0" in exI, - simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps - abc_lm_v.simps nth_append abc_lm_s.simps, simp) -apply(drule_tac x = x and suf_lm = suf_lm in 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 \ list = drop (Suc n) lm" -apply(induct n arbitrary: lm a list, simp) -apply(case_tac "lm", simp, simp) -done - -lemma empty_boxes_ex: "\length lm \ n\ - \ \stp. abc_steps_l (0, lm) (empty_boxes n) stp = - (2*n, 0\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 \ length lm" "drop n lm = a # list" - thus "\bstp. abc_steps_l (0, 0\n @ a # list) [Dec n 2, Goto 0] bstp = - (Suc (Suc 0), 0 # 0\n @ drop (Suc n) lm)" - apply(insert mv_box_ex'[of "0\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) +lemma empty_one_box_correct: + "{\nl. nl = 0 \ n @ x # lm} [Dec n 2, Goto 0] {\nl. nl = 0 # 0 \ n @ lm}" +proof(induct x) + case 0 + thus "?case" + by(simp add: abc_Hoare_halt_def, + rule_tac x = 1 in exI, simp add: abc_steps_l.simps + abc_step_l.simps abc_fetch.simps abc_lm_v.simps nth_append abc_lm_s.simps + replicate_Suc[THEN sym] exp_suc del: replicate_Suc) +next + case (Suc x) + have "{\nl. nl = 0 \ n @ x # lm} [Dec n 2, Goto 0] {\nl. nl = 0 # 0 \ n @ lm}" + by fact + then obtain stp where "abc_steps_l (0, 0 \ n @ x # lm) [Dec n 2, Goto 0] stp + = (Suc (Suc 0), 0 # 0 \ n @ lm)" + apply(auto simp: abc_Hoare_halt_def) + by(case_tac "abc_steps_l (0, 0 \ n @ x # lm) [Dec n 2, Goto 0] na", simp) + moreover have "abc_steps_l (0, 0\n @ Suc x # lm) [Dec n 2, Goto 0] (Suc (Suc 0)) + = (0, 0 \ n @ x # lm)" + by(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_lm_v.simps + nth_append abc_lm_s.simps list_update.simps list_update_append) + ultimately have "abc_steps_l (0, 0\n @ Suc x # lm) [Dec n 2, Goto 0] (Suc (Suc 0) + stp) + = (Suc (Suc 0), 0 # 0\n @ lm)" + by(simp only: abc_steps_add) + thus "?case" + apply(simp add: abc_Hoare_halt_def) + apply(rule_tac x = "Suc (Suc stp)" in exI, simp) done qed +lemma empty_boxes_correct: + "length lm \ n \ + {\ nl. nl = lm} empty_boxes n {\ nl. nl = 0\n @ drop n lm}" +proof(induct n) + case 0 + thus "?case" + by(simp add: empty_boxes.simps abc_Hoare_halt_def, + rule_tac x = 0 in exI, simp add: abc_steps_l.simps) +next + case (Suc n) + have ind: "n \ length lm \ {\nl. nl = lm} empty_boxes n {\nl. nl = 0 \ n @ drop n lm}" by fact + have h: "Suc n \ length lm" by fact + have "{\nl. nl = lm} empty_boxes n [+] [Dec n 2, Goto 0] {\nl. nl = 0 # 0 \ n @ drop (Suc n) lm}" + proof(rule_tac abc_Hoare_plus_halt) + show "{\nl. nl = lm} empty_boxes n {\nl. nl = 0 \ n @ drop n lm}" + using h + by(rule_tac ind, simp) + next + show "{\nl. nl = 0 \ n @ drop n lm} [Dec n 2, Goto 0] {\nl. nl = 0 # 0 \ n @ drop (Suc n) lm}" + using empty_one_box_correct[of n "lm ! n" "drop (Suc n) lm"] + using h + by(simp add: nth_drop') + qed + thus "?case" + by(simp add: empty_boxes.simps) +qed -lemma mv_box_paras_prog_ex: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_ci f = (a, aa, ba); - Max (set (Suc n # ba # map (\(aprog, p, n). n) - (map rec_ci (gs)))) = pstr\ - \ \ ap bp cp. aprog = ap [+] bp [+] cp \ - length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + - 6 *length gs + 3 * n + length a + 3 \ bp = empty_boxes (length gs)" -apply(simp add: rec_ci.simps) -apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] - mv_boxes 0 (Suc (max (Suc n) (Max - (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) n - [+] mv_boxes (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs) [+] - a [+] mv_box aa (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))))" - in exI, simp add: cn_merge_gs_len) -apply(rule_tac x = " mv_box (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] - mv_boxes (Suc (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) 0 n" in exI, - auto simp: abc_append_commute) +lemma [simp]: "length gs \ ffp \ + length gs + (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) = + max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" +apply(rule_tac le_add_diff_inverse) +apply(rule_tac min_max.le_supI2) +apply(simp add: Max_ge_iff) 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" - "\k(aprog, p, n). n) - (map rec_ci (gs))))" - and starts: "ss = (\(ap, pos, n)\map rec_ci gs. length ap) + - 6 * length gs + 3 * n + length a + 3" - shows "\stp. abc_steps_l - (ss, ys @ 0\(pstr - length ys) @ rs # 0\length ys - @ lm @ 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp = - (ss + 2 * length gs, 0\pstr @ rs # 0\length ys @ lm @ - 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm)" + +lemma clean_paras: + "ffp \ length gs \ + {\nl. nl = map (\i. rec_exec i xs) gs @ + 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @ + rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything} + empty_boxes (length gs) + {\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ + rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything}" +proof- + let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" + assume h: "length gs \ ffp" + let ?lm = "map (\i. rec_exec i xs) gs @ 0 \ (?ft - length gs) @ + rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" + have "{\ nl. nl = ?lm} empty_boxes (length gs) {\ nl. nl = 0\length gs @ drop (length gs) ?lm}" + by(rule_tac empty_boxes_correct, simp) + moreover have "0\length gs @ drop (length gs) ?lm + = 0 \ ?ft @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" + using h + by(simp add: replicate_merge_anywhere) + ultimately show "?thesis" + by metis +qed + + +lemma restore_rs: + "{\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ + rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything} + mv_box (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) + {\nl. nl = 0 \ length xs @ + rec_exec (Cn (length xs) f gs) xs # + 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - (length xs)) @ + 0 \ length gs @ xs @ anything}" proof - - from h and pdef and starts have k1: - "\ ap bp cp. aprog = ap [+] bp [+] cp \ - length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + - 6 *length gs + 3 * n + length a + 3 - \ bp = empty_boxes (length ys)" - by(drule_tac 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 \ pstr" - apply simp - apply(rule_tac min_max.le_supI2, simp) + let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" + let ?lm = "0\(length xs) @ 0\(?ft - (length xs)) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" + thm mv_box_correct + have "{\ nl. nl = ?lm} mv_box ?ft (length xs) {\ nl. nl = ?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0]}" + by(rule_tac mv_box_correct, simp, simp) + moreover have "?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0] + = 0 \ length xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft - (length xs)) @ 0 \ length gs @ xs @ anything" + apply(auto simp: list_update_append nth_append) + apply(case_tac ?ft, simp_all add: Suc_diff_le list_update.simps) + apply(simp add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc) done - from k1 show "?thesis" - proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE) - fix ap bp apa cp - assume "aprog = ap [+] bp [+] cp \ - length ap = (\(ap, pos, n)\map rec_ci gs. length ap) + - 6 * length gs + 3 * n + length a + 3 \ - bp = empty_boxes (length ys)" - thus"?thesis" - apply(simp, rule_tac abc_append_exc1, simp_all add: starts h) - apply(insert empty_boxes_ex[of - "length gs" "ys @ 0\(pstr - (length gs)) @ rs # - 0\length gs @ lm @ 0\(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 \ length gs = aa \ pstr \ ba", simp) - apply(simp add: j1 j2 j3) - done - qed + ultimately show "?thesis" + by(simp add: replicate_merge_anywhere) +qed + +lemma restore_orgin_paras: + "{\nl. nl = 0 \ length xs @ + rec_exec (Cn (length xs) f gs) xs # + 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length xs) @ 0 \ length gs @ xs @ anything} + mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs) + {\nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ + (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}" +proof - + let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" + thm mv_boxes_correct2 + have "{\ nl. nl = [] @ 0\(length xs) @ (rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft - length xs) @ 0 \ length gs) @ xs @ anything} + mv_boxes (Suc ?ft + length gs) 0 (length xs) + {\ nl. nl = [] @ xs @ (rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft - length xs) @ 0 \ length gs) @ 0\length xs @ anything}" + by(rule_tac mv_boxes_correct2, auto) + thus "?thesis" + by(simp add: replicate_merge_anywhere) qed -lemma restore_rs_prog_ex: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_ci f = (a, aa, ba); - Max (set (Suc n # ba # map (\(aprog, p, n). n) - (map rec_ci (gs)))) = pstr; - ss = (\(ap, pos, n)\map rec_ci gs. length ap) + - 8 * length gs + 3 * n + length a + 3\ - \ \ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = ss \ - bp = 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 (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] - mv_boxes 0 (Suc (max (Suc n) (Max (insert ba (((\(aprog, p, n). n) - \ rec_ci) ` set gs))) + length gs)) n [+] - mv_boxes (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs) [+] - a [+] mv_box aa (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] - empty_boxes (length gs)" in exI, simp add: cn_merge_gs_len) -apply(rule_tac x = "mv_boxes (Suc (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) - + length gs)) 0 n" - in exI, auto simp: abc_append_commute) -done - -lemma exp_add: "a\(b+c) = a\b @ a\c" -apply(simp add:replicate_add) -done - -lemma [simp]: "n < pstr \ (0\pstr)[n := rs] @ [0::nat] = 0\n @ rs # 0\(pstr - n)" -using list_update_length[of "0\n" "0::nat" "0\(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" - "\k(aprog, p, n). n) - (map rec_ci (gs))))" - and starts: "ss = (\(ap, pos, n)\map rec_ci gs. length ap) + - 8 * length gs + 3 * n + length a + 3" - shows "\stp. abc_steps_l - (ss, 0\pstr @ rs # 0\length ys @ lm @ - 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp = - (ss + 3, 0\n @ rs # 0\(pstr - n) @ 0\length ys @ lm @ - 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm)" +lemma compile_cn_correct': + assumes f_ind: + "\ anything r. rec_exec f (map (\g. rec_exec g xs) gs) = rec_exec (Cn (length xs) f gs) xs \ + {\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (ffp - far) @ anything} fap + {\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (ffp - Suc far) @ anything}" + and compile: "rec_ci f = (fap, far, ffp)" + and term_f: "terminate f (map (\g. rec_exec g xs) gs)" + and g_cond: "\g\set gs. terminate g xs \ + (\x xa xb. rec_ci g = (x, xa, xb) \ + (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc}))" + shows + "{\nl. nl = xs @ 0 # 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything} + cn_merge_gs (map rec_ci gs) (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) [+] + (mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs) [+] + (mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs) [+] + (fap [+] (mv_box far (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) [+] + (empty_boxes (length gs) [+] + (mv_box (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) [+] + mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs))))))) + {\nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # +0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}" proof - - from h and pdef and starts have k1: - "\ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = ss \ - 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 \ length ap = ss \ - bp = 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\pstr @ rs # 0\length gs @ - lm @ 0\(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 + let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" + let ?A = "cn_merge_gs (map rec_ci gs) ?ft" + let ?B = "mv_boxes 0 (Suc (?ft+length gs)) (length xs)" + let ?C = "mv_boxes ?ft 0 (length gs)" + let ?D = fap + let ?E = "mv_box far ?ft" + let ?F = "empty_boxes (length gs)" + let ?G = "mv_box ?ft (length xs)" + let ?H = "mv_boxes (Suc (?ft + length gs)) 0 (length xs)" + let ?P1 = "\nl. nl = xs @ 0 # 0 \ (?ft + length gs) @ anything" + let ?S = "\nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft + length gs) @ anything" + let ?Q1 = "\ nl. nl = xs @ 0\(?ft - length xs) @ map (\ i. rec_exec i xs) gs @ 0\(Suc (length xs)) @ anything" + show "{?P1} (?A [+] (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))))) {?S}" + proof(rule_tac abc_Hoare_plus_halt) + show "{?P1} ?A {?Q1}" + using g_cond + by(rule_tac compile_cn_gs_correct, auto) + next + let ?Q2 = "\nl. nl = 0 \ ?ft @ + map (\i. rec_exec i xs) gs @ 0 # xs @ anything" + show "{?Q1} (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H)))))) {?S}" + proof(rule_tac abc_Hoare_plus_halt) + show "{?Q1} ?B {?Q2}" + by(rule_tac save_paras) + next + let ?Q3 = "\ nl. nl = map (\i. rec_exec i xs) gs @ 0\?ft @ 0 # xs @ anything" + show "{?Q2} (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))) {?S}" + proof(rule_tac abc_Hoare_plus_halt) + have "ffp \ length gs" + using compile term_f + apply(subgoal_tac "length gs = far") + apply(drule_tac footprint_ge, simp) + by(drule_tac param_pattern, auto) + thus "{?Q2} ?C {?Q3}" + by(erule_tac restore_new_paras) + next + let ?Q4 = "\ nl. nl = map (\i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0\?ft @ xs @ anything" + have a: "far = length gs" + using compile term_f + by(drule_tac param_pattern, auto) + have b:"?ft \ ffp" + by auto + have c: "ffp > far" + using compile + by(erule_tac footprint_ge) + show "{?Q3} (?D [+] (?E [+] (?F [+] (?G [+] ?H)))) {?S}" + proof(rule_tac abc_Hoare_plus_halt) + have "{\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (ffp - far) @ 0\(?ft - ffp + far) @ 0 # xs @ anything} fap + {\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec (Cn (length xs) f gs) xs # + 0 \ (ffp - Suc far) @ 0\(?ft - ffp + far) @ 0 # xs @ anything}" + by(rule_tac f_ind, simp add: rec_exec.simps) + thus "{?Q3} fap {?Q4}" + using a b c + by(simp add: replicate_merge_anywhere, + case_tac "?ft", simp_all add: exp_suc del: replicate_Suc) + next + let ?Q5 = "\nl. nl = map (\i. rec_exec i xs) gs @ + 0\(?ft - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0\(length gs)@ xs @ anything" + show "{?Q4} (?E [+] (?F [+] (?G [+] ?H))) {?S}" + proof(rule_tac abc_Hoare_plus_halt) + from a b c show "{?Q4} ?E {?Q5}" + by(erule_tac save_rs, simp_all) + next + let ?Q6 = "\nl. nl = 0\?ft @ rec_exec (Cn (length xs) f gs) xs # 0\(length gs)@ xs @ anything" + show "{?Q5} (?F [+] (?G [+] ?H)) {?S}" + proof(rule_tac abc_Hoare_plus_halt) + have "length gs \ ffp" using a b c + by simp + thus "{?Q5} ?F {?Q6}" + by(erule_tac clean_paras) + next + let ?Q7 = "\nl. nl = 0\length xs @ rec_exec (Cn (length xs) f gs) xs # 0\(?ft - (length xs)) @ 0\(length gs)@ xs @ anything" + show "{?Q6} (?G [+] ?H) {?S}" + proof(rule_tac abc_Hoare_plus_halt) + show "{?Q6} ?G {?Q7}" + by(rule_tac restore_rs) + next + show "{?Q7} ?H {?S}" + by(rule_tac restore_orgin_paras) + qed + qed + qed + qed + qed + qed qed qed -lemma [simp]:"xs \ [] \ length xs \ Suc 0" -by(case_tac xs, auto) - -lemma [simp]: "n < max (Suc n) (max ba (Max (((\(aprog, p, n). n) o - rec_ci) ` set gs)))" -by(simp) - -lemma restore_paras_prog_ex: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_ci f = (a, aa, ba); - Max (set (Suc n # ba # map (\(aprog, p, n). n) - (map rec_ci (gs)))) = pstr; - ss = (\(ap, pos, n)\map rec_ci gs. length ap) + - 8 * length gs + 3 * n + length a + 6\ - \ \ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = ss \ - bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n" -apply(simp add: rec_ci.simps) -apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) - [+] mv_boxes 0 (Suc (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs))) - + length gs)) n [+] mv_boxes (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs) [+] - a [+] mv_box aa (max (Suc n) - (Max (insert ba (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] - empty_boxes (length gs) [+] - mv_box (max (Suc n) (Max (insert ba - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n" in exI, simp add: cn_merge_gs_len) -apply(rule_tac x = "[]" in exI, auto simp: abc_append_commute) -done - -lemma restore_paras: - assumes h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" - "rec_calc_rel (Cn n f gs) lm rs" - "\k(aprog, p, n). n) - (map rec_ci (gs))))" - and starts: "ss = (\(ap, pos, n)\map rec_ci gs. length ap) + - 8 * length gs + 3 * n + length a + 6" - shows "\stp. abc_steps_l (ss, 0\n @ rs # 0\(pstr - n+ length ys) @ - lm @ 0\(a_md - Suc (pstr + length ys + n)) @ suf_lm) - aprog stp = (ss + 3 * n, lm @ rs # 0\(a_md - Suc n) @ suf_lm)" -proof - - from h and pdef and starts have k1: - "\ ap bp cp. aprog = ap [+] bp [+] cp \ length ap = ss \ - bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n" - by(drule_tac restore_paras_prog_ex, auto) - from k1 show "?thesis" - proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE) - fix ap bp apa cp - assume "aprog = ap [+] bp [+] cp \ length ap = ss \ - bp = mv_boxes (pstr + Suc (length gs)) 0 n" - thus"?thesis" - apply(simp, rule_tac abc_append_exc1, simp_all add: starts h) - apply(insert mv_boxes_ex2[of n "pstr + Suc (length gs)" 0 "[]" - "rs # 0\(pstr - n + length gs)" "lm" - "0\(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp) - apply(subgoal_tac "pstr > n \ - a_md > pstr + length gs + n \ length lm = n" , simp add: exponent_add_iff h) - using h pdef - apply(simp) - apply(frule_tac a = a and - aa = aa and ba = ba in ci_cn_md_def, simp, simp) - apply(subgoal_tac "length lm = rs_pos", - simp add: ci_cn_para_eq, erule_tac para_pattern, simp) - done - qed -qed - -lemma ci_cn_length: - "\rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); - rec_calc_rel (Cn n f gs) lm rs; - rec_ci f = (a, aa, ba)\ - \ length aprog = (\(ap, pos, n)\map rec_ci gs. length ap) + - 8 * length gs + 6 * n + length a + 6" -apply(simp add: rec_ci.simps, auto simp: cn_merge_gs_len) -done - -lemma cn_case: - assumes ind1: - "\x aprog a_md rs_pos rs suf_lm lm. - \x \ set gs; - rec_ci x = (aprog, rs_pos, a_md); - rec_calc_rel x lm rs\ - \ \stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (length aprog, lm @ [rs] @ 0\(a_md - rs_pos - 1) @ suf_lm)" - and ind2: - "\x ap fp arity r anything args. - \x = map rec_ci gs; rec_ci f = (ap, arity, fp); rec_calc_rel f args r\ - \ \stp. abc_steps_l (0, args @ 0 \ (fp - arity) @ anything) ap stp = - (length ap, args @ [r] @ 0 \ (fp - arity - 1) @ anything)" - and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" - "rec_calc_rel (Cn n f gs) lm rs" - shows "\stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp - = (length aprog, lm @ [rs] @ 0\(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 (\(aprog, p, n). n) - (map rec_ci (gs)))))" - let ?gs_len = "listsum (map (\ (ap, pos, n). length ap) - (map rec_ci (gs)))" - assume g: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" - "rec_calc_rel (Cn n f gs) lm rs" - "\k stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (?gs_len + 3 * length gs, lm @ 0\(?pstr - n) @ ys @ - 0\(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 ind1, auto) - done - from g have k2: - "\ stp. abc_steps_l (?gs_len + 3 * length gs, lm @ - 0\(?pstr - n) @ ys @ 0\(a_md - ?pstr - length ys) @ suf_lm) aprog stp = - (?gs_len + 3 * length gs + 3 * n, 0\?pstr @ ys @ 0 # lm @ - 0\(a_md - Suc (?pstr + length ys + n)) @ suf_lm)" - apply(erule_tac ba = c in save_paras, auto intro: ci_cn_para_eq) - done - from g have k3: - "\ stp. abc_steps_l (?gs_len + 3 * length gs + 3 * n, - 0\?pstr @ ys @ 0 # lm @ 0\(a_md - Suc (?pstr + length ys + n)) @ suf_lm) aprog stp = - (?gs_len + 6 * length gs + 3 * n, - ys @ 0\?pstr @ 0 # lm @ 0\(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: - "\stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n, - ys @ 0\?pstr @ 0 # lm @ 0\(a_md - Suc (?pstr + length ys + n)) @ suf_lm) aprog stp = - (?gs_len + 6 * length gs + 3 * n + length a, - ys @ rs # 0\?pstr @ lm @ 0\(a_md - Suc (?pstr + length ys + n)) @ suf_lm)" - apply(rule_tac ba = c in calc_cn_f, rule_tac ind2, auto) - done - from g h have k5: - "\ stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + length a, - ys @ rs # 0\?pstr @ lm @ 0\(a_md - Suc (?pstr + length ys + n)) @ suf_lm) - aprog stp = - (?gs_len + 6 * length gs + 3 * n + length a + 3, - ys @ 0\(?pstr - length ys) @ rs # 0\length ys @ lm @ - 0\(a_md - Suc (?pstr + length ys + n)) @ suf_lm)" - apply(rule_tac save_rs, auto simp: h) - done - from g have k6: - "\ stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + - length a + 3, ys @ 0\(?pstr - length ys) @ rs # 0\length ys @ lm @ - 0\(a_md - Suc (?pstr + length ys + n)) @ suf_lm) - aprog stp = - (?gs_len + 8 * length gs + 3 *n + length a + 3, - 0\?pstr @ rs # 0\length ys @ lm @ - 0\(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: - "\ stp. abc_steps_l (?gs_len + 8 * length gs + 3 *n + - length a + 3, 0\?pstr @ rs # 0\length ys @ lm @ - 0\(a_md -Suc (?pstr + length ys + n)) @ suf_lm) aprog stp = - (?gs_len + 8 * length gs + 3 * n + length a + 6, - 0\n @ rs # 0\(?pstr - n) @ 0\length ys @ lm @ - 0\(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: "\ stp. abc_steps_l (?gs_len + 8 * length gs + - 3 * n + length a + 6, - 0\n @ rs # 0\(?pstr - n) @ 0\length ys @ lm @ - 0\(a_md -Suc (?pstr + length ys + n)) @ suf_lm) aprog stp = - (?gs_len + 8 * length gs + 6 * n + length a + 6, - lm @ rs # 0\(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 - "\stp. abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp = - (length aprog, lm @ [rs] @ 0\(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) +lemma compile_cn_correct: + assumes termi_f: "terminate f (map (\g. rec_exec g xs) gs)" + and f_ind: "\ap arity fp anything. + rec_ci f = (ap, arity, fp) + \ {\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (fp - arity) @ anything} ap + {\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec f (map (\g. rec_exec g xs) gs) # 0 \ (fp - Suc arity) @ anything}" + and g_cond: + "\g\set gs. terminate g xs \ + (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc}))" + and compile: "rec_ci (Cn n f gs) = (ap, arity, fp)" + and len: "length xs = n" + shows "{\nl. nl = xs @ 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ rec_exec (Cn n f gs) xs # 0 \ (fp - Suc arity) @ anything}" +proof(case_tac "rec_ci f") + fix fap far ffp + assume h: "rec_ci f = (fap, far, ffp)" + then have f_newind: "\ anything .{\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (ffp - far) @ anything} fap + {\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec f (map (\g. rec_exec g xs) gs) # 0 \ (ffp - Suc far) @ anything}" + by(rule_tac f_ind, simp_all) + thus "{\nl. nl = xs @ 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ rec_exec (Cn n f gs) xs # 0 \ (fp - Suc arity) @ anything}" + using compile len h termi_f g_cond + apply(auto simp: rec_ci.simps abc_comp_commute) + apply(rule_tac compile_cn_correct', simp_all) 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 [simp]: + "\length xs = n; ft = max (n+3) (max fft gft)\ + \ {\nl. nl = xs @ 0 # 0 \ (ft - n) @ anything} mv_box n ft + {\nl. nl = xs @ 0 # 0 \ (ft - n) @ anything}" +using mv_box_correct[of n ft "xs @ 0 # 0 \ (ft - n) @ anything"] +by(auto) + +lemma [simp]: "length xs < max (length xs + 3) (max fft gft)" +by auto + +lemma save_init_rs: + "\length xs = n; ft = max (n+3) (max fft gft)\ + \ {\nl. nl = xs @ rec_exec f xs # 0 \ (ft - n) @ anything} mv_box n (Suc n) + {\nl. nl = xs @ 0 # rec_exec f xs # 0 \ (ft - Suc n) @ anything}" +using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \ (ft - n) @ anything"] +apply(auto simp: list_update_append list_update.simps nth_append split: if_splits) +apply(case_tac "(max (length xs + 3) (max fft gft))", simp_all add: list_update.simps Suc_diff_le) +done + +lemma [simp]: "n + (2::nat) < max (n + 3) (max fft gft)" +by auto + +lemma [simp]: "n < max (n + (3::nat)) (max fft gft)" +by auto + +lemma [simp]: + "length xs = n \ + {\nl. nl = xs @ x # 0 \ (max (n + (3::nat)) (max fft gft) - n) @ anything} mv_box n (max (n + 3) (max fft gft)) + {\nl. nl = xs @ 0 \ (max (n + 3) (max fft gft) - n) @ x # anything}" +proof - + assume h: "length xs = n" + let ?ft = "max (n+3) (max fft gft)" + let ?lm = "xs @ x # 0\(?ft - Suc n) @ 0 # anything" + have g: "?ft > n + 2" + by simp + thm mv_box_correct + have a: "{\ nl. nl = ?lm} mv_box n ?ft {\ nl. nl = ?lm[?ft := ?lm!n + ?lm!?ft, n := 0]}" + using h + by(rule_tac mv_box_correct, auto) + have b:"?lm = xs @ x # 0 \ (max (n + 3) (max fft gft) - n) @ anything" + by(case_tac ?ft, simp_all add: Suc_diff_le exp_suc del: replicate_Suc) + have c: "?lm[?ft := ?lm!n + ?lm!?ft, n := 0] = xs @ 0 \ (max (n + 3) (max fft gft) - n) @ x # anything" + using h g + apply(auto simp: nth_append list_update_append split: if_splits) + using list_update_append[of "x # 0 \ (max (length xs + 3) (max fft gft) - Suc (length xs))" "0 # anything" + "max (length xs + 3) (max fft gft) - length xs" "x"] + apply(auto simp: if_splits) + apply(simp add: list_update.simps replicate_Suc[THEN sym] del: replicate_Suc) + done + from a c show "?thesis" + using h + apply(simp) + using b + by simp +qed + +lemma [simp]: "max n (Suc n) < Suc (Suc (max (n + 3) (max fft gft) + length anything - Suc 0))" +by arith + +lemma [simp]: "Suc n < max (n + 3) (max fft gft)" +by arith + +lemma [simp]: + "length xs = n + \ {\nl. nl = xs @ rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc n) @ x # anything} mv_box n (Suc n) + {\nl. nl = xs @ 0 # rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc (Suc n)) @ x # anything}" +using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc n) @ x # anything"] +apply(simp add: nth_append list_update_append list_update.simps) +apply(case_tac "max (n + 3) (max fft gft)", simp_all) +apply(case_tac nat, simp_all add: Suc_diff_le list_update.simps) +done -lemma recursive_compile_correct: - "\rec_ci recf = (ap, arity, fp); - rec_calc_rel recf args r\ - \ (\ stp. (abc_steps_l (0, args @ 0\(fp - arity) @ anything) ap stp) = - (length ap, args@[r]@0\(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: - "\ap fp arity r anything args. - \aa = ap \ ba = arity \ ca = fp; rec_calc_rel f args r\ \ - \stp. abc_steps_l (0, args @ 0\(fp - arity) @ anything) ap stp = - (length ap, args @ r # 0\(fp - Suc arity) @ anything)" - and g_ind: - "\x xa y xb ya ap fp arity r anything args. - \x = (aa, ba, ca); xa = aa \ y = (ba, ca); xb = ba \ ya = ca; - a = ap \ b = arity \ c = fp; rec_calc_rel g args r\ - \ \stp. abc_steps_l (0, args @ 0\(fp - arity) @ anything) ap stp = - (length ap, args @ r # 0\(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: - "\ args r anything. rec_calc_rel f args r \ - \stp. abc_steps_l (0, args @ 0\(ca - ba) @ anything) aa stp = - (length aa, args @ r # 0\(ca - Suc ba) @ anything)" - and ng_ind: - "\ args r anything. rec_calc_rel g args r \ - \stp. abc_steps_l (0, args @ 0\(c - b) @ anything) a stp = - (length a, args @ r # 0\(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 - "\stp. abc_steps_l (0, args @ 0\(fp - arity) @ anything) ap stp = - (length ap, args @ r # 0\(fp - Suc arity) @ anything)" - apply(auto intro: nf_ind ng_ind pr_case) +lemma abc_append_frist_steps_eq_pre: + assumes notfinal: "abc_notfinal (abc_steps_l (0, lm) A n) A" + and notnull: "A \ []" + shows "abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n" +using notfinal +proof(induct n) + case 0 + thus "?case" + by(simp add: abc_steps_l.simps) +next + case (Suc n) + have ind: "abc_notfinal (abc_steps_l (0, lm) A n) A \ abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n" + by fact + have h: "abc_notfinal (abc_steps_l (0, lm) A (Suc n)) A" by fact + then have a: "abc_notfinal (abc_steps_l (0, lm) A n) A" + by(simp add: notfinal_Suc) + then have b: "abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n" + using ind by simp + obtain s lm' where c: "abc_steps_l (0, lm) A n = (s, lm')" + by (metis prod.exhaust) + then have d: "s < length A \ abc_steps_l (0, lm) (A @ B) n = (s, lm')" + using a b by simp + thus "?case" + using c + by(simp add: abc_step_red2 abc_fetch.simps abc_step_l.simps nth_append) +qed + +lemma abc_append_first_step_eq_pre: + "st < length A + \ abc_step_l (st, lm) (abc_fetch st (A @ B)) = + abc_step_l (st, lm) (abc_fetch st A)" +by(simp add: abc_step_l.simps abc_fetch.simps nth_append) + +lemma abc_append_frist_steps_halt_eq': + assumes final: "abc_steps_l (0, lm) A n = (length A, lm')" + and notnull: "A \ []" + shows "\ n'. abc_steps_l (0, lm) (A @ B) n' = (length A, lm')" +proof - + have "\ n'. abc_notfinal (abc_steps_l (0, lm) A n') A \ + abc_final (abc_steps_l (0, lm) A (Suc n')) A" + using assms + by(rule_tac n = n in abc_before_final, simp_all) + then obtain na where a: + "abc_notfinal (abc_steps_l (0, lm) A na) A \ + abc_final (abc_steps_l (0, lm) A (Suc na)) A" .. + obtain sa lma where b: "abc_steps_l (0, lm) A na = (sa, lma)" + by (metis prod.exhaust) + then have c: "abc_steps_l (0, lm) (A @ B) na = (sa, lma)" + using a abc_append_frist_steps_eq_pre[of lm A na B] assms + by simp + have d: "sa < length A" using b a by simp + then have e: "abc_step_l (sa, lma) (abc_fetch sa (A @ B)) = + abc_step_l (sa, lma) (abc_fetch sa A)" + by(rule_tac abc_append_first_step_eq_pre) + from a have "abc_steps_l (0, lm) A (Suc na) = (length A, lm')" + using final equal_when_halt + by(case_tac "abc_steps_l (0, lm) A (Suc na)" , simp) + then have "abc_steps_l (0, lm) (A @ B) (Suc na) = (length A, lm')" + using a b c e + by(simp add: abc_step_red2) + thus "?thesis" + by blast +qed + +lemma abc_append_frist_steps_halt_eq: + assumes final: "abc_steps_l (0, lm) A n = (length A, lm')" + shows "\ n'. abc_steps_l (0, lm) (A @ B) n' = (length A, lm')" +using final +apply(case_tac "A = []") +apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps abc_exec_null) +apply(rule_tac abc_append_frist_steps_halt_eq', simp_all) +done + +lemma [simp]: "Suc (Suc (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))) + = max (length xs + 3) (max fft gft) - (length xs)" +by arith + +lemma [simp]: "\ft = max (n + 3) (max fft gft); length xs = n\ \ + {\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything} + [Dec ft (length gap + 7)] + {\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything}" +apply(simp add: abc_Hoare_halt_def) +apply(rule_tac x = 1 in exI) +apply(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps list_update_append) +using list_update_length +[of "(x - Suc y) # rec_exec (Pr (length xs) f g) (xs @ [x - Suc y]) # + 0 \ (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" "Suc y" anything y] +apply(simp) +done + +lemma adjust_paras': + "length xs = n \ {\nl. nl = xs @ x # y # anything} [Inc n] [+] [Dec (Suc n) 2, Goto 0] + {\nl. nl = xs @ Suc x # 0 # anything}" +proof(rule_tac abc_Hoare_plus_halt) + assume "length xs = n" + thus "{\nl. nl = xs @ x # y # anything} [Inc n] {\ nl. nl = xs @ Suc x # y # anything}" + apply(simp add: abc_Hoare_halt_def) + apply(rule_tac x = 1 in exI, simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_comp.simps + abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps) done next - fix ap fp arity r anything args - assume h: - "rec_ci z = (ap, arity, fp)" "rec_calc_rel z args r" - thus "\stp. abc_steps_l (0, args @ 0\(fp - arity) @ anything) ap stp = - (length ap, args @ [r] @ 0\(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 - "\stp. abc_steps_l (0, args @ 0\(fp - arity) @ anything) ap stp = - (length ap, args @ [r] @ 0\(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 "\stp. abc_steps_l (0, args @ 0\(fp - arity) @ anything) ap stp - = (length ap, args @ [r] @ 0\(fp - arity - 1) @ anything)" - by(erule_tac id_case) -next - fix n f gs ap fp arity r anything args - assume ind1: "\x ap fp arity r anything args. - \x \ set gs; - rec_ci x = (ap, arity, fp); - rec_calc_rel x args r\ - \ \stp. abc_steps_l (0, args @ 0\(fp - arity) @ anything) ap stp = - (length ap, args @ [r] @ 0\(fp - arity - 1) @ anything)" - and ind2: - "\x ap fp arity r anything args. - \x = map rec_ci gs; rec_ci f = (ap, arity, fp); rec_calc_rel f args r\ - \ \stp. abc_steps_l (0, args @ 0 \ (fp - arity) @ anything) ap stp = - (length ap, args @ [r] @ 0 \ (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 - "\stp. abc_steps_l (0, args @ 0\(fp - arity) @ anything) - ap stp = (length ap, args @ [r] @ 0\(fp - arity - 1) @ anything)" - apply(rule_tac cn_case) - apply(rule_tac ind1, simp, simp, simp) - apply(rule_tac ind2, auto) - done -next - fix n f ap fp arity r anything args - assume ind: - "\ap fp arity r anything args. - \rec_ci f = (ap, arity, fp); rec_calc_rel f args r\ \ - \stp. abc_steps_l (0, args @ 0\(fp - arity) @ anything) ap stp = - (length ap, args @ [r] @ 0\(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 - "\stp. abc_steps_l (0, args @ 0\(fp - arity) @ anything) ap stp = - (length ap, args @ [r] @ 0\(fp - arity - 1) @ anything)" - apply(rule_tac mn_case, rule_tac ind, auto) - done -qed - -lemma abc_append_uhalt1: - "\\ stp. (\ (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp); - p = ap [+] bp [+] cp\ - \ \ stp. (\ (ss, e). ss < length p) - (abc_steps_l (length ap, lm) p stp)" -apply(auto) -apply(erule_tac x = stp in allE, auto) -apply(frule_tac ap = ap and cp = cp in abc_append_state_in_exc, auto) -done - - -lemma abc_append_unhalt2: - "\abc_steps_l (0, am) ap stp = (length ap, lm); bp \ []; - \ stp. (\ (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp); - p = ap [+] bp [+] cp\ - \ \ stp. (\ (ss, e). ss < length p) (abc_steps_l (0, am) p stp)" -proof - - assume h: - "abc_steps_l (0, am) ap stp = (length ap, lm)" - "bp \ []" - "\ stp. (\ (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp)" - "p = ap [+] bp [+] cp" - have "\ stp. (abc_steps_l (0, am) p stp) = (length ap, lm)" - using h - apply(simp add: abc_append.simps) - apply(rule_tac abc_add_exc1, auto) - done - from this obtain stpa where g1: - "(abc_steps_l (0, am) p stpa) = (length ap, lm)" .. - moreover have g2: "\ stp. (\ (ss, e). ss < length p) - (abc_steps_l (length ap, lm) p stp)" - using h - apply(erule_tac abc_append_uhalt1, simp) - done - moreover from g1 and g2 have - "\ stp. (\ (ss, e). ss < length p) - (abc_steps_l (0, am) p (stpa + stp))" - apply(simp add: abc_steps_add) - done - thus "\ stp. (\ (ss, e). ss < length p) - (abc_steps_l (0, am) p stp)" - apply(rule_tac allI, auto) - apply(case_tac "stp \ stpa") - apply(erule_tac x = "stp - stpa" in allE, simp) - proof - - fix stp a b - assume g3: "abc_steps_l (0, am) p stp = (a, b)" - "\ stpa \ stp" - thus "a < length p" - using g1 h - apply(case_tac "a < length p", simp, simp) - apply(subgoal_tac "\ d. stpa = stp + d") - using abc_state_keep[of p a b "stpa - stp"] - apply(erule_tac exE, simp add: abc_steps_add) - apply(rule_tac x = "stpa - stp" in exI, simp) + assume h: "length xs = n" + thus "{\nl. nl = xs @ Suc x # y # anything} [Dec (Suc n) 2, Goto 0] {\nl. nl = xs @ Suc x # 0 # anything}" + proof(induct y) + case 0 + thus "?case" + apply(simp add: abc_Hoare_halt_def) + apply(rule_tac x = 1 in exI, simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_comp.simps + abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps) done + next + case (Suc y) + have "length xs = n \ + {\nl. nl = xs @ Suc x # y # anything} [Dec (Suc n) 2, Goto 0] {\nl. nl = xs @ Suc x # 0 # anything}" by fact + then obtain stp where + "abc_steps_l (0, xs @ Suc x # y # anything) [Dec (Suc n) 2, Goto 0] stp = (2, xs @ Suc x # 0 # anything)" + using h + apply(auto simp: abc_Hoare_halt_def) + by(case_tac "abc_steps_l (0, xs @ Suc x # y # anything) [Dec (Suc (length xs)) 2, Goto 0] n", + simp_all add: numeral_2_eq_2) + moreover have "abc_steps_l (0, xs @ Suc x # Suc y # anything) [Dec (Suc n) 2, Goto 0] 2 = + (0, xs @ Suc x # y # anything)" + using h + by(simp add: abc_steps_l.simps numeral_2_eq_2 abc_step_l.simps abc_fetch.simps + abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps) + ultimately show "?case" + apply(simp add: abc_Hoare_halt_def) + by(rule_tac x = "2 + stp" in exI, simp only: abc_steps_add, simp) 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 adjust_paras: + "length xs = n \ {\nl. nl = xs @ x # y # anything} [Inc n, Dec (Suc n) 3, Goto (Suc 0)] + {\nl. nl = xs @ Suc x # 0 # anything}" +using adjust_paras'[of xs n x y anything] +by(simp add: abc_comp.simps abc_shift.simps numeral_2_eq_2 numeral_3_eq_3) + +lemma [simp]: "\rec_ci g = (gap, gar, gft); \yx\ \ gar = Suc (Suc n)" + apply(erule_tac x = y in allE, simp) + apply(drule_tac param_pattern, auto) + done + +lemma loop_back': + assumes h: "length A = length gap + 4" "length xs = n" + and le: "y \ x" + shows "\ stp. abc_steps_l (length A, xs @ m # (y - x) # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp + = (length A, xs @ m # y # 0 # anything)" + using le +proof(induct x) + case 0 + thus "?case" + using h + by(rule_tac x = 0 in exI, + auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_s.simps abc_lm_v.simps) +next + case (Suc x) + have "x \ y \ \stp. abc_steps_l (length A, xs @ m # (y - x) # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp = + (length A, xs @ m # y # 0 # anything)" by fact + moreover have "Suc x \ y" by fact + moreover then have "\ stp. abc_steps_l (length A, xs @ m # (y - Suc x) # Suc x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp + = (length A, xs @ m # (y - x) # x # anything)" + using h + apply(rule_tac x = 3 in exI) + by(simp add: abc_steps_l.simps numeral_3_eq_3 abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps + list_update_append list_update.simps) + ultimately show "?case" + apply(auto) + apply(rule_tac x = "stpa + stp" in exI) + by(simp add: abc_steps_add) +qed + -lemma Mn_unhalt: - assumes mn_rf: "rf = Mn n f" - and compiled_mnrf: "rec_ci rf = (aprog, rs_pos, a_md)" - and compiled_f: "rec_ci f = (aprog', rs_pos', a_md')" - and args: "length lm = n" - and unhalt_condition: "\ y. (\ rs. rec_calc_rel f (lm @ [y]) rs \ rs \ 0)" - shows "\ stp. case abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) - aprog stp of (ss, e) \ ss < length aprog" - using mn_rf compiled_mnrf compiled_f args unhalt_condition -proof(rule_tac allI) - fix stp - assume h: "rf = Mn n f" - "rec_ci rf = (aprog, rs_pos, a_md)" - "rec_ci f = (aprog', rs_pos', a_md')" - "\y. \rs. rec_calc_rel f (lm @ [y]) rs \ rs \ 0" "length lm = n" - have "\stpa \ stp. abc_steps_l (0, lm @ 0 # 0\(a_md - Suc rs_pos) @ suf_lm) aprog stpa - = (0, lm @ stp # 0\(a_md - Suc rs_pos) @ suf_lm)" - proof(induct stp, auto) - show "\stpa. abc_steps_l (0, lm @ 0 # 0\(a_md - Suc rs_pos) @ suf_lm) - aprog stpa = (0, lm @ 0 # 0\(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 \ stpa" - and g2: "abc_steps_l (0, lm @ 0 # 0\(a_md - Suc rs_pos) @ suf_lm) - aprog stpa - = (0, lm @ stp # 0\(a_md - Suc rs_pos) @ suf_lm)" - have "\rs. rec_calc_rel f (lm @ [stp]) rs \ rs \ 0" - using h - apply(erule_tac x = stp in allE, simp) - done - from this obtain rs where g3: - "rec_calc_rel f (lm @ [stp]) rs \ rs \ 0" .. - hence "\ stpb. abc_steps_l (0, lm @ stp # 0\(a_md - Suc rs_pos) @ - suf_lm) aprog stpb - = (0, lm @ Suc stp # 0\(a_md - Suc rs_pos) @ suf_lm)" - using h - apply(rule_tac mn_ind_step) - apply(rule_tac recursive_compile_correct, simp, simp) +lemma loop_back: + assumes h: "length A = length gap + 4" "length xs = n" + shows "\ stp. abc_steps_l (length A, xs @ m # 0 # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp + = (0, xs @ m # x # 0 # anything)" +using loop_back'[of A gap xs n x x m anything] assms +apply(auto) +apply(rule_tac x = "stp + 1" in exI) +apply(simp only: abc_steps_add, simp) +apply(simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps) +done + +lemma rec_exec_pr_0_simps: "rec_exec (Pr n f g) (xs @ [0]) = rec_exec f xs" + by(simp add: rec_exec.simps) + +lemma rec_exec_pr_Suc_simps: "rec_exec (Pr n f g) (xs @ [Suc y]) + = rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])" +apply(induct y) +apply(simp add: rec_exec.simps) +apply(simp add: rec_exec.simps) +done + +lemma [simp]: "Suc (max (n + 3) (max fft gft) - Suc (Suc (Suc n))) = max (n + 3) (max fft gft) - Suc (Suc n)" +by arith + +lemma pr_loop: + assumes code: "code = ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ + [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" + and len: "length xs = n" + and g_ind: "\ yanything. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything} gap + {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything})" + and compile_g: "rec_ci g = (gap, gar, gft)" + and termi_g: "\ y x" + shows + "\stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) + code stp = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (ft - Suc (Suc n)) @ y # anything)" +proof - + let ?A = "[Dec ft (length gap + 7)]" + let ?B = "gap" + let ?C = "[Inc n, Dec (Suc n) 3, Goto (Suc 0)]" + let ?D = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" + have "\ stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) + ((?A [+] (?B [+] ?C)) @ ?D) stp = (length (?A [+] (?B [+] ?C)), + xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) + # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" + proof - + have "\ stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) + ((?A [+] (?B [+] ?C))) stp = (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 # + rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" 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 + have "{\ nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything} + (?A [+] (?B [+] ?C)) + {\ nl. nl = xs @ (x - y) # 0 # + rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything}" + proof(rule_tac abc_Hoare_plus_halt) + show "{\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything} + [Dec ft (length gap + 7)] + {\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything}" + using ft len + by(simp) + next + show + "{\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything} + ?B [+] ?C + {\nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything}" + proof(rule_tac abc_Hoare_plus_halt) + have a: "gar = Suc (Suc n)" + using compile_g termi_g len less + by simp + have b: "gft > gar" + using compile_g + by(erule_tac footprint_ge) + show "{\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything} gap + {\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # + rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything}" + proof - + have + "{\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (gft - gar) @ 0\(ft - gft) @ y # anything} gap + {\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # + rec_exec g (xs @ [(x - Suc y), rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (gft - Suc gar) @ 0\(ft - gft) @ y # anything}" + using g_ind less by simp + thus "?thesis" + using a b ft + by(simp add: replicate_merge_anywhere numeral_3_eq_3) + qed + next + show "{\nl. nl = xs @ (x - Suc y) # + rec_exec (Pr n f g) (xs @ [x - Suc y]) # + rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything} + [Inc n, Dec (Suc n) 3, Goto (Suc 0)] + {\nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) + (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything}" + using len less + using adjust_paras[of xs n "x - Suc y" " rec_exec (Pr n f g) (xs @ [x - Suc y])" + " rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # + 0 \ (ft - Suc (Suc (Suc n))) @ y # anything"] + by(simp add: Suc_diff_Suc) + qed + qed + thus "?thesis" + by(simp add: abc_Hoare_halt_def, auto, rule_tac x = na in exI, case_tac "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # + 0 \ (ft - Suc (Suc n)) @ Suc y # anything) + ([Dec ft (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) na", simp) qed - thus "\stpa\Suc stp. abc_steps_l (0, lm @ 0 # 0\(a_md - Suc rs_pos) @ - suf_lm) aprog stpa - = (0, lm @ Suc stp # 0\(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 + then obtain stpa where "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) + ((?A [+] (?B [+] ?C))) stpa = (length (?A [+] (?B [+] ?C)), + xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) + # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" .. + thus "?thesis" + by(erule_tac abc_append_frist_steps_halt_eq) qed - from this obtain stpa where - "stp \ stpa \ abc_steps_l (0, lm @ 0 # 0\(a_md - Suc rs_pos) @ suf_lm) - aprog stpa = (0, lm @ stp # 0\(a_md - Suc rs_pos) @ suf_lm)" .. - thus "case abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog stp - of (ss, e) \ ss < length aprog" - apply(case_tac "abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suf_lm) aprog - stp", simp, case_tac "a \ length aprog", - simp, simp) - apply(subgoal_tac "\ d. stpa = stp + d", erule_tac exE) - apply(subgoal_tac "lm @ 0\(a_md - rs_pos) @ suf_lm = lm @ 0 # - 0\(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 + moreover have + "\ stp. abc_steps_l (length (?A [+] (?B [+] ?C)), + xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything) + ((?A [+] (?B [+] ?C)) @ ?D) stp = (0, xs @ (x - y) # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # + 0 # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" + using len + by(rule_tac loop_back, simp_all) + moreover have "rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) = rec_exec (Pr n f g) (xs @ [x - y])" + using less + thm rec_exec.simps + apply(case_tac "x - y", simp_all add: rec_exec_pr_Suc_simps) + by(subgoal_tac "nat = x - Suc y", simp, arith) + ultimately show "?thesis" + using code ft + by(auto, rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add replicate_Suc_iff_anywhere del: replicate_Suc) +qed -lemma abc_append_cons_eq[intro!]: - "\ap = bp; cp = dp\ \ ap [+] cp = bp [+] dp" -by simp - -lemma cn_merge_gs_split: - "\i < length gs; rec_ci (gs!i) = (ga, gb, gc)\ \ - cn_merge_gs (map rec_ci gs) p = - cn_merge_gs (map rec_ci (take i gs)) p [+] ga [+] - 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]) +lemma [simp]: + "length xs = n \ abc_lm_s (xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) + (max fft gft) - Suc (Suc n)) @ 0 # anything) (max (n + 3) (max fft gft)) 0 = + xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc n) @ anything" +apply(simp add: abc_lm_s.simps) +using list_update_length[of "xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc (Suc n))" + 0 anything 0] +apply(auto simp: Suc_diff_Suc) +apply(simp add: exp_suc[THEN sym] Suc_diff_Suc del: replicate_Suc) done -text {* - Correctness of the complier (non-terminating case for Mn). There are many cases when a - recursive function does not terminate. For the purpose of Uiversal Turing Machine, we only - need to prove the case for @{text "Mn"} and @{text "Cn"}. - This lemma is for @{text "Cn"}. For @{text "Cn f g1 g2 \gi, gi+1, \ gn"}, this lemma describes what - happens when every one of @{text "g1, g2, \ gi"} terminates, but - @{text "gi+1"} does not terminate, so that whole function @{text "Cn f g1 g2 \gi, gi+1, \ gn"} - does not terminate. - *} +lemma [simp]: + "(xs @ x # rec_exec (Pr (length xs) f g) (xs @ [x]) # 0 \ (max (length xs + 3) + (max fft gft) - Suc (Suc (length xs))) @ 0 # anything) ! + max (length xs + 3) (max fft gft) = 0" +using nth_append_length[of "xs @ x # rec_exec (Pr (length xs) f g) (xs @ [x]) # + 0 \ (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" 0 anything] +by(simp) + +lemma pr_loop_correct: + assumes less: "y \ x" + and len: "length xs = n" + and compile_g: "rec_ci g = (gap, gar, gft)" + and termi_g: "\ y yanything. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything} gap + {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything})" + shows "{\nl. nl = xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (max (n + 3) (max fft gft) - Suc (Suc n)) @ y # anything} + ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)] + {\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc n) @ anything}" + using less +proof(induct y) + case 0 + thus "?case" + using len + apply(simp add: abc_Hoare_halt_def) + apply(rule_tac x = 1 in exI) + by(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps + nth_append abc_comp.simps abc_shift.simps, simp add: abc_lm_v.simps) +next + case (Suc y) + let ?ft = "max (n + 3) (max fft gft)" + let ?C = "[Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] + [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" + have ind: "y \ x \ + {\nl. nl = xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything} + ?C {\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything}" by fact + have less: "Suc y \ x" by fact + have stp1: + "\ stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (?ft - Suc (Suc n)) @ Suc y # anything) + ?C stp = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything)" + using assms less + by(rule_tac pr_loop, auto) + then obtain stp1 where a: + "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (?ft - Suc (Suc n)) @ Suc y # anything) + ?C stp1 = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything)" .. + moreover have + "\ stp. abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything) + ?C stp = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything)" + using ind less + by(auto simp: abc_Hoare_halt_def, case_tac "abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) + (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything) ?C na", rule_tac x = na in exI, simp) + then obtain stp2 where b: + "abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything) + ?C stp2 = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything)" .. + from a b show "?case" + by(simp add: abc_Hoare_halt_def, rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add) +qed -lemma cn_gi_uhalt: - assumes cn_recf: "rf = Cn n f gs" - and compiled_cn_recf: "rec_ci rf = (aprog, rs_pos, a_md)" - and args_length: "length lm = n" - and exist_unhalt_recf: "i < length gs" "gi = gs ! i" - and complied_unhalt_recf: "rec_ci gi = (ga, gb, gc)" "gb = n" - and all_halt_before_gi: "\ j < i. (\ rs. rec_calc_rel (gs!j) lm rs)" - and unhalt_condition: "\ slm. \ stp. case abc_steps_l (0, lm @ 0\(gc - gb) @ slm) - ga stp of (se, e) \ se < length ga" - shows " \ stp. case abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suflm) aprog - stp of (ss, e) \ ss < length aprog" - using cn_recf compiled_cn_recf args_length exist_unhalt_recf complied_unhalt_recf - all_halt_before_gi unhalt_condition -proof(case_tac "rec_ci f", simp) - fix a b c - assume h1: "rf = Cn n f gs" - "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" - "length lm = n" - "gi = gs ! i" - "rec_ci (gs!i) = (ga, n, gc)" - "gb = n" "rec_ci f = (a, b, c)" - and h2: "\jrs. rec_calc_rel (gs ! j) lm rs" - "i < length gs" - and ind: - "\ slm. \ stp. case abc_steps_l (0, lm @ 0\(gc - n) @ slm) ga stp of (se, e) \ se < length ga" - have h3: "rs_pos = n" - using h1 - by(rule_tac ci_cn_para_eq, simp) - let ?ggs = "take i gs" - have "\ ys. (length ys = i \ - (\ k < i. rec_calc_rel (?ggs ! k) lm (ys ! k)))" - using h2 - apply(induct i, simp, simp) - apply(erule_tac exE) - apply(erule_tac x = ia in allE, simp) - apply(erule_tac exE) - apply(rule_tac x = "ys @ [x]" in exI, simp add: nth_append, auto) - apply(subgoal_tac "k = length ys", simp, simp) - done - from this obtain ys where g1: - "(length ys = i \ (\ k < i. rec_calc_rel (?ggs ! k) - lm (ys ! k)))" .. - let ?pstr = "Max (set (Suc n # c # map (\(aprog, p, n). n) - (map rec_ci (gs))))" - have "\stp. abc_steps_l (0, lm @ 0\(a_md - n) @ suflm) - (cn_merge_gs (map rec_ci ?ggs) ?pstr) stp = - (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) ?ggs) + - 3 * length ?ggs, lm @ 0\(?pstr - n) @ ys @ 0\(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 \ set gs", simp) - using h2 - apply(rule_tac A = "set (take i gs)" in subsetD, - simp add: set_take_subset, simp) - done - from this obtain stpa where g2: - "abc_steps_l (0, lm @ 0\(a_md - n) @ suflm) - (cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa = - (listsum (map ((\(ap, pos, n). length ap) \ rec_ci) ?ggs) + - 3 * length ?ggs, lm @ 0\(?pstr - n) @ ys @ 0\(a_md -(?pstr + length ?ggs)) @ - suflm)" .. - moreover have - "\ cp. aprog = (cn_merge_gs - (map rec_ci ?ggs) ?pstr) [+] ga [+] cp" - using h1 - apply(simp add: rec_ci.simps) - apply(rule_tac x = "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 - (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + - length gs)) n [+] mv_boxes (max (Suc n) (Max (insert c - (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) 0 (length gs) [+] - a [+] mv_box b (max (Suc n) - (Max (insert c (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) [+] - empty_boxes (length gs) [+] mv_box (max (Suc n) - (Max (insert c (((\(aprog, p, n). n) \ rec_ci) ` set gs)))) n [+] - mv_boxes (Suc (max (Suc n) (Max (insert c - (((\(aprog, p, n). n) \ rec_ci) ` set gs))) + length gs)) 0 n" in exI) - apply(simp add: abc_append_commute [THEN sym]) - apply(auto) - using cn_merge_gs_split[of i gs ga "length lm" gc - "(max (Suc (length lm)) - (Max (insert c (((\(aprog, p, n). n) \ rec_ci) ` set gs))))"] - h2 - apply(simp) - done - from this obtain cp where g3: - "aprog = (cn_merge_gs (map rec_ci ?ggs) ?pstr) [+] ga [+] cp" .. - show "\ stp. case abc_steps_l (0, lm @ 0\(a_md - rs_pos) @ suflm) - aprog stp of (ss, e) \ ss < length aprog" - proof(rule_tac abc_append_unhalt2) - show "abc_steps_l (0, lm @ 0\(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\(?pstr - n) @ ys @ 0\(a_md -(?pstr + length ?ggs)) @ suflm)" - using h3 g2 - apply(simp add: cn_merge_gs_length) - done +lemma compile_pr_correct': + assumes termi_g: "\ y yanything. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything} gap + {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything})" + and termi_f: "terminate f xs" + and f_ind: "\ anything. {\nl. nl = xs @ 0 \ (fft - far) @ anything} fap {\nl. nl = xs @ rec_exec f xs # 0 \ (fft - Suc far) @ anything}" + and len: "length xs = n" + and compile1: "rec_ci f = (fap, far, fft)" + and compile2: "rec_ci g = (gap, gar, gft)" + shows + "{\nl. nl = xs @ x # 0 \ (max (n + 3) (max fft gft) - n) @ anything} + mv_box n (max (n + 3) (max fft gft)) [+] + (fap [+] (mv_box n (Suc n) [+] + ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ + [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]))) + {\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc n) @ anything}" +proof - + let ?ft = "max (n+3) (max fft gft)" + let ?A = "mv_box n ?ft" + let ?B = "fap" + let ?C = "mv_box n (Suc n)" + let ?D = "[Dec ?ft (length gap + 7)]" + let ?E = "gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]" + let ?F = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" + let ?P = "\nl. nl = xs @ x # 0 \ (?ft - n) @ anything" + let ?S = "\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything" + let ?Q1 = "\nl. nl = xs @ 0 \ (?ft - n) @ x # anything" + show "{?P} (?A [+] (?B [+] (?C [+] (?D [+] ?E @ ?F)))) {?S}" + proof(rule_tac abc_Hoare_plus_halt) + show "{?P} ?A {?Q1}" + using len by simp next - show "ga \ []" - using h1 - apply(simp add: rec_ci_not_null) - done - next - show "\stp. case abc_steps_l (0, lm @ 0\(?pstr - n) @ ys - @ 0\(a_md - (?pstr + length (take i gs))) @ suflm) ga stp of - (ss, e) \ ss < length ga" - using ind[of "0\(?pstr - gc) @ ys @ 0\(a_md - (?pstr + length (take i gs))) - @ suflm"] - apply(subgoal_tac "lm @ 0\(?pstr - n) @ ys - @ 0\(a_md - (?pstr + length (take i gs))) @ suflm - = lm @ 0\(gc - n) @ - 0\(?pstr - gc) @ ys @ 0\(a_md - (?pstr + length (take i gs))) @ suflm", simp) - apply(simp add: replicate_add[THEN sym]) - apply(subgoal_tac "gc > n \ ?pstr \ gc") - apply(erule_tac conjE) - apply(simp add: h1) - using h1 - apply(auto) - apply(rule_tac min_max.le_supI2) - apply(rule_tac Max_ge, simp, simp) - apply(rule_tac disjI2) - using h2 - 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 + let ?Q2 = "\nl. nl = xs @ rec_exec f xs # 0 \ (?ft - Suc n) @ x # anything" + have a: "?ft \ fft" + by arith + have b: "far = n" + using compile1 termi_f len + by(drule_tac param_pattern, auto) + have c: "fft > far" + using compile1 + by(simp add: footprint_ge) + show "{?Q1} (?B [+] (?C [+] (?D [+] ?E @ ?F))) {?S}" + proof(rule_tac abc_Hoare_plus_halt) + have "{\nl. nl = xs @ 0 \ (fft - far) @ 0\(?ft - fft) @ x # anything} fap + {\nl. nl = xs @ rec_exec f xs # 0 \ (fft - Suc far) @ 0\(?ft - fft) @ x # anything}" + by(rule_tac f_ind) + moreover have "fft - far + ?ft - fft = ?ft - far" + using a b c by arith + moreover have "fft - Suc n + ?ft - fft = ?ft - Suc n" + using a b c by arith + ultimately show "{?Q1} ?B {?Q2}" + using b + by(simp add: replicate_merge_anywhere) + next + let ?Q3 = "\ nl. nl = xs @ 0 # rec_exec f xs # 0\(?ft - Suc (Suc n)) @ x # anything" + show "{?Q2} (?C [+] (?D [+] ?E @ ?F)) {?S}" + proof(rule_tac abc_Hoare_plus_halt) + show "{?Q2} (?C) {?Q3}" + using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc n) @ x # anything"] + using len + by(auto) + next + show "{?Q3} (?D [+] ?E @ ?F) {?S}" + using pr_loop_correct[of x x xs n g gap gar gft f fft anything] assms + by(simp add: rec_exec_pr_0_simps) + qed + qed qed +qed + +lemma compile_pr_correct: + assumes g_ind: "\y + (\x xa xb. rec_ci g = (x, xa, xb) \ + (\xc. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (xb - xa) @ xc} x + {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (xb - Suc xa) @ xc}))" + and termi_f: "terminate f xs" + and f_ind: + "\ap arity fp anything. + rec_ci f = (ap, arity, fp) \ {\nl. nl = xs @ 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ rec_exec f xs # 0 \ (fp - Suc arity) @ anything}" + and len: "length xs = n" + and compile: "rec_ci (Pr n f g) = (ap, arity, fp)" + shows "{\nl. nl = xs @ x # 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (fp - Suc arity) @ anything}" +proof(case_tac "rec_ci f", case_tac "rec_ci g") + fix fap far fft gap gar gft + assume h: "rec_ci f = (fap, far, fft)" "rec_ci g = (gap, gar, gft)" + have g: + "\y + (\anything. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything} gap + {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything}))" + using g_ind h + by(auto) + hence termi_g: "\ y yanything. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything} gap + {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything})" + by auto + have f_newind: "\ anything. {\nl. nl = xs @ 0 \ (fft - far) @ anything} fap {\nl. nl = xs @ rec_exec f xs # 0 \ (fft - Suc far) @ anything}" + using h + by(rule_tac f_ind, simp) + show "?thesis" + using termi_f termi_g h compile + apply(simp add: rec_ci.simps abc_comp_commute, auto) + using g_newind f_newind len + by(rule_tac compile_pr_correct', simp_all) qed -lemma recursive_compile_correct_spec: - "\rec_ci re = (ap, ary, fp); - rec_calc_rel re args r\ - \ (\ stp. (abc_steps_l (0, args @ 0\(fp - ary)) ap stp) = - (length ap, args@[r]@0\(fp - ary - 1)))" -using recursive_compile_correct[of re ap ary fp args r "[]"] -by simp +fun mn_ind_inv :: + "nat \ nat list \ nat \ nat \ nat \ nat list \ nat list \ bool" + where + "mn_ind_inv (as, lm') ss x rsx suf_lm lm = + (if as = ss then lm' = lm @ x # rsx # suf_lm + else if as = ss + 1 then + \y. (lm' = lm @ x # y # suf_lm) \ y \ rsx + else if as = ss + 2 then + \y. (lm' = lm @ x # y # suf_lm) \ y \ rsx + else if as = ss + 3 then lm' = lm @ x # 0 # suf_lm + else if as = ss + 4 then lm' = lm @ Suc x # 0 # suf_lm + else if as = 0 then lm' = lm @ Suc x # 0 # suf_lm + else False +)" + +fun mn_stage1 :: "nat \ nat list \ nat \ nat \ nat" + where + "mn_stage1 (as, lm) ss n = + (if as = 0 then 0 + else if as = ss + 4 then 1 + else if as = ss + 3 then 2 + else if as = ss + 2 \ as = ss + 1 then 3 + else if as = ss then 4 + else 0 +)" + +fun mn_stage2 :: "nat \ nat list \ nat \ nat \ nat" + where + "mn_stage2 (as, lm) ss n = + (if as = ss + 1 \ as = ss + 2 then (lm ! (Suc n)) + else 0)" + +fun mn_stage3 :: "nat \ nat list \ nat \ nat \ nat" + where + "mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)" + + +fun mn_measure :: "((nat \ nat list) \ nat \ nat) \ + (nat \ nat \ nat)" + where + "mn_measure ((as, lm), ss, n) = + (mn_stage1 (as, lm) ss n, mn_stage2 (as, lm) ss n, + mn_stage3 (as, lm) ss n)" + +definition mn_LE :: "(((nat \ nat list) \ nat \ nat) \ + ((nat \ nat list) \ nat \ nat)) set" + where "mn_LE \ (inv_image lex_triple mn_measure)" + +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 [simp]: + "0 < rsx \ + \y. (xs @ x # rsx # anything)[Suc (length xs) := rsx - Suc 0] = xs @ x # y # anything \ y \ rsx" +apply(rule_tac x = "rsx - 1" in exI) +apply(simp add: list_update_append list_update.simps) +done + +lemma [simp]: + "\y \ rsx; 0 < y\ + \ \ya. (xs @ x # y # anything)[Suc (length xs) := y - Suc 0] = xs @ x # ya # anything \ ya \ rsx" +apply(rule_tac x = "y - 1" in exI) +apply(simp add: list_update_append list_update.simps) +done + +lemma abc_comp_null[simp]: "(A [+] B = []) = (A = [] \ B = [])" +by(auto simp: abc_comp.simps abc_shift.simps) + +lemma rec_ci_not_null[simp]: "(rec_ci f \ ([], a, b))" +apply(case_tac f, auto simp: rec_ci_z_def rec_ci_s_def rec_ci.simps addition.simps rec_ci_id.simps) +apply(case_tac "rec_ci recf", auto simp: mv_box.simps) +apply(case_tac "rec_ci recf1", case_tac "rec_ci recf2", simp) +apply(case_tac "rec_ci recf", simp) +done + +lemma mn_correct: + assumes compile: "rec_ci rf = (fap, far, fft)" + and ge: "0 < rsx" + and len: "length xs = arity" + and B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" + and f: "f = (\ stp. (abc_steps_l (length fap, xs @ x # rsx # anything) (fap @ B) stp, (length fap), arity)) " + and P: "P =(\ ((as, lm), ss, arity). as = 0)" + and Q: "Q = (\ ((as, lm), ss, arity). mn_ind_inv (as, lm) (length fap) x rsx anything xs)" + shows "\ stp. P (f stp) \ Q (f stp)" +proof(rule_tac halt_lemma2) + show "wf mn_LE" + using wf_mn_le by simp +next + show "Q (f 0)" + by(auto simp: Q f abc_steps_l.simps mn_ind_inv.simps) +next + have "fap \ []" + using compile by auto + thus "\ P (f 0)" + by(auto simp: f P abc_steps_l.simps) +next + have "fap \ []" + using compile by auto + then have "\ stp. \\ P (f stp); Q (f stp)\ \ Q (f (Suc stp)) \ (f (Suc stp), f stp) \ mn_LE" + using ge len + apply(case_tac "(abc_steps_l (length fap, xs @ x # rsx # anything) (fap @ B) stp)") + apply(simp add: abc_step_red2 B f P Q) + apply(auto split:if_splits simp add:abc_steps_l.simps mn_ind_inv.simps abc_steps_zero B abc_fetch.simps nth_append) + by(auto simp: 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 abc_fetch.simps + split: if_splits) + thus "\stp. \ P (f stp) \ Q (f stp) \ Q (f (Suc stp)) \ (f (Suc stp), f stp) \ mn_LE" + by(auto) +qed + +lemma abc_Hoare_haltE: + "{\ nl. nl = lm1} p {\ nl. nl = lm2} + \ \ stp. abc_steps_l (0, lm1) p stp = (length p, lm2)" +apply(auto simp: abc_Hoare_halt_def) +apply(rule_tac x = n in exI) +apply(case_tac "abc_steps_l (0, lm1) p n", auto) +done + +lemma mn_loop: + assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" + and ft: "ft = max (Suc arity) fft" + and len: "length xs = arity" + and far: "far = Suc arity" + and ind: " (\xc. ({\nl. nl = xs @ x # 0 \ (fft - far) @ xc} fap + {\nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \ (fft - Suc far) @ xc}))" + and exec_less: "rec_exec f (xs @ [x]) > 0" + and compile: "rec_ci f = (fap, far, fft)" + shows "\ stp > 0. abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = + (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" +proof - + have "\ stp. abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = + (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything)" + proof - + have "\ stp. abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) fap stp = + (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything)" + proof - + have "{\nl. nl = xs @ x # 0 \ (fft - far) @ 0\(ft - fft) @ anything} fap + {\nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \ (fft - Suc far) @ 0\(ft - fft) @ anything}" + using ind by simp + moreover have "fft > far" + using compile + by(erule_tac footprint_ge) + ultimately show "?thesis" + using ft far + apply(drule_tac abc_Hoare_haltE) + by(simp add: replicate_merge_anywhere max_absorb2) + qed + then obtain stp where "abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) fap stp = + (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything)" .. + thus "?thesis" + by(erule_tac abc_append_frist_steps_halt_eq) + qed + moreover have + "\ stp > 0. abc_steps_l (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything) (fap @ B) stp = + (0, xs @ Suc x # 0 # 0 \ (ft - Suc (Suc arity)) @ anything)" + using mn_correct[of f fap far fft "rec_exec f (xs @ [x])" xs arity B + "(\stp. (abc_steps_l (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything) (fap @ B) stp, length fap, arity))" + x "0 \ (ft - Suc (Suc arity)) @ anything" "(\((as, lm), ss, arity). as = 0)" + "(\((as, lm), ss, aritya). mn_ind_inv (as, lm) (length fap) x (rec_exec f (xs @ [x])) (0 \ (ft - Suc (Suc arity)) @ anything) xs) "] + B compile exec_less len + apply(subgoal_tac "fap \ []", auto) + apply(rule_tac x = stp in exI, auto simp: mn_ind_inv.simps) + by(case_tac "stp = 0", simp_all add: abc_steps_l.simps) + moreover have "fft > far" + using compile + by(erule_tac footprint_ge) + ultimately show "?thesis" + using ft far + apply(auto) + by(rule_tac x = "stp + stpa" in exI, + simp add: abc_steps_add replicate_Suc[THEN sym] diff_Suc_Suc del: replicate_Suc) +qed + +lemma mn_loop_correct': + assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" + and ft: "ft = max (Suc arity) fft" + and len: "length xs = arity" + and ind_all: "\i\x. (\xc. ({\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap + {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}))" + and exec_ge: "\ i\x. rec_exec f (xs @ [i]) > 0" + and compile: "rec_ci f = (fap, far, fft)" + and far: "far = Suc arity" + shows "\ stp > x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = + (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" +using ind_all exec_ge +proof(induct x) + case 0 + thus "?case" + using assms + by(rule_tac mn_loop, simp_all) +next + case (Suc x) + have ind': "\\i\x. \xc. {\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}; + \i\x. 0 < rec_exec f (xs @ [i])\ \ + \stp > x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" by fact + have exec_ge: "\i\Suc x. 0 < rec_exec f (xs @ [i])" by fact + have ind_all: "\i\Suc x. \xc. {\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap + {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}" by fact + have ind: "\stp > x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = + (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" using ind' exec_ge ind_all by simp + have stp: "\ stp > 0. abc_steps_l (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = + (0, xs @ Suc (Suc x) # 0 \ (ft - Suc arity) @ anything)" + using ind_all exec_ge B ft len far compile + by(rule_tac mn_loop, simp_all) + from ind stp show "?case" + apply(auto) + by(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add) +qed + +lemma mn_loop_correct: + assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" + and ft: "ft = max (Suc arity) fft" + and len: "length xs = arity" + and ind_all: "\i\x. (\xc. ({\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap + {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}))" + and exec_ge: "\ i\x. rec_exec f (xs @ [i]) > 0" + and compile: "rec_ci f = (fap, far, fft)" + and far: "far = Suc arity" + shows "\ stp. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = + (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" +proof - + have "\stp>x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" + using assms + by(rule_tac mn_loop_correct', simp_all) + thus "?thesis" + by(auto) +qed + +lemma compile_mn_correct': + assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" + and ft: "ft = max (Suc arity) fft" + and len: "length xs = arity" + and termi_f: "terminate f (xs @ [r])" + and f_ind: "\anything. {\nl. nl = xs @ r # 0 \ (fft - far) @ anything} fap + {\nl. nl = xs @ r # 0 # 0 \ (fft - Suc far) @ anything}" + and ind_all: "\i < r. (\xc. ({\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap + {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}))" + and exec_less: "\ i 0" + and exec: "rec_exec f (xs @ [r]) = 0" + and compile: "rec_ci f = (fap, far, fft)" + shows "{\nl. nl = xs @ 0 \ (max (Suc arity) fft - arity) @ anything} + fap @ B + {\nl. nl = xs @ rec_exec (Mn arity f) xs # 0 \ (max (Suc arity) fft - Suc arity) @ anything}" +proof - + have a: "far = Suc arity" + using len compile termi_f + by(drule_tac param_pattern, auto) + have b: "fft > far" + using compile + by(erule_tac footprint_ge) + have "\ stp. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = + (0, xs @ r # 0 \ (ft - Suc arity) @ anything)" + using assms a + apply(case_tac r, rule_tac x = 0 in exI, simp add: abc_steps_l.simps, simp) + by(rule_tac mn_loop_correct, auto) + moreover have + "\ stp. abc_steps_l (0, xs @ r # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = + (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" + proof - + have "\ stp. abc_steps_l (0, xs @ r # 0 \ (ft - Suc arity) @ anything) fap stp = + (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" + proof - + have "{\nl. nl = xs @ r # 0 \ (fft - far) @ 0\(ft - fft) @ anything} fap + {\nl. nl = xs @ r # rec_exec f (xs @ [r]) # 0 \ (fft - Suc far) @ 0\(ft - fft) @ anything}" + using f_ind exec by simp + thus "?thesis" + using ft a b + apply(drule_tac abc_Hoare_haltE) + by(simp add: replicate_merge_anywhere max_absorb2) + qed + then obtain stp where "abc_steps_l (0, xs @ r # 0 \ (ft - Suc arity) @ anything) fap stp = + (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" .. + thus "?thesis" + by(erule_tac abc_append_frist_steps_halt_eq) + qed + moreover have + "\ stp. abc_steps_l (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything) (fap @ B) stp = + (length fap + 5, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" + using ft a b len B exec + apply(rule_tac x = 1 in exI, auto) + by(auto simp: abc_steps_l.simps B abc_step_l.simps + abc_fetch.simps nth_append max_absorb2 abc_lm_v.simps abc_lm_s.simps) + moreover have "rec_exec (Mn (length xs) f) xs = r" + using exec exec_less + apply(auto simp: rec_exec.simps Least_def) + thm the_equality + apply(rule_tac the_equality, auto) + apply(metis exec_less less_not_refl3 linorder_not_less) + by (metis le_neq_implies_less less_not_refl3) + ultimately show "?thesis" + using ft a b len B exec + apply(auto simp: abc_Hoare_halt_def) + apply(rule_tac x = "stp + stpa + stpb" in exI) + by(simp add: abc_steps_add replicate_Suc_iff_anywhere max_absorb2 Suc_diff_Suc del: replicate_Suc) +qed + +lemma compile_mn_correct: + assumes len: "length xs = n" + and termi_f: "terminate f (xs @ [r])" + and f_ind: "\ap arity fp anything. rec_ci f = (ap, arity, fp) \ + {\nl. nl = xs @ r # 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ r # 0 # 0 \ (fp - Suc arity) @ anything}" + and exec: "rec_exec f (xs @ [r]) = 0" + and ind_all: + "\i + (\x xa xb. rec_ci f = (x, xa, xb) \ + (\xc. {\nl. nl = xs @ i # 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (xb - Suc xa) @ xc})) \ + 0 < rec_exec f (xs @ [i])" + and compile: "rec_ci (Mn n f) = (ap, arity, fp)" + shows "{\nl. nl = xs @ 0 \ (fp - arity) @ anything} ap + {\nl. nl = xs @ rec_exec (Mn n f) xs # 0 \ (fp - Suc arity) @ anything}" +proof(case_tac "rec_ci f") + fix fap far fft + assume h: "rec_ci f = (fap, far, fft)" + hence f_newind: + "\anything. {\nl. nl = xs @ r # 0 \ (fft - far) @ anything} fap + {\nl. nl = xs @ r # 0 # 0 \ (fft - Suc far) @ anything}" + by(rule_tac f_ind, simp) + have newind_all: + "\i < r. (\xc. ({\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap + {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}))" + using ind_all h + by(auto) + have all_less: "\ i 0" + using ind_all by auto + show "?thesis" + using h compile f_newind newind_all all_less len termi_f exec + apply(auto simp: rec_ci.simps) + by(rule_tac compile_mn_correct', auto) +qed + +lemma recursive_compile_correct: + "\terminate recf args; rec_ci recf = (ap, arity, fp)\ + \ {\ nl. nl = args @ 0\(fp - arity) @ anything} ap + {\ nl. nl = args@ rec_exec recf args # 0\(fp - Suc arity) @ anything}" +apply(induct arbitrary: ap arity fp anything r rule: terminate.induct) +apply(simp_all add: compile_s_correct compile_z_correct compile_id_correct + compile_cn_correct compile_pr_correct compile_mn_correct) +done definition dummy_abc :: "nat \ abc_inst list" where @@ -4449,84 +2217,51 @@ where "abc_list_crsp xs ys = (\ n. xs = ys @ 0\n \ ys = xs @ 0\n)" -lemma [intro]: "abc_list_crsp (lm @ 0\m) lm" -apply(auto simp: abc_list_crsp_def) -done +lemma abc_list_crsp_simp1[intro]: "abc_list_crsp (lm @ 0\m) lm" +by(auto simp: abc_list_crsp_def) + lemma abc_list_crsp_lm_v: "abc_list_crsp lma lmb \ abc_lm_v lma n = abc_lm_v lmb n" -apply(auto simp: abc_list_crsp_def abc_lm_v.simps +by(auto simp: abc_list_crsp_def abc_lm_v.simps nth_append) + + +lemma abc_list_crsp_elim: + "\abc_list_crsp lma lmb; \ n. lma = lmb @ 0\n \ lmb = lma @ 0\n \ P \ \ P" +by(auto simp: abc_list_crsp_def) + +lemma [simp]: + "\abc_list_crsp lma lmb; m < length lma; m < length lmb\ \ + abc_list_crsp (lma[m := n]) (lmb[m := n])" +by(auto simp: abc_list_crsp_def list_update_append) + +lemma [simp]: + "\abc_list_crsp lma lmb; m < length lma; \ m < length lmb\ \ + abc_list_crsp (lma[m := n]) (lmb @ 0 \ (m - length lmb) @ [n])" +apply(auto simp: abc_list_crsp_def list_update_append) +apply(rule_tac x = "na + length lmb - Suc m" in exI) +apply(rule_tac disjI1) +apply(simp add: upd_conv_take_nth_drop min_absorb1) done -lemma rep_app_cons_iff: - "k < n \ replicate n a[k:=b] = - replicate k a @ b # replicate (n - k - 1) a" -apply(induct n arbitrary: k, simp) -apply(simp split:nat.splits) +lemma [simp]: + "\abc_list_crsp lma lmb; \ m < length lma; m < length lmb\ \ + abc_list_crsp (lma @ 0 \ (m - length lma) @ [n]) (lmb[m := n])" +apply(auto simp: abc_list_crsp_def list_update_append) +apply(rule_tac x = "na + length lma - Suc m" in exI) +apply(rule_tac disjI2) +apply(simp add: upd_conv_take_nth_drop min_absorb1) done +lemma [simp]: "\abc_list_crsp lma lmb; \ m < length lma; \ m < length lmb\ \ + abc_list_crsp (lma @ 0 \ (m - length lma) @ [n]) (lmb @ 0 \ (m - length lmb) @ [n])" + by(auto simp: abc_list_crsp_def list_update_append replicate_merge_anywhere) + lemma abc_list_crsp_lm_s: "abc_list_crsp lma lmb \ abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)" -apply(auto simp: abc_list_crsp_def abc_lm_v.simps abc_lm_s.simps) -apply(simp_all add: list_update_append, auto) -proof - - fix na - assume h: "m < length lmb + na" " \ m < length lmb" - hence "m - length lmb < na" by simp - hence "replicate na 0[(m- length lmb):= n] = - replicate (m - length lmb) 0 @ n # - replicate (na - (m - length lmb) - 1) 0" - apply(erule_tac rep_app_cons_iff) - done - thus "\nb. replicate na 0[m - length lmb := n] = - replicate (m - length lmb) 0 @ n # replicate nb 0 \ - replicate (m - length lmb) 0 @ [n] = - replicate na 0[m - length lmb := n] @ replicate nb 0" - apply(auto) - done -next - fix na - assume h: "\ m < length lmb + na" - show - "\nb. replicate na 0 @ replicate (m - (length lmb + na)) 0 @ [n] = - replicate (m - length lmb) 0 @ n # replicate nb 0 \ - replicate (m - length lmb) 0 @ [n] = - replicate na 0 @ - replicate (m - (length lmb + na)) 0 @ n # replicate nb 0" - apply(rule_tac x = 0 in exI, simp, auto) - using h - apply(simp add: replicate_add[THEN sym]) - done -next - fix na - assume h: "\ m < length lma" "m < length lma + na" - hence "m - length lma < na" by simp - hence - "replicate na 0[(m- length lma):= n] = replicate (m - length lma) - 0 @ n # replicate (na - (m - length lma) - 1) 0" - apply(erule_tac rep_app_cons_iff) - done - thus "\nb. replicate (m - length lma) 0 @ [n] = - replicate na 0[m - length lma := n] @ replicate nb 0 - \ replicate na 0[m - length lma := n] = - replicate (m - length lma) 0 @ n # replicate nb 0" - apply(auto) - done -next - fix na - assume "\ m < length lma + na" - thus " \nb. replicate (m - length lma) 0 @ [n] = - replicate na 0 @ - replicate (m - (length lma + na)) 0 @ n # replicate nb 0 - \ replicate na 0 @ - replicate (m - (length lma + na)) 0 @ [n] = - replicate (m - length lma) 0 @ n # replicate nb 0" - apply(rule_tac x = 0 in exI, simp, auto) - apply(simp add: replicate_add[THEN sym]) - done -qed +by(auto simp: abc_lm_s.simps) lemma abc_list_crsp_step: "\abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); @@ -4577,210 +2312,230 @@ done qed -lemma recursive_compile_correct_norm: - "\rec_ci re = (aprog, rs_pos, a_md); - rec_calc_rel re lm rs\ - \ (\ stp lm' m. (abc_steps_l (0, lm) aprog stp) = - (length aprog, lm') \ abc_list_crsp lm' (lm @ rs # 0\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 list_crsp_simp2: "abc_list_crsp (lm1 @ 0\n) lm2 \ abc_list_crsp lm1 lm2" +proof(induct n) + case 0 + thus "?case" + by(auto simp: abc_list_crsp_def) +next + case (Suc n) + have ind: "abc_list_crsp (lm1 @ 0 \ n) lm2 \ abc_list_crsp lm1 lm2" by fact + have h: "abc_list_crsp (lm1 @ 0 \ Suc n) lm2" by fact + then have "abc_list_crsp (lm1 @ 0 \ n) lm2" + apply(auto simp: exp_suc abc_list_crsp_def del: replicate_Suc) + apply(case_tac n, simp_all add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc, auto) + apply(rule_tac x = 1 in exI, simp) + by(rule_tac x = "Suc n" in exI, simp, simp add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc) + thus "?case" + using ind + by auto +qed -lemma [simp]: "length (dummy_abc (length lm)) = 3" -apply(simp add: dummy_abc_def) -done - -lemma [simp]: "dummy_abc (length lm) \ []" -apply(simp add: dummy_abc_def) +lemma recursive_compile_correct_norm': + "\rec_ci f = (ap, arity, ft); + terminate f args\ + \ \ stp rl. (abc_steps_l (0, args) ap stp) = (length ap, rl) \ abc_list_crsp (args @ [rec_exec f args]) rl" + using recursive_compile_correct[of f args ap arity ft "[]"] +apply(auto simp: abc_Hoare_halt_def) +apply(rule_tac x = n in exI) +apply(case_tac "abc_steps_l (0, args @ 0 \ (ft - arity)) ap n", auto) +apply(drule_tac abc_list_crsp_steps, auto) +apply(rule_tac list_crsp_simp2, auto) done -lemma dummy_abc_steps_ex: - "\bstp. abc_steps_l (0, lm') (dummy_abc (length lm)) bstp = - ((Suc (Suc (Suc 0))), abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)))" -apply(rule_tac x = "Suc (Suc (Suc 0))" in exI) -apply(auto simp: abc_steps_l.simps abc_step_l.simps - dummy_abc_def abc_fetch.simps) -apply(auto simp: abc_lm_s.simps abc_lm_v.simps nth_append) -apply(simp add: butlast_append) -done - -lemma [simp]: - "\Suc (length lm) - length lm' \ n; \ length lm < length lm'; lm @ rs # 0 \ m = lm' @ 0 \ n\ - \ lm' @ 0 \ Suc (length lm - length lm') = lm @ [rs]" -apply(subgoal_tac "n > m") -apply(subgoal_tac "\ 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) +lemma [simp]: + assumes a: "args @ [rec_exec f args] = lm @ 0 \ n" + and b: "length args < length lm" + shows "\m. lm = args @ rec_exec f args # 0 \ m" +using assms +apply(case_tac n, simp) +apply(rule_tac x = 0 in exI, simp) apply(drule_tac length_equal, simp) done -lemma [elim]: - "lm @ rs # 0\m = lm' @ 0\n \ - \m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) = - lm @ rs # 0\m" -proof(cases "length lm' > length lm") - case True - assume h: "lm @ rs # 0\m = lm' @ 0\n" "length lm < length lm'" - hence "m \ n" - apply(drule_tac length_equal) - apply(simp) - done - hence "\ d. m = d + n" - apply(rule_tac x = "m - n" in exI, simp) - done - from this obtain d where "m = d + n" .. - from h and this show "?thesis" - apply(auto simp: abc_lm_s.simps abc_lm_v.simps - replicate_add) - done +lemma [simp]: +"\args @ [rec_exec f args] = lm @ 0 \ n; \ length args < length lm\ + \ \m. (lm @ 0 \ (length args - length lm) @ [Suc 0])[length args := 0] = + args @ rec_exec f args # 0 \ m" +apply(case_tac n, simp_all add: exp_suc list_update_append list_update.simps del: replicate_Suc) +apply(subgoal_tac "length args = Suc (length lm)", simp) +apply(rule_tac x = "Suc (Suc 0)" in exI, simp) +apply(drule_tac length_equal, simp, auto) +done + +lemma compile_append_dummy_correct: + assumes compile: "rec_ci f = (ap, ary, fp)" + and termi: "terminate f args" + shows "{\ nl. nl = args} (ap [+] dummy_abc (length args)) {\ nl. (\ m. nl = args @ rec_exec f args # 0\m)}" +proof(rule_tac abc_Hoare_plus_halt) + show "{\nl. nl = args} ap {\ nl. abc_list_crsp (args @ [rec_exec f args]) nl}" + using compile termi recursive_compile_correct_norm'[of f ap ary fp args] + apply(auto simp: abc_Hoare_halt_def) + by(rule_tac x = stp in exI, simp) next - case False - assume h:"lm @ rs # 0\m = lm' @ 0\n" - and g: "\ length lm < length lm'" - have "take (Suc (length lm)) (lm @ rs # 0\m) = - take (Suc (length lm)) (lm' @ 0\n)" - using h by simp - moreover have "n \ (Suc (length lm) - length lm')" - using h g - apply(drule_tac length_equal) - apply(simp) - done - ultimately show - "\m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) = - lm @ rs # 0\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 + show "{abc_list_crsp (args @ [rec_exec f args])} dummy_abc (length args) + {\nl. \m. nl = args @ rec_exec f args # 0 \ m}" + apply(auto simp: dummy_abc_def abc_Hoare_halt_def) + apply(rule_tac x = 3 in exI) + by(auto simp: abc_steps_l.simps abc_list_crsp_def abc_step_l.simps numeral_3_eq_3 abc_fetch.simps + abc_lm_v.simps nth_append abc_lm_s.simps) qed -lemma [elim]: - "abc_list_crsp lm' (lm @ rs # 0\m) - \ \m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) - = lm @ rs # 0\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) +lemma cn_merge_gs_split: + "\i < length gs; rec_ci (gs!i) = (ga, gb, gc)\ \ + cn_merge_gs (map rec_ci gs) p = cn_merge_gs (map rec_ci (take i gs)) p [+] (ga [+] + 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_comp_commute[THEN sym]) done -lemma abc_append_dummy_complie: - "\rec_ci recf = (ap, ary, fp); - rec_calc_rel recf args r; - length args = k\ - \ (\ stp m. (abc_steps_l (0, args) (ap [+] dummy_abc k) stp) = - (length ap + 3, args @ r # 0\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\m)" - have "\stp. abc_steps_l (0, args) (ap [+] - (dummy_abc (length args))) stp = (length ap + 3, - abc_lm_s lm' (length args) (abc_lm_v lm' (length args)))" - using h - apply(rule_tac bm = lm' in abc_append_exc2, - auto intro: dummy_abc_steps_ex simp: numeral_3_eq_3) - done - thus "\stp m. abc_steps_l (0, args) (ap [+] - dummy_abc (length args)) stp = (Suc (Suc (Suc (length ap))), args @ r # 0\m)" - using h - apply(erule_tac exE) - apply(rule_tac x = stpa in exI, auto) - done +lemma cn_unhalt_case: + assumes compile1: "rec_ci (Cn n f gs) = (ap, ar, ft) \ length args = ar" + and g: "i < length gs" + and compile2: "rec_ci (gs!i) = (gap, gar, gft) \ gar = length args" + and g_unhalt: "\ anything. {\ nl. nl = args @ 0\(gft - gar) @ anything} gap \" + and g_ind: "\ apj arj ftj j anything. \j < i; rec_ci (gs!j) = (apj, arj, ftj)\ + \ {\ nl. nl = args @ 0\(ftj - arj) @ anything} apj {\ nl. nl = args @ rec_exec (gs!j) args # 0\(ftj - Suc arj) @ anything}" + and all_termi: "\ j nl. nl = args @ 0\(ft - ar) @ anything} ap \" +using compile1 +apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_comp_commute) +proof(rule_tac abc_Hoare_plus_unhalt1) + fix fap far fft + let ?ft = "max (Suc (length args)) (Max (insert fft ((\(aprog, p, n). n) ` rec_ci ` set gs)))" + let ?Q = "\nl. nl = args @ 0\ (?ft - length args) @ map (\i. rec_exec i args) (take i gs) @ + 0\(length gs - i) @ 0\ Suc (length args) @ anything" + have "cn_merge_gs (map rec_ci gs) ?ft = + cn_merge_gs (map rec_ci (take i gs)) ?ft [+] (gap [+] + mv_box gar (?ft + i)) [+] cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?ft + Suc i)" + using g compile2 cn_merge_gs_split by simp + thus "{\nl. nl = args @ 0 # 0 \ (?ft + length gs) @ anything} (cn_merge_gs (map rec_ci gs) ?ft) \" + proof(simp, rule_tac abc_Hoare_plus_unhalt1, rule_tac abc_Hoare_plus_unhalt2, + rule_tac abc_Hoare_plus_unhalt1) + let ?Q_tmp = "\nl. nl = args @ 0\ (gft - gar) @ 0\(?ft - (length args) - (gft -gar)) @ map (\i. rec_exec i args) (take i gs) @ + 0\(length gs - i) @ 0\ Suc (length args) @ anything" + have a: "{?Q_tmp} gap \" + using g_unhalt[of "0 \ (?ft - (length args) - (gft - gar)) @ + map (\i. rec_exec i args) (take i gs) @ 0 \ (length gs - i) @ 0 \ Suc (length args) @ anything"] + by simp + moreover have "?ft \ gft" + using g compile2 + apply(rule_tac min_max.le_supI2, rule_tac Max_ge, simp, rule_tac insertI2) + apply(rule_tac x = "rec_ci (gs ! i)" in image_eqI, simp) + by(rule_tac x = "gs!i" in image_eqI, simp, simp) + then have b:"?Q_tmp = ?Q" + using compile2 + apply(rule_tac arg_cong) + by(simp add: replicate_merge_anywhere) + thus "{?Q} gap \" + using a by simp + next + show "{\nl. nl = args @ 0 # 0 \ (?ft + length gs) @ anything} + cn_merge_gs (map rec_ci (take i gs)) ?ft + {\nl. nl = args @ 0 \ (?ft - length args) @ + map (\i. rec_exec i args) (take i gs) @ 0 \ (length gs - i) @ 0 \ Suc (length args) @ anything}" + using all_termi + apply(rule_tac compile_cn_gs_correct', auto simp: set_conv_nth) + by(drule_tac apj = x and arj = xa and ftj = xb and j = ia and anything = xc in g_ind, auto) + qed qed -lemma [simp]: "length (dummy_abc k) = 3" -apply(simp add: dummy_abc_def) -done + -lemma [simp]: "length args = k \ abc_lm_v (args @ r # 0\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, @ Bk \ rn) ires" -apply(auto simp: crsp.simps start_of.simps) -done - -(* cccc *) +lemma mn_unhalt_case': + assumes compile: "rec_ci f = (a, b, c)" + and all_termi: "\i. terminate f (args @ [i]) \ 0 < rec_exec f (args @ [i])" + and B: "B = [Dec (Suc (length args)) (length a + 5), Dec (Suc (length args)) (length a + 3), + Goto (Suc (length a)), Inc (length args), Goto 0]" + shows "{\nl. nl = args @ 0 \ (max (Suc (length args)) c - length args) @ anything} + a @ B \" +proof(rule_tac abc_Hoare_unhaltI, auto) + fix n + have a: "b = Suc (length args)" + using all_termi compile + apply(erule_tac x = 0 in allE) + by(auto, drule_tac param_pattern,auto) + moreover have b: "c > b" + using compile by(elim footprint_ge) + ultimately have c: "max (Suc (length args)) c = c" by arith + have "\ stp > n. abc_steps_l (0, args @ 0 # 0\(c - Suc (length args)) @ anything) (a @ B) stp + = (0, args @ Suc n # 0\(c - Suc (length args)) @ anything)" + using assms a b c + proof(rule_tac mn_loop_correct', auto) + fix i xc + show "{\nl. nl = args @ i # 0 \ (c - Suc (length args)) @ xc} a + {\nl. nl = args @ i # rec_exec f (args @ [i]) # 0 \ (c - Suc (Suc (length args))) @ xc}" + using all_termi recursive_compile_correct[of f "args @ [i]" a b c xc] compile a + by(simp) + qed + then obtain stp where d: "stp > n \ abc_steps_l (0, args @ 0 # 0\(c - Suc (length args)) @ anything) (a @ B) stp + = (0, args @ Suc n # 0\(c - Suc (length args)) @ anything)" .. + then obtain d where e: "stp = n + Suc d" + by (metis add_Suc_right less_iff_Suc_add) + obtain s nl where f: "abc_steps_l (0, args @ 0 # 0\(c - Suc (length args)) @ anything) (a @ B) n = (s, nl)" + by (metis prod.exhaust) + have g: "s < length (a @ B)" + using d e f + apply(rule_tac classical, simp only: abc_steps_add) + by(simp add: halt_steps2 leI) + from f g show "abc_notfinal (abc_steps_l (0, args @ 0 \ + (max (Suc (length args)) c - length args) @ anything) (a @ B) n) (a @ B)" + using c b a + by(simp add: replicate_Suc_iff_anywhere Suc_diff_Suc del: replicate_Suc) +qed + +lemma mn_unhalt_case: + assumes compile: "rec_ci (Mn n f) = (ap, ar, ft) \ length args = ar" + and all_term: "\ i. terminate f (args @ [i]) \ rec_exec f (args @ [i]) > 0" + shows "{\ nl. nl = args @ 0\(ft - ar) @ anything} ap \ " + using assms +apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_comp_commute) +by(rule_tac mn_unhalt_case', simp_all) fun tm_of_rec :: "recf \ instr list" where "tm_of_rec recf = (let (ap, k, fp) = rec_ci recf in let tp = tm_of (ap [+] dummy_abc k) in - tp @ (shift (mopup k) (length tp div 2)))" + tp @ (shift (mopup k) (length tp div 2)))" -lemma recursive_compile_to_tm_correct: - "\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)\ - \ \ stp m l. steps0 (Suc 0, Bk # Bk # ires, @ Bk\rn) - (tp @ shift (mopup k) (length tp div 2)) stp - = (0, Bk\m @ Bk # Bk # ires, Oc\Suc r @ Bk\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 recursive_compile_to_tm_correct1: + assumes compile: "rec_ci recf = (ap, ary, fp)" + and termi: " terminate recf args" + and tp: "tp = tm_of (ap [+] dummy_abc (length args))" + shows "\ stp m l. steps0 (Suc 0, Bk # Bk # ires, @ Bk\rn) + (tp @ shift (mopup (length args)) (length tp div 2)) stp = (0, Bk\m @ Bk # Bk # ires, Oc\Suc (rec_exec recf args) @ Bk\l)" +proof - + have "{\nl. nl = args} ap [+] dummy_abc (length args) {\nl. \m. nl = args @ rec_exec recf args # 0 \ m}" + using compile termi compile + by(rule_tac compile_append_dummy_correct, auto) + then obtain stp m where h: "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) stp = + (length (ap [+] dummy_abc (length args)), args @ rec_exec recf args # 0\m) " + apply(simp add: abc_Hoare_halt_def, auto) + by(case_tac "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) n", auto) + thus "?thesis" + using assms tp + by(rule_tac lm = args and stp = stp and am = "args @ rec_exec recf args # 0 \ m" + in compile_correct_halt, auto simp: crsp.simps start_of.simps length_abc_comp abc_lm_v.simps) +qed lemma recursive_compile_to_tm_correct2: - assumes "rec_ci recf = (ap, ary, fp)" - and "rec_calc_rel recf args r" - and "length args = k" - and "tp = tm_of (ap [+] dummy_abc k)" - shows "\ m n. {\tp. tp = ([Bk, Bk], )} - (tp @ (shift (mopup k) (length tp div 2))) - {\tp. tp = (Bk \ m, Oc \ (Suc r) @ Bk \ n)}" -using recursive_compile_to_tm_correct[where ires="[]" and rn="0", OF assms(1-3) _ assms(4)] -apply(simp add: Hoare_halt_def) -apply(drule_tac x="layout_of (ap [+] dummy_abc k)" in meta_spec) -apply(auto) -apply(rule_tac x="m + 2" in exI) -apply(rule_tac x="l" in exI) -apply(rule_tac x="stp" in exI) -apply(auto) -by (metis append_Nil2 replicate_app_Cons_same) + assumes compile: "rec_ci recf = (ap, ary, fp)" + and termi: " terminate recf args" + shows "\ stp m l. steps0 (Suc 0, [Bk, Bk], ) (tm_of_rec recf) stp = + (0, Bk\Suc (Suc m), Oc\Suc (rec_exec recf args) @ Bk\l)" +using recursive_compile_to_tm_correct1[of recf ap ary fp args "tm_of (ap [+] dummy_abc (length args))" "[]" 0] +assms param_pattern[of recf args ap ary fp] +by(simp add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc tm_of_rec_def) lemma recursive_compile_to_tm_correct3: - assumes "rec_calc_rel recf args r" - shows "{\tp. tp = ([Bk, Bk], )} tm_of_rec recf {\tp. \k l. tp = (Bk \ k, @ Bk \ l)}" -using recursive_compile_to_tm_correct2[OF _ assms] -apply(auto) -apply(case_tac "rec_ci recf") -apply(auto) -apply(drule_tac x="a" in meta_spec) -apply(drule_tac x="b" in meta_spec) -apply(drule_tac x="c" in meta_spec) -apply(drule_tac x="length args" in meta_spec) -apply(drule_tac x="tm_of (a [+] dummy_abc (length args))" in meta_spec) -apply(auto) -apply(simp add: tape_of_nat_abv) -apply(subgoal_tac "b = length args") -apply(simp add: Hoare_halt_def) -apply(auto)[1] -apply(rule_tac x="na" in exI) -apply(auto)[1] -apply(case_tac "steps0 (Suc 0, [Bk, Bk], ) - (tm_of (a [+] dummy_abc (length args)) @ - shift (mopup (length args)) - (listsum - (layout_of (a [+] dummy_abc (length args))))) - na") -apply(simp) -by (metis assms para_pattern) - + assumes compile: "rec_ci recf = (ap, ary, fp)" + and termi: "terminate recf args" + shows "{\ (l, r). l = [Bk, Bk] \ r = } (tm_of_rec recf) + {\ (l, r). \ m i. l = Bk\(Suc (Suc m)) \ r = Oc\Suc (rec_exec recf args) @ Bk \ i}" +using recursive_compile_to_tm_correct2[of recf ap ary fp args] assms +apply(simp add: Hoare_halt_def, auto) +apply(rule_tac x = stp in exI, simp) +done lemma [simp]: "list_all (\(acn, s). s \ Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))) xs \ @@ -4908,12 +2663,13 @@ apply(rule_tac start_of_all_le) done -lemma concat_in: "i < length (concat xs) \ \k < length xs. concat xs ! i \ set (xs ! k)" -apply(induct xs rule: list_tl_induct, simp, simp) -apply(case_tac "i < length (concat list)", simp) +lemma concat_in: "i < length (concat xs) \ + \k < length xs. concat xs ! i \ set (xs ! k)" +apply(induct xs rule: rev_induct, simp, simp) +apply(case_tac "i < length (concat xs)", 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(rule_tac x = "length xs" in exI, simp) apply(simp add: nth_append) done @@ -4950,9 +2706,9 @@ done lemma [intro]: "listsum (map (length \ (\(x, y). ci ly x y)) zs) mod 2 = 0" -apply(induct zs rule: list_tl_induct, simp) -apply(case_tac a, simp) -apply(subgoal_tac "length (ci ly aa b) mod 2 = 0") +apply(induct zs rule: rev_induct, simp) +apply(case_tac x, simp) +apply(subgoal_tac "length (ci ly a b) mod 2 = 0") apply(auto) done @@ -4984,16 +2740,6 @@ lemma [simp]: "length mopup_b = 12" apply(simp add: mopup_b_def) done -(* -lemma [elim]: "\na < 4 * n; tshift (mop_bef n) q ! na = (a, b)\ \ - b \ q + (2 * n + 6)" -apply(induct n, simp, simp add: mop_bef.simps nth_append tshift_append shift_length) -apply(case_tac "na < 4*n", simp, simp) -apply(subgoal_tac "na = 4*n \ na = 1 + 4*n \ na = 2 + 4*n \ na = 3 + 4*n", - auto simp: shift_length) -apply(simp_all add: tshift.simps) -done -*) lemma mp_up_all_le: "list_all (\(acn, s). s \ q + (2 * n + 6)) [(R, Suc (Suc (2 * n + q))), (R, Suc (2 * n + q)), @@ -5028,18 +2774,23 @@ apply(auto simp: mopup.simps) done -lemma t_compiled_correct: - "\tp = tm_of ap; ly = layout_of ap; mop_ss = start_of ly (length ap)\ \ +lemma wf_tm_from_abacus: + "tp = tm_of ap \ 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 - - - - - - - +lemma wf_tm_from_recf: + assumes compile: "tp = tm_of_rec recf" + shows "tm_wf0 tp" +proof - + obtain a b c where "rec_ci recf = (a, b, c)" + by (metis prod_cases3) + thus "?thesis" + using compile + using wf_tm_from_abacus[of "tm_of (a [+] dummy_abc b)" "(a [+] dummy_abc b)" b] + by simp +qed + +end \ No newline at end of file diff -r e9ef4ada308b -r d8e6f0798e23 thys/UF.thy --- a/thys/UF.thy Thu Mar 14 20:43:43 2013 +0000 +++ b/thys/UF.thy Wed Mar 27 09:47:02 2013 +0000 @@ -158,60 +158,6 @@ it always returns meaningful results for primitive recursive functions. *} -function rec_exec :: "recf \ nat list \ nat" - where - "rec_exec z xs = 0" | - "rec_exec s xs = (Suc (xs ! 0))" | - "rec_exec (id m n) xs = (xs ! n)" | - "rec_exec (Cn n f gs) xs = - (let ys = (map (\ a. rec_exec a xs) gs) in - rec_exec f ys)" | - "rec_exec (Pr n f g) xs = - (if last xs = 0 then - rec_exec f (butlast xs) - else rec_exec g (butlast xs @ [last xs - 1] @ - [rec_exec (Pr n f g) (butlast xs @ [last xs - 1])]))" | - "rec_exec (Mn n f) xs = (LEAST x. rec_exec f (xs @ [x]) = 0)" -by pat_completeness auto -termination -proof - show "wf (measures [\ (r, xs). size r, (\ (r, xs). last xs)])" - by auto -next - fix n f gs xs x - assume "(x::recf) \ set gs" - thus "((x, xs), Cn n f gs, xs) \ - measures [\(r, xs). size r, \(r, xs). last xs]" - by(induct gs, auto) -next - fix n f gs xs x - assume "x = map (\a. rec_exec a xs) gs" - "\x. x \ set gs \ rec_exec_dom (x, xs)" - thus "((f, x), Cn n f gs, xs) \ - measures [\(r, xs). size r, \(r, xs). last xs]" - by(auto) -next - fix n f g xs - show "((f, butlast xs), Pr n f g, xs) \ - measures [\(r, xs). size r, \(r, xs). last xs]" - by auto -next - fix n f g xs - assume "last xs \ (0::nat)" thus - "((Pr n f g, butlast xs @ [last xs - 1]), Pr n f g, xs) - \ measures [\(r, xs). size r, \(r, xs). last xs]" - by auto -next - fix n f g xs - show "((g, butlast xs @ [last xs - 1] @ [rec_exec (Pr n f g) (butlast xs @ [last xs - 1])]), - Pr n f g, xs) \ measures [\(r, xs). size r, \(r, xs). last xs]" - by auto -next - fix n f xs x - show "((f, xs @ [x]), Mn n f, xs) \ - measures [\(r, xs). size r, \(r, xs). last xs]" - by auto -qed declare rec_exec.simps[simp del] constn.simps[simp del] @@ -280,7 +226,6 @@ else 1)" by(induct_tac y, auto simp: rec_exec.simps sg_lemma rec_conj_def mult_lemma) - text {* Correctness of @{text "rec_disj"}. *} @@ -2948,190 +2893,6 @@ The lemma relates the interpreter of primitive fucntions with the calculation relation of general recursive functions. *} -lemma prime_rel_exec_eq: "primerec r (length xs) - \ rec_calc_rel r xs rs = (rec_exec r xs = rs)" -proof(induct r xs arbitrary: rs rule: rec_exec.induct, simp_all) - fix xs rs - assume "primerec z (length (xs::nat list))" - hence "length xs = Suc 0" by(erule_tac prime_z_reverse, simp) - thus "rec_calc_rel z xs rs = (rec_exec z xs = rs)" - apply(case_tac xs, simp, auto) - apply(erule_tac calc_z_reverse, simp add: rec_exec.simps) - apply(simp add: rec_exec.simps, rule_tac calc_z) - done -next - fix xs rs - assume "primerec s (length (xs::nat list))" - hence "length xs = Suc 0" .. - thus "rec_calc_rel s xs rs = (rec_exec s xs = rs)" - by(case_tac xs, auto simp: rec_exec.simps intro: calc_s - elim: calc_s_reverse) -next - fix m n xs rs - assume "primerec (recf.id m n) (length (xs::nat list))" - thus - "rec_calc_rel (recf.id m n) xs rs = - (rec_exec (recf.id m n) xs = rs)" - apply(erule_tac prime_id_reverse) - apply(simp add: rec_exec.simps, auto) - apply(erule_tac calc_id_reverse, simp) - apply(rule_tac calc_id, auto) - done -next - fix n f gs xs rs - assume ind1: - "\x rs. \x \ set gs; primerec x (length xs)\ \ - rec_calc_rel x xs rs = (rec_exec x xs = rs)" - and ind2: - "\x rs. \x = map (\a. rec_exec a xs) gs; - primerec f (length gs)\ \ - rec_calc_rel f (map (\a. rec_exec a xs) gs) rs = - (rec_exec f (map (\a. rec_exec a xs) gs) = rs)" - and h: "primerec (Cn n f gs) (length xs)" - show "rec_calc_rel (Cn n f gs) xs rs = - (rec_exec (Cn n f gs) xs = rs)" - proof(auto simp: rec_exec.simps, erule_tac calc_cn_reverse, auto) - fix ys - assume g1:"\ka. rec_exec a xs) gs) rs = - (rec_exec f (map (\a. rec_exec a xs) gs) = rs)" - apply(rule_tac ind2, auto) - using h - apply(erule_tac prime_cn_reverse, simp) - done - moreover have "ys = (map (\a. rec_exec a xs) gs)" - proof(rule_tac nth_equalityI, auto simp: g2) - fix i - assume "i < length gs" thus "ys ! i = rec_exec (gs!i) xs" - using ind1[of "gs ! i" "ys ! i"] g1 h - apply(erule_tac prime_cn_reverse, simp) - done - qed - ultimately show "rec_exec f (map (\a. rec_exec a xs) gs) = rs" - using g3 - by(simp) - next - from h show - "rec_calc_rel (Cn n f gs) xs - (rec_exec f (map (\a. rec_exec a xs) gs))" - apply(rule_tac rs = "(map (\a. rec_exec a xs) gs)" in calc_cn, - auto) - apply(erule_tac [!] prime_cn_reverse, auto) - proof - - fix k - assume "k < length gs" "primerec f (length gs)" - "\iia. rec_exec a xs) gs) - (rec_exec f (map (\a. rec_exec a xs) gs))" - using ind2[of "(map (\a. rec_exec a xs) gs)" - "(rec_exec f (map (\a. rec_exec a xs) gs))"] - by simp - qed - qed -next - fix n f g xs rs - assume ind1: - "\rs. \last xs = 0; primerec f (length xs - Suc 0)\ - \ rec_calc_rel f (butlast xs) rs = - (rec_exec f (butlast xs) = rs)" - and ind2 : - "\rs. \0 < last xs; - primerec (Pr n f g) (Suc (length xs - Suc 0))\ \ - rec_calc_rel (Pr n f g) (butlast xs @ [last xs - Suc 0]) rs - = (rec_exec (Pr n f g) (butlast xs @ [last xs - Suc 0]) = rs)" - and ind3: - "\rs. \0 < last xs; primerec g (Suc (Suc (length xs - Suc 0)))\ - \ rec_calc_rel g (butlast xs @ - [last xs - Suc 0, rec_exec (Pr n f g) - (butlast xs @ [last xs - Suc 0])]) rs = - (rec_exec g (butlast xs @ [last xs - Suc 0, - rec_exec (Pr n f g) - (butlast xs @ [last xs - Suc 0])]) = rs)" - and h: "primerec (Pr n f g) (length (xs::nat list))" - show "rec_calc_rel (Pr n f g) xs rs = (rec_exec (Pr n f g) xs = rs)" - proof(auto) - assume "rec_calc_rel (Pr n f g) xs rs" - thus "rec_exec (Pr n f g) xs = rs" - proof(erule_tac calc_pr_reverse) - fix l - assume g: "xs = l @ [0]" - "rec_calc_rel f l rs" - "n = length l" - thus "rec_exec (Pr n f g) xs = rs" - using ind1[of rs] h - apply(simp add: rec_exec.simps, - erule_tac prime_pr_reverse, simp) - done - next - fix l y ry - assume d:"xs = l @ [Suc y]" - "rec_calc_rel (Pr (length l) f g) (l @ [y]) ry" - "n = length l" - "rec_calc_rel g (l @ [y, ry]) rs" - moreover hence "primerec g (Suc (Suc n))" using h - proof(erule_tac prime_pr_reverse) - assume "primerec g (Suc (Suc n))" "length xs = Suc n" - thus "?thesis" by simp - qed - ultimately show "rec_exec (Pr n f g) xs = rs" - apply(simp) - using ind3[of rs] - apply(simp add: rec_pr_Suc_simp_rewrite) - using ind2[of ry] h - apply(simp) - done - qed - next - show "rec_calc_rel (Pr n f g) xs (rec_exec (Pr n f g) xs)" - proof - - have "rec_calc_rel (Pr n f g) (butlast xs @ [last xs]) - (rec_exec (Pr n f g) (butlast xs @ [last xs]))" - using h - apply(erule_tac prime_pr_reverse, simp) - apply(case_tac "last xs", simp) - apply(rule_tac calc_pr_zero, simp) - using ind1[of "rec_exec (Pr n f g) (butlast xs @ [0])"] - apply(simp add: rec_exec.simps, simp, simp, simp) - 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] @@ -3141,11 +2902,6 @@ @{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 @@ -3283,12 +3039,6 @@ 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}, @@ -3303,38 +3053,61 @@ @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ done -lemma nonstop_eq[simp]: - "rec_calc_rel rec_nonstop [m, r, t] rs = - (rec_exec rec_nonstop [m, r, t] = rs)" -apply(rule prime_rel_exec_eq, auto) -done - -lemma halt_lemma': - "rec_calc_rel rec_halt [m, r] t = - (rec_calc_rel rec_nonstop [m, r, t] 0 \ - (\ t'< t. - (\ y. rec_calc_rel rec_nonstop [m, r, t'] y \ - y \ 0)))" -apply(auto simp: rec_halt_def) -apply(erule calc_mn_reverse, simp) -apply(erule_tac calc_mn_reverse) -apply(erule_tac x = t' in allE, simp) -apply(rule_tac calc_mn, simp_all) -done +lemma primerec_terminate: + "\primerec f x; length xs = x\ \ terminate f xs" +proof(induct arbitrary: xs rule: primerec.induct) + fix xs + assume "length (xs::nat list) = Suc 0" thus "terminate z xs" + by(case_tac xs, auto intro: termi_z) +next + fix xs + assume "length (xs::nat list) = Suc 0" thus "terminate s xs" + by(case_tac xs, auto intro: termi_s) +next + fix n m xs + assume "n < m" "length (xs::nat list) = m" thus "terminate (id m n) xs" + by(erule_tac termi_id, simp) +next + fix f k gs m n xs + assume ind: "\i (\x. length x = m \ terminate (gs ! i) x)" + and ind2: "\ xs. length xs = k \ terminate f xs" + and h: "primerec f k" "length gs = k" "m = n" "length (xs::nat list) = m" + have "terminate f (map (\g. rec_exec g xs) gs)" + using ind2[of "(map (\g. rec_exec g xs) gs)"] h + by simp + moreover have "\g\set gs. terminate g xs" + using ind h + by(auto simp: set_conv_nth) + ultimately show "terminate (Cn n f gs) xs" + using h + by(rule_tac termi_cn, auto) +next + fix f n g m xs + assume ind1: "\xs. length xs = n \ terminate f xs" + and ind2: "\xs. length xs = Suc (Suc n) \ terminate g xs" + and h: "primerec f n" " primerec g (Suc (Suc n))" " m = Suc n" "length (xs::nat list) = m" + have "\y - (\ t'< t. (\ y. rec_exec rec_nonstop [m, r, t'] = y - \ y \ 0)))" -using halt_lemma'[of m r t] -by simp - + text {*F: universal machine*} text {* @@ -3370,11 +3143,6 @@ 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 {* @@ -3393,77 +3161,42 @@ done lemma [simp]: - "\ys \ []; k < length ys\ \ + "\ys \ []; + k < length ys\ \ (get_fstn_args (length ys) (length ys) ! k) = id (length ys) (k)" by(erule_tac get_fstn_args_nth) -lemma calc_rel_get_pren: - "\ys \ []; k < length ys\ \ - rec_calc_rel (get_fstn_args (length ys) (length ys) ! k) ys - (ys ! k)" -apply(simp) -apply(rule_tac calc_id, auto) +lemma terminate_halt_lemma: + "\rec_exec rec_nonstop ([m, r] @ [t]) = 0; + \i \ terminate rec_halt [m, r]" +apply(simp add: rec_halt_def) +apply(rule_tac termi_mn, auto) +apply(rule_tac [!] primerec_terminate, auto) done -lemma [elim]: - "\xs \ []; k < Suc (length xs)\ \ - rec_calc_rel (get_fstn_args (Suc (length xs)) - (Suc (length xs)) ! k) (m # xs) ((m # xs) ! k)" -using calc_rel_get_pren[of "m#xs" k] -apply(simp) -done text {* The correctness of @{text "rec_F"}, halt case. *} -lemma F_lemma: - "rec_calc_rel rec_halt [m, r] t \ - rec_calc_rel rec_F [m, r] (valu (rght (conf m r t)))" + +lemma F_lemma: "rec_exec rec_halt [m, r] = t \ rec_exec rec_F [m, r] = (valu (rght (conf m r t)))" +by(simp add: rec_F_def rec_exec.simps value_lemma right_lemma conf_lemma halt_lemma) + +lemma terminate_F_lemma: "terminate rec_halt [m, r] \ terminate rec_F [m, r]" apply(simp add: rec_F_def) -apply(rule_tac rs = "[rght (conf m r t)]" in calc_cn, - auto simp: value_lemma) -apply(rule_tac rs = "[conf m r t]" in calc_cn, - auto simp: right_lemma) -apply(rule_tac rs = "[m, r, t]" in calc_cn, auto) -apply(subgoal_tac " k = 0 \ k = Suc 0 \ k = Suc (Suc 0)", - auto simp:nth_append) -apply(rule_tac [1-2] calc_id, simp_all add: conf_lemma) +apply(rule_tac termi_cn, auto) +apply(rule_tac primerec_terminate, auto) +apply(rule_tac termi_cn, auto) +apply(rule_tac primerec_terminate, auto) +apply(rule_tac termi_cn, auto) +apply(rule_tac primerec_terminate, auto) +apply(rule_tac [!] termi_id, auto) done - text {* The correctness of @{text "rec_F"}, nonhalt case. *} -lemma F_lemma2: - "\ t. \ rec_calc_rel rec_halt [m, r] t \ - \ rs. \ rec_calc_rel rec_F [m, r] rs" -apply(auto simp: rec_F_def) -apply(erule_tac calc_cn_reverse, simp (no_asm_use))+ -proof - - fix rs rsa rsb rsc - assume h: - "\t. \ rec_calc_rel rec_halt [m, r] t" - "length rsa = Suc 0" - "rec_calc_rel rec_valu rsa rs" - "length rsb = Suc 0" - "rec_calc_rel rec_right rsb (rsa ! 0)" - "length rsc = (Suc (Suc (Suc 0)))" - "rec_calc_rel rec_conf rsc (rsb ! 0)" - and g: "\ksteps0 (Suc 0, Bk\l, ) tp stp = (0, Bk\m, Oc\rs@Bk\n); + tm_wf (tp,0); 0 \ terminate rec_halt [code tp, (bl2wc ())]" +apply(frule_tac halt_least_step, auto) +thm terminate_halt_lemma +apply(rule_tac t = stpa in terminate_halt_lemma) +apply(simp_all add: nonstop_lemma, auto) +done + +lemma terminate_F: + "\steps0 (Suc 0, Bk\l, ) tp stp = (0, Bk\m, Oc\rs@Bk\n); + tm_wf (tp,0); 0 \ terminate rec_F [code tp, (bl2wc ())]" +apply(drule_tac terminate_halt, simp_all) +apply(erule_tac terminate_F_lemma) +done + lemma F_correct: "\steps0 (Suc 0, Bk\l, ) tp stp = (0, Bk\m, Oc\rs@Bk\n); tm_wf (tp,0); 0 - \ rec_calc_rel rec_F [code tp, (bl2wc ())] (rs - Suc 0)" + \ rec_exec rec_F [code tp, (bl2wc ())] = (rs - Suc 0)" apply(frule_tac halt_least_step, auto) apply(frule_tac nonstop_t_eq, auto simp: nonstop_lemma) using rec_t_eq_steps[of tp l lm stp] @@ -4880,24 +4629,20 @@ using halt_state_keep[of "code tp" lm stpa stp] by(simp) moreover have g2: - "rec_calc_rel rec_halt [code tp, (bl2wc ())] stpa" + "rec_exec rec_halt [code tp, (bl2wc ())] = stpa" using h - apply(simp add: halt_lemma nonstop_lemma, auto) - done + by(auto simp: rec_exec.simps rec_halt_def nonstop_lemma intro!: Least_equality) show - "rec_calc_rel rec_F [code tp, (bl2wc ())] (rs - Suc 0)" + "rec_exec rec_F [code tp, (bl2wc ())] = (rs - Suc 0)" proof - have - "rec_calc_rel rec_F [code tp, (bl2wc ())] - (valu (rght (conf (code tp) (bl2wc ()) stpa)))" - apply(rule F_lemma) using g2 h by auto - moreover have "valu (rght (conf (code tp) (bl2wc ()) stpa)) = rs - Suc 0" using g1 apply(simp add: valu.simps trpl_code.simps bl2wc.simps bl2nat_append lg_power) done - ultimately show "?thesis" by simp + thus "?thesis" + by(simp add: rec_exec.simps F_lemma g2) qed qed diff -r e9ef4ada308b -r d8e6f0798e23 thys/UTM.thy --- a/thys/UTM.thy Thu Mar 14 20:43:43 2013 +0000 +++ b/thys/UTM.thy Wed Mar 27 09:47:02 2013 +0000 @@ -1214,22 +1214,11 @@ apply(simp add: t_twice_def mopup.simps t_twice_compile_def) done -lemma [intro]: "rec_calc_rel (recf.id (Suc 0) 0) [rs] rs" - apply(rule_tac calc_id, simp_all) - done - -lemma [intro]: "rec_calc_rel (constn 2) [rs] 2" -using prime_rel_exec_eq[of "constn 2" "[rs]" 2] -apply(subgoal_tac "primerec (constn 2) 1", auto) -done - -lemma [intro]: "rec_calc_rel rec_mult [rs, 2] (2 * rs)" -using prime_rel_exec_eq[of "rec_mult" "[rs, 2]" "2*rs"] -apply(subgoal_tac "primerec rec_mult (Suc (Suc 0))", auto) -done - declare start_of.simps[simp del] +lemma twice_lemma: "rec_exec rec_twice [rs] = 2*rs" +by(auto simp: rec_twice_def rec_exec.simps) + lemma t_twice_correct: "\stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (tm_of abc_twice @ shift (mopup (Suc 0)) ((length (tm_of abc_twice) div 2))) stp = @@ -1237,28 +1226,23 @@ proof(case_tac "rec_ci rec_twice") fix a b c assume h: "rec_ci rec_twice = (a, b, c)" - have "\stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\(n)) (tm_of abc_twice @ shift (mopup 1) - (length (tm_of abc_twice) div 2)) stp = (0, Bk\(m) @ Bk # Bk # ires, Oc\(Suc (2*rs)) @ Bk\(l))" - proof(rule_tac recursive_compile_to_tm_correct) + have "\stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\(n)) (tm_of abc_twice @ shift (mopup (length [rs])) + (length (tm_of abc_twice) div 2)) stp = (0, Bk\(m) @ Bk # Bk # ires, Oc\(Suc (rec_exec rec_twice [rs])) @ Bk\(l))" + thm recursive_compile_to_tm_correct1 + proof(rule_tac recursive_compile_to_tm_correct1) show "rec_ci rec_twice = (a, b, c)" by (simp add: h) next - show "rec_calc_rel rec_twice [rs] (2 * rs)" - apply(simp add: rec_twice_def) - apply(rule_tac rs = "[rs, 2]" in calc_cn, simp_all) - apply(rule_tac allI, case_tac k, auto) - done + show "terminate rec_twice [rs]" + apply(rule_tac primerec_terminate, auto) + apply(simp add: rec_twice_def, auto simp: constn.simps numeral_2_eq_2) + by(auto) next - show "length [rs] = 1" by simp - next - show "layout_of (a [+] dummy_abc 1) = layout_of (a [+] dummy_abc 1)" by simp - next - show "tm_of abc_twice = tm_of (a [+] dummy_abc 1)" + show "tm_of abc_twice = tm_of (a [+] dummy_abc (length [rs]))" using h - apply(simp add: abc_twice_def) - done + by(simp add: abc_twice_def) qed thus "?thesis" - apply(simp add: tape_of_nl_abv tape_of_nat_list.simps tape_of_nat_abv) + apply(simp add: tape_of_nl_abv tape_of_nat_list.simps tape_of_nat_abv rec_exec.simps twice_lemma) done qed @@ -1386,8 +1370,7 @@ lemma [simp]: " tm_wf (t_twice_compile, 0)" apply(simp only: t_twice_compile_def) -apply(rule_tac t_compiled_correct) -apply(simp_all add: abc_twice_def) +apply(rule_tac wf_tm_from_abacus, simp) done lemma t_twice_change_term_state: @@ -1503,7 +1486,7 @@ apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI, rule_tac x = rn in exI) apply(simp add: t_wcode_main_def) - apply(simp add: replicate_Suc[THEN sym] exp_add[THEN sym] del: replicate_Suc) + apply(simp add: replicate_Suc[THEN sym] replicate_add [THEN sym] del: replicate_Suc) done from this obtain stpb lnb rnb where stp2: "steps0 (13, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (Suc rs)) @ Bk\(rna)) t_wcode_main stpb = @@ -2018,15 +2001,13 @@ apply(simp add: t_fourtimes_len_def t_fourtimes_def mopup.simps t_fourtimes_compile_def) done -lemma [intro]: "rec_calc_rel (constn 4) [rs] 4" -using prime_rel_exec_eq[of "constn 4" "[rs]" 4] -apply(subgoal_tac "primerec (constn 4) 1", auto) -done - -lemma [intro]: "rec_calc_rel rec_mult [rs, 4] (4 * rs)" -using prime_rel_exec_eq[of "rec_mult" "[rs, 4]" "4*rs"] -apply(subgoal_tac "primerec rec_mult 2", auto simp: numeral_2_eq_2) -done +lemma [intro]: "primerec rec_fourtimes (Suc 0)" +apply(auto simp: rec_fourtimes_def numeral_4_eq_4 constn.simps) +by auto + +lemma fourtimes_lemma: "rec_exec rec_fourtimes [rs] = 4 * rs" +by(simp add: rec_exec.simps rec_fourtimes_def) + lemma t_fourtimes_correct: "\stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) @@ -2035,35 +2016,28 @@ proof(case_tac "rec_ci rec_fourtimes") fix a b c assume h: "rec_ci rec_fourtimes = (a, b, c)" - have "\stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\(n)) (tm_of abc_fourtimes @ shift (mopup 1) - (length (tm_of abc_fourtimes) div 2)) stp = (0, Bk\(m) @ Bk # Bk # ires, Oc\(Suc (4*rs)) @ Bk\(l))" - proof(rule_tac recursive_compile_to_tm_correct) + have "\stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\(n)) (tm_of abc_fourtimes @ shift (mopup (length [rs])) + (length (tm_of abc_fourtimes) div 2)) stp = (0, Bk\(m) @ Bk # Bk # ires, Oc\(Suc (rec_exec rec_fourtimes [rs])) @ Bk\(l))" + thm recursive_compile_to_tm_correct1 + proof(rule_tac recursive_compile_to_tm_correct1) show "rec_ci rec_fourtimes = (a, b, c)" by (simp add: h) next - show "rec_calc_rel rec_fourtimes [rs] (4 * rs)" - apply(simp add: rec_fourtimes_def) - apply(rule_tac rs = "[rs, 4]" in calc_cn, simp_all) - apply(rule_tac allI, case_tac k, auto simp: mult_lemma) - done + show "terminate rec_fourtimes [rs]" + apply(rule_tac primerec_terminate) + by auto next - show "length [rs] = 1" by simp - next - show "layout_of (a [+] dummy_abc 1) = layout_of (a [+] dummy_abc 1)" by simp - next - show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc 1)" + show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc (length [rs]))" using h - apply(simp add: abc_fourtimes_def) - done + by(simp add: abc_fourtimes_def) qed thus "?thesis" - apply(simp add: tape_of_nl_abv tape_of_nat_list.simps tape_of_nat_abv) + apply(simp add: tape_of_nl_abv tape_of_nat_list.simps tape_of_nat_abv fourtimes_lemma) done qed lemma wf_fourtimes[intro]: "tm_wf (t_fourtimes_compile, 0)" apply(simp only: t_fourtimes_compile_def) -apply(rule_tac t_compiled_correct) -apply(simp_all add: abc_twice_def) +apply(rule_tac wf_tm_from_abacus, simp) done lemma t_fourtimes_change_term_state: @@ -2175,7 +2149,7 @@ = (L, Suc 0)" apply(subgoal_tac "14 = Suc 13") apply(simp only: fetch.simps add_Suc nth_of.simps t_wcode_main_def) -apply(simp add:length_append length_shift Parity.two_times_even_div_two even_twice_len t_fourtimes_len_def) +apply(simp add:length_append length_shift Parity.two_times_even_div_two even_twice_len t_fourtimes_len_def nth_append) by arith lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) Bk @@ -2228,7 +2202,7 @@ apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI, rule_tac x = rn in exI, simp) - apply(simp add: replicate_Suc[THEN sym] exp_add[THEN sym] del: replicate_Suc) + apply(simp add: replicate_Suc[THEN sym] replicate_add[THEN sym] del: replicate_Suc) done from this obtain stpb lnb rnb where stp2: "steps0 (t_twice_len + 14, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (rs + 1)) @ Bk\(rna)) @@ -2254,8 +2228,6 @@ done qed -(**********************************************************) - fun wcode_on_left_moving_3_B :: "bin_inv_t" where "wcode_on_left_moving_3_B ires rs (l, r) = @@ -4742,7 +4714,7 @@ apply(rule_tac x = "Suc m" in exI, simp only: exp_ind) apply(simp only: exp_ind, simp) apply(subgoal_tac "m = length la + nata") -apply(rule_tac x = "m - nata" in exI, simp add: exp_add) +apply(rule_tac x = "m - nata" in exI, simp add: replicate_add) apply(drule_tac length_equal, simp) apply(simp only: exp_ind[THEN sym] replicate_Suc[THEN sym] replicate_add[THEN sym]) apply(rule_tac x = "m + Suc (Suc n)" in exI, simp) @@ -4757,40 +4729,36 @@ proof - obtain ap arity fp where a: "rec_ci rec_F = (ap, arity, fp)" by (metis prod_cases3) - moreover have b: "rec_calc_rel rec_F [code tp, (bl2wc ())] (rs - Suc 0)" + moreover have b: "rec_exec rec_F [code tp, (bl2wc ())] = (rs - Suc 0)" using assms apply(rule_tac F_correct, simp_all) done have "\ stp m l. steps0 (Suc 0, Bk # Bk # [], <[code tp, bl2wc ()]> @ Bk\i) - (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp - = (0, Bk\m @ Bk # Bk # [], Oc\Suc (rs - 1) @ Bk\l)" - proof(rule_tac recursive_compile_to_tm_correct) + (F_tprog @ shift (mopup (length [code tp, bl2wc ()])) (length F_tprog div 2)) stp + = (0, Bk\m @ Bk # Bk # [], Oc\Suc (rec_exec rec_F [code tp, (bl2wc ())]) @ Bk\l)" + proof(rule_tac recursive_compile_to_tm_correct1) show "rec_ci rec_F = (ap, arity, fp)" using a by simp next - show "rec_calc_rel rec_F [code tp, bl2wc ()] (rs - 1)" - using b by simp - next - show "length [code tp, bl2wc ()] = 2" by simp + show "terminate rec_F [code tp, bl2wc ()]" + using assms + by(rule_tac terminate_F, simp_all) next - show "layout_of (ap [+] dummy_abc 2) = layout_of (ap [+] dummy_abc 2)" - by simp - next - show "F_tprog = tm_of (ap [+] dummy_abc 2)" + show "F_tprog = tm_of (ap [+] dummy_abc (length [code tp, bl2wc ()]))" using a apply(simp add: F_tprog_def F_aprog_def numeral_2_eq_2) done qed then obtain stp m l where "steps0 (Suc 0, Bk # Bk # [], <[code tp, bl2wc ()]> @ Bk\i) - (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp - = (0, Bk\m @ Bk # Bk # [], Oc\Suc (rs - 1) @ Bk\l)" by blast + (F_tprog @ shift (mopup (length [code tp, (bl2wc ())])) (length F_tprog div 2)) stp + = (0, Bk\m @ Bk # Bk # [], Oc\Suc (rec_exec rec_F [code tp, (bl2wc ())]) @ Bk\l)" by blast hence "\ m. steps0 (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\i) (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp = (0, Bk\m, Oc\Suc (rs - 1) @ Bk\l)" proof - assume g: "steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]> @ Bk \ i) - (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp = - (0, Bk \ m @ [Bk, Bk], Oc \ Suc (rs - 1) @ Bk \ l)" + (F_tprog @ shift (mopup (length [code tp, bl2wc ()])) (length F_tprog div 2)) stp = + (0, Bk \ m @ [Bk, Bk], Oc \ Suc ((rec_exec rec_F [code tp, bl2wc ()])) @ Bk \ l)" moreover have "tinres [Bk, Bk] [Bk]" apply(auto simp: tinres_def) done @@ -4800,7 +4768,8 @@ (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp", auto) done ultimately show "?thesis" - apply(drule_tac tinres_steps1, auto) + using b + apply(drule_tac la = "Bk\m @ [Bk, Bk]" in tinres_steps1, auto simp: numeral_2_eq_2) done qed thus "?thesis" @@ -4819,7 +4788,7 @@ lemma [intro]: "tm_wf (t_utm, 0)" apply(simp only: t_utm_def F_tprog_def) -apply(rule_tac t_compiled_correct, auto) +apply(rule_tac wf_tm_from_abacus, auto) done lemma UTM_halt_lemma_pre: @@ -4993,141 +4962,162 @@ lemma nonstop_true: "\tm_wf (tp, 0); \ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))\ - \ \y. rec_calc_rel rec_nonstop - ([code tp, bl2wc (), y]) (Suc 0)" + \ \y. rec_exec rec_nonstop ([code tp, bl2wc (), y]) = (Suc 0)" apply(rule_tac allI, erule_tac x = y in allE) apply(case_tac "steps0 (Suc 0, Bk\(l), ) tp y", simp) apply(rule_tac nonstop_t_uhalt_eq, simp_all) done -declare ci_cn_para_eq[simp] +lemma cn_arity: "rec_ci (Cn n f gs) = (a, b, c) \ b = n" +by(case_tac "rec_ci f", simp add: rec_ci.simps) + +lemma mn_arity: "rec_ci (Mn n f) = (a, b, c) \ b = n" +by(case_tac "rec_ci f", simp add: rec_ci.simps) lemma F_aprog_uhalt: - "\tm_wf (tp,0); - \ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp)); - rec_ci rec_F = (F_ap, rs_pos, a_md)\ - \ \ stp. case abc_steps_l (0, [code tp, bl2wc ()] @ 0\(a_md - rs_pos ) - @ suflm) (F_ap) stp of (ss, e) \ ss < length (F_ap)" -apply(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf - ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])])") -apply(simp only: rec_F_def, rule_tac i = 0 and ga = a and gb = b and - gc = c in cn_gi_uhalt, simp, simp, simp, simp, simp, simp, simp) -apply(simp add: ci_cn_para_eq) -apply(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_conf - ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt]))") -apply(rule_tac rf = "(Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf - ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])])" - and n = "Suc (Suc 0)" and f = rec_right and - gs = "[Cn (Suc (Suc 0)) rec_conf - ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])]" - and i = 0 and ga = aa and gb = ba and gc = ca in - cn_gi_uhalt) -apply(simp, simp, simp, simp, simp, simp, simp, - simp add: ci_cn_para_eq) -apply(case_tac "rec_ci rec_halt") -apply(rule_tac rf = "(Cn (Suc (Suc 0)) rec_conf - ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt]))" - and n = "Suc (Suc 0)" and f = "rec_conf" and - gs = "([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])" and - i = "Suc (Suc 0)" and gi = "rec_halt" and ga = ab and gb = bb and - gc = cb in cn_gi_uhalt) -apply(simp, simp, simp, simp, simp add: nth_append, simp, - simp add: nth_append, simp add: rec_halt_def) -apply(simp only: rec_halt_def) -apply(case_tac [!] "rec_ci ((rec_nonstop))") -apply(rule_tac allI, rule_tac impI, simp) -apply(case_tac j, simp) -apply(rule_tac x = "code tp" in exI, rule_tac calc_id, simp, simp, simp, simp) -apply(rule_tac x = "bl2wc ()" in exI, rule_tac calc_id, simp, simp, simp) -apply(rule_tac rf = "Mn (Suc (Suc 0)) (rec_nonstop)" - and f = "(rec_nonstop)" and n = "Suc (Suc 0)" - and aprog' = ac and rs_pos' = bc and a_md' = cc in Mn_unhalt) -apply(simp, simp add: rec_halt_def , simp, simp) -apply(drule_tac nonstop_true, simp_all) -apply(rule_tac allI) -apply(erule_tac x = y in allE)+ -apply(simp) -done + assumes wf_tm: "tm_wf (tp,0)" + and unhalt: "\ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))" + and compile: "rec_ci rec_F = (F_ap, rs_pos, a_md)" + shows "{\ nl. nl = [code tp, bl2wc ()] @ 0\(a_md - rs_pos ) @ suflm} (F_ap) \" + using compile +proof(simp only: rec_F_def) + assume h: "rec_ci (Cn (Suc (Suc 0)) rec_valu [Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) + rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt]]]) = + (F_ap, rs_pos, a_md)" + moreover hence "rs_pos = Suc (Suc 0)" + using cn_arity + by simp + moreover obtain ap1 ar1 ft1 where a: "rec_ci + (Cn (Suc (Suc 0)) rec_right + [Cn (Suc (Suc 0)) rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt]]) = (ap1, ar1, ft1)" + by(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) + rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt]])", auto) + moreover hence b: "ar1 = Suc (Suc 0)" + using cn_arity by simp + ultimately show "?thesis" + proof(rule_tac i = 0 in cn_unhalt_case, auto) + fix anything + obtain ap2 ar2 ft2 where c: + "rec_ci (Cn (Suc (Suc 0)) rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt]) + = (ap2, ar2, ft2)" + by(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_conf + [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt])", auto) + moreover hence d:"ar2 = Suc (Suc 0)" + using cn_arity by simp + ultimately have "{\nl. nl = [code tp, bl2wc ()] @ 0 \ (ft1 - Suc (Suc 0)) @ anything} ap1 \" + using a b c d + proof(rule_tac i = 0 in cn_unhalt_case, auto) + fix anything + obtain ap3 ar3 ft3 where e: "rec_ci rec_halt = (ap3, ar3, ft3)" + by(case_tac "rec_ci rec_halt", auto) + hence f: "ar3 = Suc (Suc 0)" + using mn_arity + by(simp add: rec_halt_def) + have "{\nl. nl = [code tp, bl2wc ()] @ 0 \ (ft2 - Suc (Suc 0)) @ anything} ap2 \" + using c d e f + proof(rule_tac i = 2 in cn_unhalt_case, auto simp: rec_halt_def) + fix anything + have "{\nl. nl = [code tp, bl2wc ()] @ 0 \ (ft3 - Suc (Suc 0)) @ anything} ap3 \" + using e f + proof(rule_tac mn_unhalt_case, auto simp: rec_halt_def) + fix i + show "terminate rec_nonstop [code tp, bl2wc (), i]" + by(rule_tac primerec_terminate, auto) + next + fix i + show "0 < rec_exec rec_nonstop [code tp, bl2wc (), i]" + using assms + by(drule_tac nonstop_true, auto) + qed + thus "{\nl. nl = code tp # bl2wc () # 0 \ (ft3 - Suc (Suc 0)) @ anything} ap3 \" by simp + next + fix apj arj ftj j anything + assume "j<2" "rec_ci ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j) = (apj, arj, ftj)" + hence "{\nl. nl = [code tp, bl2wc ()] @ 0 \ (ftj - arj) @ anything} apj + {\nl. nl = [code tp, bl2wc ()] @ + rec_exec ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j) [code tp, bl2wc ()] # + 0 \ (ftj - Suc arj) @ anything}" + apply(rule_tac recursive_compile_correct) + apply(case_tac j, auto) + apply(rule_tac [!] primerec_terminate) + by(auto) + thus "{\nl. nl = code tp # bl2wc () # 0 \ (ftj - arj) @ anything} apj + {\nl. nl = code tp # bl2wc () # rec_exec ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) + (Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j) [code tp, bl2wc ()] # 0 \ (ftj - Suc arj) @ anything}" + by simp + next + fix j + assume "(j::nat) < 2" + thus "terminate ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j) + [code tp, bl2wc ()]" + by(case_tac j, auto intro!: primerec_terminate) + qed + thus "{\nl. nl = code tp # bl2wc () # 0 \ (ft2 - Suc (Suc 0)) @ anything} ap2 \" + by simp + qed + thus "{\nl. nl = code tp # bl2wc () # 0 \ (ft1 - Suc (Suc 0)) @ anything} ap1 \" by simp + qed +qed lemma uabc_uhalt': "\tm_wf (tp, 0); \ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp)); rec_ci rec_F = (ap, pos, md)\ - \ \ stp. case abc_steps_l (0, [code tp, bl2wc ()]) ap stp of (ss, e) - \ ss < length ap" + \ {\ nl. nl = [code tp, bl2wc ()]} ap \" proof(frule_tac F_ap = ap and rs_pos = pos and a_md = md - and suflm = "[]" in F_aprog_uhalt, auto) - fix stp a b + and suflm = "[]" in F_aprog_uhalt, auto simp: abc_Hoare_unhalt_def, + case_tac "abc_steps_l (0, [code tp, bl2wc ()]) ap n", simp) + fix n a b assume h: - "\stp. case abc_steps_l (0, code tp # bl2wc () # 0\(md - pos)) ap stp of - (ss, e) \ ss < length ap" - "abc_steps_l (0, [code tp, bl2wc ()]) ap stp = (a, b)" + "\n. abc_notfinal (abc_steps_l (0, code tp # bl2wc () # 0 \ (md - pos)) ap n) ap" + "abc_steps_l (0, [code tp, bl2wc ()]) ap n = (a, b)" "tm_wf (tp, 0)" "rec_ci rec_F = (ap, pos, md)" - moreover have "ap \ []" - using h apply(rule_tac rec_ci_not_null, simp) - done + moreover have a: "ap \ []" + using h rec_ci_not_null[of "rec_F" pos md] by auto ultimately show "a < length ap" - proof(erule_tac x = stp in allE, - case_tac "abc_steps_l (0, code tp # bl2wc () # 0\(md - pos)) ap stp", simp) - fix aa ba - assume g: "aa < length ap" - "abc_steps_l (0, code tp # bl2wc () # 0\(md - pos)) ap stp = (aa, ba)" - "ap \ []" + proof(erule_tac x = n in allE) + assume g: "abc_notfinal (abc_steps_l (0, code tp # bl2wc () # 0 \ (md - pos)) ap n) ap" + obtain ss nl where b : "abc_steps_l (0, code tp # bl2wc () # 0 \ (md - pos)) ap n = (ss, nl)" + by (metis prod.exhaust) + then have c: "ss < length ap" + using g by simp thus "?thesis" + using a b c using abc_list_crsp_steps[of "[code tp, bl2wc ()]" - "md - pos" ap stp aa ba] h - apply(simp) - done + "md - pos" ap n ss nl] h + by(simp) qed qed lemma uabc_uhalt: "\tm_wf (tp, 0); \ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))\ - \ \ stp. case abc_steps_l (0, [code tp, bl2wc ()]) F_aprog - stp of (ss, e) \ ss < length F_aprog" + \ {\ nl. nl = [code tp, bl2wc ()]} F_aprog \ " apply(case_tac "rec_ci rec_F", simp add: F_aprog_def) apply(drule_tac ap = a and pos = b and md = c in uabc_uhalt', simp_all) -proof - - fix a b c - assume - "\stp. case abc_steps_l (0, [code tp, bl2wc ()]) a stp of (ss, e) - \ ss < length a" - "rec_ci rec_F = (a, b, c)" - thus - "\stp. case abc_steps_l (0, [code tp, bl2wc ()]) - (a [+] dummy_abc (Suc (Suc 0))) stp of (ss, e) \ - ss < Suc (Suc (Suc (length a)))" - using abc_append_uhalt1[of a "[code tp, bl2wc ()]" - "a [+] dummy_abc (Suc (Suc 0))" "[]" "dummy_abc (Suc (Suc 0))"] - apply(simp) - done -qed +apply(rule_tac abc_Hoare_plus_unhalt1, simp) +done lemma tutm_uhalt': assumes tm_wf: "tm_wf (tp,0)" and unhalt: "\ stp. (\ TSTD (steps0 (1, Bk\(l), ) tp stp))" shows "\ stp. \ is_final (steps0 (1, [Bk, Bk], <[code tp, bl2wc ()]>) t_utm stp)" unfolding t_utm_def -proof(rule_tac compile_correct_unhalt) - show "layout_of F_aprog = layout_of F_aprog" by simp -next +proof(rule_tac compile_correct_unhalt, auto) show "F_tprog = tm_of F_aprog" by(simp add: F_tprog_def) next - show "crsp (layout_of F_aprog) (0, [code tp, bl2wc ()]) (1, [Bk, Bk], <[code tp, bl2wc ()]>) []" + show "crsp (layout_of F_aprog) (0, [code tp, bl2wc ()]) (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) []" by(auto simp: crsp.simps start_of.simps) next - show "length F_tprog div 2 = length F_tprog div 2" by simp -next - show "\stp. case abc_steps_l (0, [code tp, bl2wc ()]) F_aprog stp of (as, am) \ as < length F_aprog" + fix stp a b + show "abc_steps_l (0, [code tp, bl2wc ()]) F_aprog stp = (a, b) \ a < length F_aprog" using assms - apply(erule_tac uabc_uhalt, simp) - done + apply(drule_tac uabc_uhalt, auto simp: abc_Hoare_unhalt_def) + by(erule_tac x = stp in allE, erule_tac x = stp in allE, simp) qed - lemma tinres_commute: "tinres r r' \ tinres r' r" apply(auto simp: tinres_def) done @@ -5166,7 +5156,7 @@ apply(case_tac "m > Suc (Suc 0)") apply(rule_tac x = "m - Suc (Suc 0)" in exI) apply(case_tac m, simp_all add: , case_tac nat, simp_all add: replicate_Suc) -apply(rule_tac x = "2 - m" in exI, simp add: exp_add[THEN sym]) +apply(rule_tac x = "2 - m" in exI, simp add: replicate_add[THEN sym]) apply(simp only: numeral_2_eq_2, simp add: replicate_Suc) done @@ -5330,4 +5320,5 @@ using assms(2-3) apply(simp add: tape_of_nat_abv) done + end \ No newline at end of file