diff --git a/thys/Dependent_SIFUM_Refinement/Examples/Eg1Eg2.thy b/thys/Dependent_SIFUM_Refinement/Examples/Eg1Eg2.thy --- a/thys/Dependent_SIFUM_Refinement/Examples/Eg1Eg2.thy +++ b/thys/Dependent_SIFUM_Refinement/Examples/Eg1Eg2.thy @@ -1,197 +1,197 @@ theory Eg1Eg2 imports Eg1 Eg2 "../CompositionalRefinement" begin locale sifum_refinement_eg1_eg2 = A: sifum_example + C: sifum_example2 primrec Eg2_var\<^sub>C_of_Eg1 :: "Eg1.var \ Eg2.var\<^sub>C" where "Eg2_var\<^sub>C_of_Eg1 control_var = control_var\<^sub>C" | "Eg2_var\<^sub>C_of_Eg1 buffer = buffer\<^sub>C" | "Eg2_var\<^sub>C_of_Eg1 high_var = high_var\<^sub>C" | "Eg2_var\<^sub>C_of_Eg1 low_var = low_var\<^sub>C" | "Eg2_var\<^sub>C_of_Eg1 temp = temp\<^sub>C" sublocale sifum_refinement_eg1_eg2 \ sifum_refinement dma dma\<^sub>C \_vars \_vars\<^sub>C \ \\<^sub>C A.eval\<^sub>w C.eval\<^sub>w undefined Eg2_var\<^sub>C_of_Eg1 - supply image_cong_simp [cong del] INF_cong_simp [cong] SUP_cong_simp [cong] subst_all [simp del] subst_all' [simp del] + supply image_cong_simp [cong del] INF_cong_simp [cong] SUP_cong_simp [cong] subst_all [simp del] apply(unfold_locales) apply(erule C.eval_det, simp) apply(rule C.Var_finite) apply(simp add:\_vars\<^sub>C_def split:if_splits) apply(simp add:\_vars\<^sub>C_def dma\<^sub>C_def) apply(simp add:\_vars\<^sub>C_def dma\<^sub>C_def) apply(rule inj_onI, simp) apply(case_tac x) apply(case_tac y, simp+) apply(case_tac y, simp+) apply(case_tac y, simp+) apply(case_tac y, simp+) apply(case_tac y, simp+) apply(case_tac x\<^sub>A) apply(clarsimp simp:dma\<^sub>C_def dma_def dma_control_var_def dma_control_var\<^sub>C_def)+ apply(case_tac x\<^sub>A) apply(clarsimp simp:\_vars\<^sub>C_def \_vars_def)+ apply(rule set_eqI, clarsimp) apply(case_tac x, clarsimp) apply(rule_tac x=control_var in rev_image_eqI, clarsimp+) apply(case_tac xa, clarsimp+) apply(case_tac xb, clarsimp+) apply(case_tac xa, clarsimp+) apply(case_tac xb, clarsimp+) apply(case_tac xa, clarsimp+) apply(case_tac xb, clarsimp+) apply(case_tac xa, clarsimp+) apply(case_tac xb, clarsimp+) apply(case_tac xa, clarsimp+) apply(simp add:\\<^sub>C_def) done context sifum_refinement_eg1_eg2 begin lemma bisim_simple_\: "bisim_simple (A.\ \ \ P)" unfolding bisim_simple_def apply clarsimp apply(drule_tac lc="\c\<^sub>1\<^sub>A, mds, mem\<^sub>1\<^sub>A\\<^sub>A" and lc'="\c\<^sub>2\<^sub>A, mds, mem\<^sub>2\<^sub>A\\<^sub>A" in A.bisim_simple_\\<^sub>u) by simp (* Don't know to what extent these helpers can be made generic enough to go into any parent theories. Again, mightn't be able to (or bother to) make some generic considering the aexp evaluator is going to be specific to each language. *) lemma conc_only_vars_not_visible_abs: "(\v\<^sub>C. v\<^sub>C \ range Eg2_var\<^sub>C_of_Eg1 \ mem\<^sub>C v\<^sub>C = mem\<^sub>C' v\<^sub>C) \ mem\<^sub>A_of mem\<^sub>C = mem\<^sub>A_of mem\<^sub>C'" by (simp add: mem\<^sub>A_of_def) lemma conc_only_var_assign_not_visible_abs: "\v\<^sub>C e. v\<^sub>C \ range Eg2_var\<^sub>C_of_Eg1 \ mem\<^sub>A_of mem\<^sub>C = mem\<^sub>A_of (mem\<^sub>C(v\<^sub>C := e))" using conc_only_vars_not_visible_abs by simp lemma reg\<^sub>C_is_not_the_var\<^sub>C_of_anything: "reg\<^sub>C = Eg2_var\<^sub>C_of_Eg1 x \ False" by (induct x, clarsimp+) lemma reg\<^sub>C_not_visible_abs: "reg\<^sub>C \ range Eg2_var\<^sub>C_of_Eg1" using reg\<^sub>C_is_not_the_var\<^sub>C_of_anything by blast (* This one's pretty specific to this refinement... *) lemma reg\<^sub>C_the_only_concrete_only_var: "v\<^sub>C \ range Eg2_var\<^sub>C_of_Eg1 \ v\<^sub>C = reg\<^sub>C" apply(case_tac v\<^sub>C) apply(erule rev_notE, clarsimp, rule_tac x=control_var in range_eqI, clarsimp) apply(erule rev_notE, clarsimp, rule_tac x=buffer in range_eqI, clarsimp) apply(erule rev_notE, clarsimp, rule_tac x=high_var in range_eqI, clarsimp) apply(erule rev_notE, clarsimp, rule_tac x=low_var in range_eqI, clarsimp) apply(erule rev_notE, clarsimp, rule_tac x=temp in range_eqI, clarsimp) apply clarsimp done lemma NoRW\<^sub>A_implies_NoRW\<^sub>C: "x \ mds\<^sub>A_of mds\<^sub>C AsmNoReadOrWrite \ Eg2_var\<^sub>C_of_Eg1 x \ mds\<^sub>C AsmNoReadOrWrite" unfolding mds\<^sub>A_of_def apply clarsimp apply (simp only: Eg2_var\<^sub>C_of_Eg1_def) apply clarsimp apply (simp add: f_inv_into_f) done lemma NoWrite\<^sub>A_implies_NoWrite\<^sub>C: "x \ mds\<^sub>A_of mds\<^sub>C AsmNoWrite \ Eg2_var\<^sub>C_of_Eg1 x \ mds\<^sub>C AsmNoWrite" unfolding mds\<^sub>A_of_def apply clarsimp apply (simp only: Eg2_var\<^sub>C_of_Eg1_def) apply clarsimp apply (simp add: f_inv_into_f) done lemma assign_eval\<^sub>w_load\<^sub>A: shows "(\x \ Eg1.Load y, mds, mem\\<^sub>A, \Stop, mds, mem (x := mem y)\\<^sub>A) \ A.eval\<^sub>w" by (metis A.assign_eval\<^sub>w ev\<^sub>A.simps(2)) lemma assign_eval\<^sub>w_load\<^sub>C: shows "(\x \ Load y, mds, mem\\<^sub>C, \Stop, mds, mem (x := mem y)\\<^sub>C) \ C.eval\<^sub>w" using C.unannotated[OF C.assign, where E="[]", simplified] apply(drule_tac x=x in meta_spec) apply(drule_tac x="Load y" in meta_spec) apply(drule_tac x=mds in meta_spec) apply(drule_tac x=mem in meta_spec) apply clarsimp done lemma assign_eval\<^sub>w_const\<^sub>A: shows "(\x \ Eg1.Const c, mds, mem\, \Stop, mds, mem (x := c)\) \ A.eval\<^sub>w" by (metis A.assign_eval\<^sub>w ev\<^sub>A.simps(1)) lemma assign_eval\<^sub>w_const\<^sub>C: shows "(\x \ Const c, mds, mem\, \Stop, mds, mem (x := c)\) \ C.eval\<^sub>w" using C.unannotated[OF C.assign, where E="[]", simplified] apply(drule_tac x=x in meta_spec) apply(drule_tac x="Const c" in meta_spec) apply(drule_tac x=mds in meta_spec) apply(drule_tac x=mem in meta_spec) apply clarsimp done lemma if_seq_eval\<^sub>w_helper\<^sub>A: "(\If B T E, mds, mem\, \if ev\<^sub>B mem B then T else E, mds, mem\\<^sub>A) \ A.eval\<^sub>w \ (\If B T E ;; TAIL, mds, mem\, \if ev\<^sub>B mem B then T ;; TAIL else E ;; TAIL, mds, mem\\<^sub>A) \ A.eval\<^sub>w" using A.eval\<^sub>w.seq by auto lemma if_seq_eval\<^sub>w_helper\<^sub>C: "(\If B T E, mds, mem\, \if ev\<^sub>B\<^sub>C mem B then T else E, mds, mem\\<^sub>C) \ C.eval\<^sub>w \ (\If B T E ;; TAIL, mds, mem\, \if ev\<^sub>B\<^sub>C mem B then T ;; TAIL else E ;; TAIL, mds, mem\\<^sub>C) \ C.eval\<^sub>w" using C.eval\<^sub>w.seq by auto lemma mem_assign_refinement_helper_var: "mem\<^sub>A_of (mem\<^sub>C (Eg2_var\<^sub>C_of_Eg1 x := mem\<^sub>C (Eg2_var\<^sub>C_of_Eg1 y))) = (mem\<^sub>A_of mem\<^sub>C) (x := (mem\<^sub>A_of mem\<^sub>C) y)" apply(clarsimp simp: mem\<^sub>A_of_def) apply(rule ext, clarsimp) apply(cases x) apply(case_tac x\<^sub>A, clarsimp+)+ done lemma mem_assign_refinement_helper_const: "mem\<^sub>A_of (mem\<^sub>C (Eg2_var\<^sub>C_of_Eg1 x := c)) = (mem\<^sub>A_of mem\<^sub>C) (x := c)" apply(clarsimp simp: mem\<^sub>A_of_def) apply(rule ext, clarsimp) apply(cases x) apply(case_tac x\<^sub>A, clarsimp+)+ done lemma if_true_eval\<^sub>w\<^sub>C: shows "mem\<^sub>C x = 0 \ (\(If (Eq x 0) c\<^sub>C_then c\<^sub>C_else) ;; c\<^sub>C_tail, mds\<^sub>C, mem\<^sub>C\\<^sub>C, \(c\<^sub>C_then ;; c\<^sub>C_tail), mds\<^sub>C, mem\<^sub>C\\<^sub>C) \ C.eval\<^sub>w" using C.if_eval\<^sub>w C.eval\<^sub>w.seq ev\<^sub>B\<^sub>C.simps by presburger lemma if_false_eval\<^sub>w\<^sub>C: shows "mem\<^sub>C x \ 0 \ (\(If (Eq x 0) c\<^sub>C_then c\<^sub>C_else) ;; c\<^sub>C_tail, mds\<^sub>C, mem\<^sub>C\\<^sub>C, \(c\<^sub>C_else ;; c\<^sub>C_tail), mds\<^sub>C, mem\<^sub>C\\<^sub>C) \ C.eval\<^sub>w" using C.if_eval\<^sub>w C.eval\<^sub>w.seq ev\<^sub>B\<^sub>C.simps by presburger end end diff --git a/thys/Flyspeck-Tame/FaceDivisionProps.thy b/thys/Flyspeck-Tame/FaceDivisionProps.thy --- a/thys/Flyspeck-Tame/FaceDivisionProps.thy +++ b/thys/Flyspeck-Tame/FaceDivisionProps.thy @@ -1,4955 +1,4955 @@ (* Author: Gertrud Bauer, Tobias Nipkow *) section\Properties of Face Division\ theory FaceDivisionProps imports Plane EnumeratorProps begin subsection\Finality\ (********************* makeFaceFinal ****************************) lemma vertices_makeFaceFinal: "vertices(makeFaceFinal f g) = vertices g" by(induct g)(simp add:vertices_graph_def makeFaceFinal_def) lemma edges_makeFaceFinal: "\ (makeFaceFinal f g) = \ g" proof - { fix fs have "(\\<^bsub>f\set (makeFaceFinalFaceList f fs)\<^esub> edges f) = (\\<^bsub>f\ set fs\<^esub> edges f)" apply(unfold makeFaceFinalFaceList_def) apply(induct f) by(induct fs) simp_all } thus ?thesis by(simp add:edges_graph_def makeFaceFinal_def) qed lemma in_set_repl_setFin: "f \ set fs \ final f \ f \ set (replace f' [setFinal f'] fs)" by (induct fs) auto lemma in_set_repl: "f \ set fs \ f \ f' \ f \ set (replace f' fs' fs)" by (induct fs) auto lemma makeFaceFinals_preserve_finals: "f \ set (finals g) \ f \ set (finals (makeFaceFinal f' g))" by (induct g) (simp add:makeFaceFinal_def finals_def makeFaceFinalFaceList_def in_set_repl_setFin) lemma len_faces_makeFaceFinal[simp]: "|faces (makeFaceFinal f g)| = |faces g|" by(simp add:makeFaceFinal_def makeFaceFinalFaceList_def) lemma len_finals_makeFaceFinal: "f \ \ g \ \ final f \ |finals (makeFaceFinal f g)| = |finals g| + 1" by(simp add:makeFaceFinal_def finals_def makeFaceFinalFaceList_def length_filter_replace1) lemma len_nonFinals_makeFaceFinal: "\ \ final f; f \ \ g\ \ |nonFinals (makeFaceFinal f g)| = |nonFinals g| - 1" by(simp add:makeFaceFinal_def nonFinals_def makeFaceFinalFaceList_def length_filter_replace2) lemma set_finals_makeFaceFinal[simp]: "distinct(faces g) \ f \ \ g \ set(finals (makeFaceFinal f g)) = insert (setFinal f) (set(finals g))" by(auto simp:finals_def makeFaceFinal_def makeFaceFinalFaceList_def distinct_set_replace) lemma splitFace_preserve_final: "f \ set (finals g) \ \ final f' \ f \ set (finals (snd (snd (splitFace g i j f' ns))))" by (induct g) (auto simp add: splitFace_def finals_def split_def intro: in_set_repl) lemma splitFace_nonFinal_face: "\ final (fst (snd (splitFace g i j f' ns)))" by (simp add: splitFace_def split_def split_face_def) lemma subdivFace'_preserve_finals: "\n i f' g. f \ set (finals g) \ \ final f' \ f \ set (finals (subdivFace' g f' i n is))" proof (induct "is") case Nil then show ?case by(simp add:makeFaceFinals_preserve_finals) next case (Cons j "js") then show ?case proof (cases j) case None with Cons show ?thesis by simp next case (Some sj) with Cons show ?thesis by (auto simp: splitFace_preserve_final splitFace_nonFinal_face split_def) qed qed lemma subdivFace_pres_finals: "f \ set (finals g) \ \ final f' \ f \ set (finals (subdivFace g f' is))" by(simp add:subdivFace_def subdivFace'_preserve_finals) declare Nat.diff_is_0_eq' [simp del] (********************* section is_prefix ****************************) subsection \\is_prefix\\ definition is_prefix :: "'a list \ 'a list \ bool" where "is_prefix ls vs \ (\ bs. vs = ls @ bs)" lemma is_prefix_add: "is_prefix ls vs \ is_prefix (as @ ls) (as @ vs)" by (simp add: is_prefix_def) lemma is_prefix_hd[simp]: "is_prefix [l] vs = (l = hd vs \ vs \ [])" apply (rule iffI) apply (auto simp: is_prefix_def) apply (intro exI) apply (subgoal_tac "vs = hd vs # tl vs") apply assumption by auto lemma is_prefix_f[simp]: "is_prefix (a#as) (a#vs) = is_prefix as vs" by (auto simp: is_prefix_def) lemma splitAt_is_prefix: "ram \ set vs \ is_prefix (fst (splitAt ram vs) @ [ram]) vs" by (auto dest!: splitAt_ram simp: is_prefix_def) (******************** section is_sublist *********************************) subsection \\is_sublist\\ definition is_sublist :: "'a list \ 'a list \ bool" where "is_sublist ls vs \ (\ as bs. vs = as @ ls @ bs)" lemma is_prefix_sublist: "is_prefix ls vs \ is_sublist ls vs" by (auto simp: is_prefix_def is_sublist_def) lemma is_sublist_trans: "is_sublist as bs \ is_sublist bs cs \ is_sublist as cs" apply (simp add: is_sublist_def) apply (elim exE) apply (subgoal_tac "cs = (asaa @ asa) @ as @ (bsa @ bsaa)") apply (intro exI) apply assumption by force lemma is_sublist_add: "is_sublist as bs \ is_sublist as (xs @ bs @ ys)" apply (simp add: is_sublist_def) apply (elim exE) apply (subgoal_tac "xs @ bs @ ys = (xs @ asa) @ as @ (bsa @ ys)") apply (intro exI) apply assumption by auto lemma is_sublist_rec: "is_sublist xs ys = (if length xs > length ys then False else if xs = take (length xs) ys then True else is_sublist xs (tl ys))" proof (simp add:is_sublist_def, goal_cases) case 1 show ?case proof (standard, goal_cases) case 1 show ?case proof (standard, goal_cases) case xs: 1 show ?case proof (standard, goal_cases) case 1 show ?case by auto next case 2 show ?case proof (standard, goal_cases) case 1 have "ys = take |xs| ys @ drop |xs| ys" by simp also have "\ = [] @ xs @ drop |xs| ys" by(simp add:xs[symmetric]) finally show ?case by blast qed qed qed next case 2 show ?case proof (standard, goal_cases) case xs_neq: 1 show ?case proof (standard, goal_cases) case 1 show ?case by auto next case 2 show ?case proof (standard, goal_cases) case not_less: 1 show ?case proof (standard, goal_cases) case 1 then obtain as bs where ys: "ys = as @ xs @ bs" by blast have "as \ []" using xs_neq ys by auto then obtain a as' where "as = a # as'" by (simp add:neq_Nil_conv) blast hence "tl ys = as' @ xs @ bs" by(simp add:ys) thus ?case by blast next case 2 then obtain as bs where ys: "tl ys = as @ xs @ bs" by blast have "ys \ []" using xs_neq not_less by auto then obtain y ys' where "ys = y # ys'" by (simp add:neq_Nil_conv) blast hence "ys = (y#as) @ xs @ bs" using ys by simp thus ?case by blast qed qed qed qed qed qed lemma not_sublist_len[simp]: "|ys| < |xs| \ \ is_sublist xs ys" by(simp add:is_sublist_rec) lemma is_sublist_simp[simp]: "a \ v \ is_sublist (a#as) (v#vs) = is_sublist (a#as) vs" proof assume av: "a \ v" and subl: "is_sublist (a # as) (v # vs)" then obtain rs ts where vvs: "v#vs = rs @ (a # as) @ ts" by (auto simp: is_sublist_def) with av have "rs \ []" by auto with vvs have "tl (v#vs) = tl rs @ a # as @ ts" by auto then have "vs = tl rs @ a # as @ ts" by auto then show "is_sublist (a # as) vs" by (auto simp: is_sublist_def) next assume av: "a \ v" and subl: "is_sublist (a # as) vs" then show "is_sublist (a # as) (v # vs)" apply (auto simp: is_sublist_def) apply (intro exI) apply (subgoal_tac "v # asa @ a # as @ bs = (v # asa) @ a # as @ bs") apply assumption by auto qed lemma is_sublist_id[simp]: "is_sublist vs vs" apply (auto simp: is_sublist_def) apply (intro exI) apply (subgoal_tac "vs = [] @ vs @ []") by (assumption) auto lemma is_sublist_in: "is_sublist (a#as) vs \ a \ set vs" by (auto simp: is_sublist_def) lemma is_sublist_in1: "is_sublist [x,y] vs \ y \ set vs" by (auto simp: is_sublist_def) lemma is_sublist_notlast[simp]: "distinct vs \ x = last vs \ \ is_sublist [x,y] vs" proof assume dvs: "distinct vs" and xl: "x = last vs" and subl:"is_sublist [x, y] vs" then obtain rs ts where vs: "vs = rs @ x # y # ts" by (auto simp: is_sublist_def) define as where "as = rs @ [x]" define bs where "bs = y # ts" then have bsne: "bs \ []" by auto from as_def bs_def have vs2: "vs = as @ bs" using vs by auto with as_def have xas: "x \ set as" by auto from bsne vs2 have "last vs = last bs" by auto with xl have "x = last bs" by auto with bsne have "bs = (butlast bs) @ [x]" by auto then have "x \ set bs" by (induct bs) auto with xas vs2 dvs show False by auto qed lemma is_sublist_nth1: "is_sublist [x,y] ls \ \ i j. i < length ls \ j < length ls \ ls!i = x \ ls!j = y \ Suc i = j" proof - assume subl: "is_sublist [x,y] ls" then obtain as bs where "ls = as @ x # y # bs" by (auto simp: is_sublist_def) then have "(length as) < length ls \ (Suc (length as)) < length ls \ ls!(length as) = x \ ls!(Suc (length as)) = y \ Suc (length as) = (Suc (length as))" apply auto apply hypsubst_thin apply (induct as) by auto then show ?thesis by auto qed lemma is_sublist_nth2: "\ i j. i < length ls \ j < length ls \ ls!i = x \ ls!j = y \ Suc i = j \ is_sublist [x,y] ls " proof - assume "\ i j. i < length ls \ j < length ls \ ls!i = x \ ls!j = y \ Suc i = j" then obtain i j where vors: "i < length ls \ j < length ls \ ls!i = x \ ls!j = y \ Suc i = j" by auto then have "ls = take (Suc (Suc i)) ls @ drop (Suc (Suc i)) ls" by auto with vors have "ls = take (Suc i) ls @ [ls! (Suc i)] @ drop (Suc (Suc i)) ls" by (auto simp: take_Suc_conv_app_nth) with vors have "ls = take i ls @ [ls!i] @ [ls! (Suc i)] @ drop (Suc (Suc i)) ls" by (auto simp: take_Suc_conv_app_nth) with vors show ?thesis by (auto simp: is_sublist_def) qed lemma is_sublist_tl: "is_sublist (a # as) vs \ is_sublist as vs" apply (simp add: is_sublist_def) apply (elim exE) apply (intro exI) apply (subgoal_tac "vs = (asa @ [a]) @ as @ bs") apply assumption by auto lemma is_sublist_hd: "is_sublist (a # as) vs \ is_sublist [a] vs" apply (simp add: is_sublist_def) by auto lemma is_sublist_hd_eq[simp]: "(is_sublist [a] vs) = (a \ set vs)" apply (rule_tac iffI) apply (simp add: is_sublist_def) apply force apply (simp add: is_sublist_def) apply (induct vs) apply force apply (case_tac "a = aa") apply force apply (subgoal_tac "a \ set vs") apply simp apply (elim exE) apply (intro exI) apply (subgoal_tac "aa # vs = (aa # as) @ a # bs") apply (assumption) by auto lemma is_sublist_distinct_prefix: "is_sublist (v#as) (v # vs) \ distinct (v # vs) \ is_prefix as vs" proof - assume d: "distinct (v # vs)" and subl: "is_sublist (v # as) (v # vs)" from subl obtain rs ts where v_vs: "v # vs = rs @ (v # as) @ ts" by (simp add: is_sublist_def) auto from d have v: "v \ set vs" by auto then have "\ is_sublist (v # as) vs" by (auto dest: is_sublist_hd) with v_vs have "rs = []" apply (cases rs) by (auto simp: is_sublist_def) with v_vs show "is_prefix as vs" by (auto simp: is_prefix_def) qed lemma is_sublist_distinct[intro]: "is_sublist as vs \ distinct vs \ distinct as" by (auto simp: is_sublist_def) lemma is_sublist_y_hd: "distinct vs \ y = hd vs \ \ is_sublist [x,y] vs" proof assume d: "distinct vs" and yh: "y = hd vs" and subl: "is_sublist [x, y] vs" then obtain rs ts where vs: "vs = rs @ x # y # ts" by (auto simp: is_sublist_def) define as where "as = rs @ [x]" then have asne: "as \ []" by auto define bs where "bs = y # ts" then have bsne: "bs \ []" by auto from as_def bs_def have vs2: "vs = as @ bs" using vs by auto from bs_def have xbs: "y \ set bs" by auto from vs2 asne have "hd vs = hd as" by simp with yh have "y = hd as" by auto with asne have "y \ set as" by (induct as) auto with d xbs vs2 show False by auto qed lemma is_sublist_at1: "distinct (as @ bs) \ is_sublist [x,y] (as @ bs) \ x \ (last as) \ is_sublist [x,y] as \ is_sublist [x,y] bs" proof (cases "x \ set as") assume d: "distinct (as @ bs)" and subl: "is_sublist [x, y] (as @ bs)" and xnl: "x \ last as" define vs where "vs = as @ bs" with d have dvs: "distinct vs" by auto case True with xnl subl have ind: "is_sublist (as@bs) vs \ is_sublist [x, y] as" proof (induct as) case Nil then show ?case by force next case (Cons a as) assume ih: "\is_sublist (as@bs) vs; x \ last as; is_sublist [x,y] (as @ bs); x \ set as\ \ is_sublist [x, y] as" and subl_aas_vs: "is_sublist ((a # as) @ bs) vs" and xnl2: "x \ last (a # as)" and subl2: "is_sublist [x, y] ((a # as) @ bs)" and x: "x \ set (a # as)" then have rule1: "x \ a \ is_sublist [x,y] as" apply (cases "as = []") apply simp apply (rule_tac ih) by (auto dest: is_sublist_tl) from dvs subl_aas_vs have daas: "distinct (a # as @ bs)" apply (rule_tac is_sublist_distinct) by auto from xnl2 have asne: "x = a \ as \ []" by auto with subl2 daas have yhdas: "x = a \ y = hd as" apply simp apply (drule_tac is_sublist_distinct_prefix) by auto with asne have "x = a \ as = y # tl as" by auto with asne yhdas have "x = a \ is_prefix [x,y] (a # as)" by auto then have rule2: "x = a \ is_sublist [x,y] (a # as)" by (simp add: is_prefix_sublist) from rule1 rule2 show ?case by (cases "x = a") auto qed from vs_def d have "is_sublist [x, y] as" by (rule_tac ind) auto then show ?thesis by auto next assume d: "distinct (as @ bs)" and subl: "is_sublist [x, y] (as @ bs)" and xnl: "x \ last as" define ars where "ars = as" case False with ars_def have xars: "x \ set ars" by auto from subl have ind: "is_sublist as ars \ is_sublist [x, y] bs" proof (induct as) case Nil then show ?case by auto next case (Cons a as) assume ih: "\is_sublist as ars; is_sublist [x, y] (as @ bs)\ \ is_sublist [x, y] bs" and subl_aasbsvs: "is_sublist (a # as) ars" and subl2: "is_sublist [x, y] ((a # as) @ bs)" from subl_aasbsvs ars_def False have "x \ a" by (auto simp:is_sublist_in) with subl_aasbsvs subl2 show ?thesis apply (rule_tac ih) by (auto dest: is_sublist_tl) qed from ars_def have "is_sublist [x, y] bs" by (rule_tac ind) auto then show ?thesis by auto qed lemma is_sublist_at4: "distinct (as @ bs) \ is_sublist [x,y] (as @ bs) \ as \ [] \ x = last as \ y = hd bs" proof - assume d: "distinct (as @ bs)" and subl: "is_sublist [x,y] (as @ bs)" and asne: "as \ []" and xl: "x = last as" define vs where "vs = as @ bs" with subl have "is_sublist [x,y] vs" by auto then obtain rs ts where vs2: "vs = rs @ x # y # ts" by (auto simp: is_sublist_def) from vs_def d have dvs:"distinct vs" by auto from asne xl have as:"as = butlast as @ [x]" by auto with vs_def have vs3: "vs = butlast as @ x # bs" by auto from dvs vs2 vs3 have "rs = butlast as" apply (rule_tac dist_at1) by auto then have "rs @ [x] = butlast as @ [x]" by auto with as have "rs @ [x] = as" by auto then have "as = rs @ [x]" by auto with vs2 vs_def have "bs = y # ts" by auto then show ?thesis by auto qed lemma is_sublist_at5: "distinct (as @ bs) \ is_sublist [x,y] (as @ bs) \ is_sublist [x,y] as \ is_sublist [x,y] bs \ x = last as \ y = hd bs" apply (case_tac "as = []") apply simp apply (cases "x = last as") apply (subgoal_tac "y = hd bs") apply simp apply (rule is_sublist_at4) apply assumption+ (*apply force+ *) apply (drule_tac is_sublist_at1) by auto lemma is_sublist_rev: "is_sublist [a,b] (rev zs) = is_sublist [b,a] zs" apply (simp add: is_sublist_def) apply (intro iffI) apply (elim exE) apply (intro exI) apply (subgoal_tac "zs = (rev bs) @ b # a # rev as") apply assumption apply (subgoal_tac "rev (rev zs) = rev (as @ a # b # bs)") apply (thin_tac "rev zs = as @ a # b # bs") apply simp apply simp apply (elim exE) apply (intro exI) by force lemma is_sublist_at5'[simp]: "distinct as \ distinct bs \ set as \ set bs = {} \ is_sublist [x,y] (as @ bs) \ is_sublist [x,y] as \ is_sublist [x,y] bs \ x = last as \ y = hd bs" apply (subgoal_tac "distinct (as @ bs)") apply (drule is_sublist_at5) by auto lemma splitAt_is_sublist1R[simp]: "ram \ set vs \ is_sublist (fst (splitAt ram vs) @ [ram]) vs" apply (auto dest!: splitAt_ram simp: is_sublist_def) apply (intro exI) apply (subgoal_tac "vs = [] @ fst (splitAt ram vs) @ ram # snd (splitAt ram vs)") apply assumption by simp lemma splitAt_is_sublist2R[simp]: "ram \ set vs \ is_sublist (ram # snd (splitAt ram vs)) vs" apply (auto dest!: splitAt_ram splitAt_no_ram simp: is_sublist_def) apply (intro exI) apply (subgoal_tac "vs = fst (splitAt ram vs) @ ram # snd (splitAt ram vs) @ []") apply assumption by auto (*********************** section is_nextElem *****************************) subsection \\is_nextElem\\ definition is_nextElem :: "'a list \ 'a \ 'a \ bool" where "is_nextElem xs x y \ is_sublist [x,y] xs \ xs \ [] \ x = last xs \ y = hd xs" lemma is_nextElem_a[intro]: "is_nextElem vs a b \ a \ set vs" by (auto simp: is_nextElem_def is_sublist_def) lemma is_nextElem_b[intro]: "is_nextElem vs a b \ b \ set vs" by (auto simp: is_nextElem_def is_sublist_def) lemma is_nextElem_last_hd[intro]: "distinct vs \ is_nextElem vs x y \ x = last vs \ y = hd vs" by (auto simp: is_nextElem_def) lemma is_nextElem_last_ne[intro]: "distinct vs \ is_nextElem vs x y \ x = last vs \ vs \ []" by (auto simp: is_nextElem_def) lemma is_nextElem_sublistI: "is_sublist [x,y] vs \ is_nextElem vs x y" by (auto simp: is_nextElem_def) lemma is_nextElem_nth1: "is_nextElem ls x y \ \ i j. i < length ls \ j < length ls \ ls!i = x \ ls!j = y \ (Suc i) mod (length ls) = j" proof (cases "is_sublist [x,y] ls") assume is_nextElem: "is_nextElem ls x y" case True then show ?thesis apply (drule_tac is_sublist_nth1) by auto next assume is_nextElem: "is_nextElem ls x y" case False with is_nextElem have hl: "ls \ [] \ last ls = x \ hd ls = y" by (auto simp: is_nextElem_def) then have j: "ls!0 = y" by (cases ls) auto from hl have i: "ls!(length ls - 1) = x" by (cases ls rule: rev_exhaust) auto from i j hl have "(length ls - 1) < length ls \ 0 < length ls \ ls!(length ls - 1) = x \ ls!0 = y \ (Suc (length ls - 1)) mod (length ls) = 0" by auto then show ?thesis apply (intro exI) . qed lemma is_nextElem_nth2: " \ i j. i < length ls \ j < length ls \ ls!i = x \ ls!j = y \ (Suc i) mod (length ls) = j \ is_nextElem ls x y" proof - assume "\ i j. i < length ls \ j < length ls \ ls!i = x \ ls!j = y \ (Suc i) mod (length ls) = j" then obtain i j where vors: "i < length ls \ j < length ls \ ls!i = x \ ls!j = y \ (Suc i) mod (length ls) = j" by auto then show ?thesis proof (cases "Suc i = length ls") case True with vors have "j = 0" by auto (*ls ! i = last ls*) with True vors show ?thesis apply (auto simp: is_nextElem_def) apply (cases ls rule: rev_exhaust) apply auto apply (cases ls) by auto next case False with vors have "is_sublist [x,y] ls" apply (rule_tac is_sublist_nth2) by auto then show ?thesis by (simp add: is_nextElem_def) qed qed lemma is_nextElem_rotate1_aux: "is_nextElem (rotate m ls) x y \ is_nextElem ls x y" proof - assume is_nextElem: "is_nextElem (rotate m ls) x y" define n where "n = m mod length ls" then have rot_eq: "rotate m ls = rotate n ls" by (auto intro: rotate_conv_mod) with is_nextElem have "is_nextElem (rotate n ls) x y" by simp then obtain i j where vors:"i < length (rotate n ls) \ j < length (rotate n ls) \ (rotate n ls)!i = x \ (rotate n ls)!j = y \ (Suc i) mod (length (rotate n ls)) = j" by (drule_tac is_nextElem_nth1) auto then have lls: "0 < length ls" by auto define k where "k = (i+n) mod (length ls)" with lls have sk: "k < length ls" by simp from k_def lls vors have "ls!k = (rotate n ls)!(i mod (length ls))" by (simp add: nth_rotate) with vors have lsk: "ls!k = x" by simp define l where "l = (j+n) mod (length ls)" with lls have sl: "l < length ls" by simp from l_def lls vors have "ls!l = (rotate n ls)!(j mod (length ls))" by (simp add: nth_rotate) with vors have lsl: "ls!l = y" by simp from vors k_def l_def have "(Suc i) mod length ls = j" by simp then have "(Suc i) mod length ls = j mod length ls" by auto then have "((Suc i) mod length ls + n mod (length ls)) mod length ls = (j mod length ls + n mod (length ls)) mod length ls" by simp then have "((Suc i) + n) mod length ls = (j + n) mod length ls" by (simp add: mod_simps) with vors k_def l_def have "(Suc k) mod (length ls) = l" by (simp add: mod_simps) with sk lsk sl lsl show ?thesis by (auto intro: is_nextElem_nth2) qed lemma is_nextElem_rotate_eq[simp]: "is_nextElem (rotate m ls) x y = is_nextElem ls x y" apply (auto dest: is_nextElem_rotate1_aux) apply (rule is_nextElem_rotate1_aux) apply (subgoal_tac "is_nextElem (rotate (length ls - m mod length ls) (rotate m ls)) x y") apply assumption by simp lemma is_nextElem_congs_eq: "ls \ ms \ is_nextElem ls x y = is_nextElem ms x y" by (auto simp: congs_def) lemma is_nextElem_rev[simp]: "is_nextElem (rev zs) a b = is_nextElem zs b a" apply (simp add: is_nextElem_def is_sublist_rev) apply (case_tac "zs = []") apply simp apply simp apply (case_tac "a = hd zs") apply (case_tac "zs") apply simp apply simp apply simp apply (case_tac "a = last (rev zs) \ b = last zs") apply simp apply (case_tac "zs" rule: rev_exhaust) apply simp apply (case_tac "ys") apply simp apply simp by force lemma is_nextElem_circ: "\ distinct xs; is_nextElem xs a b; is_nextElem xs b a \ \ |xs| \ 2" apply(drule is_nextElem_nth1) apply(drule is_nextElem_nth1) apply (clarsimp) apply(rename_tac i j) apply(frule_tac i=j and j = "Suc i mod |xs|" in nth_eq_iff_index_eq) apply assumption+ apply(frule_tac j=i and i = "Suc j mod |xs|" in nth_eq_iff_index_eq) apply assumption+ apply(rule ccontr) apply(simp add: distinct_conv_nth mod_Suc) done subsection\\nextElem, sublist, is_nextElem\\ lemma is_sublist_eq: "distinct vs \ c \ y \ (nextElem vs c x = y) = is_sublist [x,y] vs" proof - assume d: "distinct vs" and c: "c \ y" have r1: "nextElem vs c x = y \ is_sublist [x,y] vs" proof - assume fn: "nextElem vs c x = y" with c show ?thesis by(drule_tac nextElem_cases)(auto simp: is_sublist_def) qed with d have r2: "is_sublist [x,y] vs \ nextElem vs c x = y" apply (simp add: is_sublist_def) apply (elim exE) by auto show ?thesis apply (intro iffI r1) by (auto intro: r2) qed lemma is_nextElem1: "distinct vs \ x \ set vs \ nextElem vs (hd vs) x = y \ is_nextElem vs x y" proof - assume d: "distinct vs" and x: "x \ set vs" and fn: "nextElem vs (hd vs) x = y" from x have r0: "vs \ []" by auto from d fn have r1: "x = last vs \ y = hd vs" by (auto) from d fn have r3: "hd vs \ y \ (\ a b. vs = a @ [x,y] @ b)" by (drule_tac nextElem_cases) auto from x obtain n where xn:"x = vs!n" and nl: "n < length vs" by (auto simp: in_set_conv_nth) define as where "as = take n vs" define bs where "bs = drop (Suc n) vs" from as_def bs_def xn nl have vs:"vs = as @ [x] @ bs" by (auto intro: id_take_nth_drop) then have r2: "x \ last vs \ y \ hd vs" proof - assume notx: "x \ last vs" from vs notx have "bs \ []" by auto with vs have r2: "vs = as @ [x, hd bs] @ tl bs" by auto with d have ineq: "hd bs \ hd vs" by (cases as) auto from d fn r2 have "y = hd bs" by auto with ineq show ?thesis by auto qed from r0 r1 r2 r3 show ?thesis apply (simp add:is_nextElem_def is_sublist_def) apply (cases "x = last vs") by auto qed lemma is_nextElem2: "distinct vs \ x \ set vs \ is_nextElem vs x y \ nextElem vs (hd vs) x = y" proof - assume d: "distinct vs" and x: "x \ set vs" and is_nextElem: "is_nextElem vs x y" then show ?thesis apply (simp add: is_nextElem_def) apply (cases "is_sublist [x,y] vs") apply (cases "y = hd vs") apply (simp add: is_sublist_def) apply (force dest: distinct_hd_not_cons) apply (subgoal_tac "hd vs \ y") apply (simp add: is_sublist_eq) by auto qed lemma nextElem_is_nextElem: "distinct xs \ x \ set xs \ is_nextElem xs x y = (nextElem xs (hd xs) x = y)" by (auto intro!: is_nextElem1 is_nextElem2) lemma nextElem_congs_eq: "xs \ ys \ distinct xs \ x \ set xs \ nextElem xs (hd xs) x = nextElem ys (hd ys) x" proof - assume eq: "xs \ ys" and dist: "distinct xs" and x: "x \ set xs" define y where "y = nextElem xs (hd xs) x" then have f1:"nextElem xs (hd xs) x = y" by auto with dist x have "is_nextElem xs x y" by (auto intro: is_nextElem1) with eq have "is_nextElem ys x y" by (simp add:is_nextElem_congs_eq) with eq dist x have f2:"nextElem ys (hd ys) x = y" by (auto simp: congs_distinct intro: is_nextElem2) from f1 f2 show ?thesis by auto qed lemma is_sublist_is_nextElem: "distinct vs \ is_nextElem vs x y \ is_sublist as vs \ x \ set as \ x \ last as \ is_sublist [x,y] as" proof - assume d: "distinct vs" and is_nextElem: "is_nextElem vs x y" and subl: "is_sublist as vs" and xin: "x \ set as" and xnl: "x \ last as" from xin have asne: "as \ []" by auto with subl have vsne: "vs \ []" by (auto simp: is_sublist_def) from subl obtain rs ts where vs: "vs = rs @ as @ ts" apply (simp add: is_sublist_def) apply (elim exE) by auto with d xnl asne have "x \ last vs" proof (cases "ts = []") case True with d xnl asne vs show ?thesis by force next define lastvs where "lastvs = last ts" case False with vs lastvs_def have vs2: "vs = rs @ as @ butlast ts @ [lastvs]" by auto with d have "lastvs \ set as" by auto with xin have "lastvs \ x" by auto with vs2 show ?thesis by auto qed with is_nextElem have subl_vs: "is_sublist [x,y] vs" by (auto simp: is_nextElem_def) from d xin vs have "\ is_sublist [x] rs" by auto then have nrs: "\ is_sublist [x,y] rs" by (auto dest: is_sublist_hd) from d xin vs have "\ is_sublist [x] ts" by auto then have nts: "\ is_sublist [x,y] ts" by (auto dest: is_sublist_hd) from d xin vs have xnrs: "x \ set rs" by auto then have notrs: "\ is_sublist [x,y] rs" by (auto simp:is_sublist_in) from xnrs have xnlrs: "rs \ [] \ x \ last rs" by (induct rs) auto from d xin vs have xnts: "x \ set ts" by auto then have notts: "\ is_sublist [x,y] ts" by (auto simp:is_sublist_in) from d vs subl_vs have "is_sublist [x,y] rs \ is_sublist [x,y] (as@ts)" apply (cases "rs = []") apply simp apply (rule_tac is_sublist_at1) by (auto intro!: xnlrs) with notrs have "is_sublist [x,y] (as@ts)" by auto with d vs xnl have "is_sublist [x,y] as \ is_sublist [x,y] ts" apply (rule_tac is_sublist_at1) by auto with notts show "is_sublist [x,y] as" by auto qed subsection \\before\\ definition before :: "'a list \ 'a \ 'a \ bool" where "before vs ram1 ram2 \ \ a b c. vs = a @ ram1 # b @ ram2 # c" lemma before_dist_fst_fst[simp]: "before vs ram1 ram2 \ distinct vs \ fst (splitAt ram2 (fst (splitAt ram1 vs))) = fst (splitAt ram1 (fst (splitAt ram2 vs)))" apply (simp add: before_def) apply (elim exE) apply (drule splitAt_dist_ram_all) by (auto dest!: pairD) lemma before_dist_fst_snd[simp]: "before vs ram1 ram2 \ distinct vs \ fst (splitAt ram2 (snd (splitAt ram1 vs))) = snd (splitAt ram1 (fst (splitAt ram2 vs)))" apply (simp add: before_def) apply (elim exE) apply (drule_tac splitAt_dist_ram_all) by (auto dest!: pairD) lemma before_dist_snd_fst[simp]: "before vs ram1 ram2 \ distinct vs \ snd (splitAt ram2 (fst (splitAt ram1 vs))) = snd (splitAt ram1 (snd (splitAt ram2 vs)))" apply (simp add: before_def) apply (elim exE) apply (drule_tac splitAt_dist_ram_all) by (auto dest!: pairD) lemma before_dist_snd_snd[simp]: "before vs ram1 ram2 \ distinct vs \ snd (splitAt ram2 (snd (splitAt ram1 vs))) = fst (splitAt ram1 (snd (splitAt ram2 vs)))" apply (simp add: before_def) apply (elim exE) apply (drule_tac splitAt_dist_ram_all) by (auto dest!: pairD) lemma before_dist_snd[simp]: "before vs ram1 ram2 \ distinct vs \ fst (splitAt ram1 (snd (splitAt ram2 vs))) = snd (splitAt ram2 vs)" apply (simp add: before_def) apply (elim exE) apply (drule_tac splitAt_dist_ram_all) by (auto dest!: pairD) lemma before_dist_fst[simp]: "before vs ram1 ram2 \ distinct vs \ fst (splitAt ram1 (fst (splitAt ram2 vs))) = fst (splitAt ram1 vs)" apply (simp add: before_def) apply (elim exE) apply (drule_tac splitAt_dist_ram_all) by (auto dest!: pairD) lemma before_or: "ram1 \ set vs \ ram2 \ set vs \ ram1 \ ram2 \ before vs ram1 ram2 \ before vs ram2 ram1" proof - assume r1: "ram1 \ set vs" and r2: "ram2 \ set vs" and r12: "ram1 \ ram2" then show ?thesis proof (cases "ram2 \ set (snd (splitAt ram1 vs))") define a where "a = fst (splitAt ram1 vs)" define b where "b = fst (splitAt ram2 (snd (splitAt ram1 vs)))" define c where "c = snd (splitAt ram2 (snd (splitAt ram1 vs)))" case True with r1 a_def b_def c_def have "vs = a @ [ram1] @ b @ [ram2] @ c" by (auto dest!: splitAt_ram) then show ?thesis apply (simp add: before_def) by auto next define ab where "ab = fst (splitAt ram1 vs)" case False with r1 r2 r12 ab_def have r2': "ram2 \ set ab" by (auto intro: splitAt_ram3) define a where "a = fst (splitAt ram2 ab)" define b where "b = snd (splitAt ram2 ab)" define c where "c = snd (splitAt ram1 vs)" from r1 ab_def c_def have "vs = ab @ [ram1] @ c" by (auto dest!: splitAt_ram) with r2' a_def b_def have "vs = (a @ [ram2] @ b) @ [ram1] @ c" by (drule_tac splitAt_ram) simp then show ?thesis apply (simp add: before_def) apply (rule disjI2) by auto qed qed lemma before_r1: "before vs r1 r2 \ r1 \ set vs" by (auto simp: before_def) lemma before_r2: "before vs r1 r2 \ r2 \ set vs" by (auto simp: before_def) lemma before_dist_r2: "distinct vs \ before vs r1 r2 \ r2 \ set (snd (splitAt r1 vs))" proof - assume d: "distinct vs" and b: "before vs r1 r2" from d b have ex1: "\! s. (vs = (fst s) @ r1 # snd (s))" apply (drule_tac before_r1) apply (rule distinct_unique1) by auto from d b ex1 show ?thesis apply (unfold before_def) proof (elim exE ex1E) fix a b c s assume vs: "vs = a @ r1 # b @ r2 # c" and "\y. vs = fst y @ r1 # snd y \ y = s" then have "\ y. vs = fst y @ r1 # snd y \ y = s" by (clarify, hypsubst_thin, auto) then have single: "\ y. vs = fst y @ r1 # snd y \ y = s" by auto define bc where "bc = b @ r2 # c" with vs have vs2: "vs = a @ r1 # bc" by auto from bc_def have r2: "r2 \ set bc" by auto define t where "t = (a,bc)" with vs2 have vs3: "vs = fst (t) @ r1 # snd (t)" by auto with single have ts: "t = s" by (rule_tac single) auto from b have "splitAt r1 vs = s" apply (drule_tac before_r1) apply (drule_tac splitAt_ram) by (rule single) auto with ts have "t = splitAt r1 vs" by simp with t_def have "bc = snd(splitAt r1 vs)" by simp with r2 show ?thesis by simp qed qed lemma before_dist_not_r2[intro]: "distinct vs \ before vs r1 r2 \ r2 \ set (fst (splitAt r1 vs))" apply (frule before_dist_r2) by (auto dest: splitAt_distinct_fst_snd) lemma before_dist_r1: "distinct vs \ before vs r1 r2 \ r1 \ set (fst (splitAt r2 vs))" proof - assume d: "distinct vs" and b: "before vs r1 r2" from d b have ex1: "\! s. (vs = (fst s) @ r2 # snd (s))" apply (drule_tac before_r2) apply (rule distinct_unique1) by auto from d b ex1 show ?thesis apply (unfold before_def) proof (elim exE ex1E) fix a b c s assume vs: "vs = a @ r1 # b @ r2 # c" and "\y. vs = fst y @ r2 # snd y \ y = s" then have "\ y. vs = fst y @ r2 # snd y \ y = s" by (clarify, hypsubst_thin, auto) then have single: "\ y. vs = fst y @ r2 # snd y \ y = s" by auto define ab where "ab = a @ r1 # b" with vs have vs2: "vs = ab @ r2 # c" by auto from ab_def have r1: "r1 \ set ab" by auto define t where "t = (ab,c)" with vs2 have vs3: "vs = fst (t) @ r2 # snd (t)" by auto with single have ts: "t = s" by (rule_tac single) auto from b have "splitAt r2 vs = s" apply (drule_tac before_r2) apply (drule_tac splitAt_ram) by (rule single) auto with ts have "t = splitAt r2 vs" by simp with t_def have "ab = fst(splitAt r2 vs)" by simp with r1 show ?thesis by simp qed qed lemma before_dist_not_r1[intro]: "distinct vs \ before vs r1 r2 \ r1 \ set (snd (splitAt r2 vs))" apply (frule before_dist_r1) by (auto dest: splitAt_distinct_fst_snd) lemma before_snd: "r2 \ set (snd (splitAt r1 vs)) \ before vs r1 r2" proof - assume r2: "r2 \ set (snd (splitAt r1 vs))" from r2 have r1: "r1 \ set vs" apply (rule_tac ccontr) apply (drule splitAt_no_ram) by simp define a where "a = fst (splitAt r1 vs)" define bc where "bc = snd (splitAt r1 vs)" define b where "b = fst (splitAt r2 bc)" define c where "c = snd (splitAt r2 bc)" from r1 a_def bc_def have vs: "vs = a @ [r1] @ bc" by (auto dest: splitAt_ram) from r2 bc_def have r2: "r2 \ set bc" by simp with b_def c_def have "bc = b @ [r2] @ c" by (auto dest: splitAt_ram) with vs show ?thesis by (simp add: before_def) auto qed lemma before_fst: "r2 \ set vs \ r1 \ set (fst (splitAt r2 vs)) \ before vs r1 r2" proof - assume r2: "r2 \ set vs" and r1: "r1 \ set (fst (splitAt r2 vs))" define ab where "ab = fst (splitAt r2 vs)" define c where "c = snd (splitAt r2 vs)" define a where "a = fst (splitAt r1 ab)" define b where "b = snd (splitAt r1 ab)" from r2 ab_def c_def have vs: "vs = ab @ [r2] @ c" by (auto dest: splitAt_ram) from r1 ab_def have r1: "r1 \ set ab" by simp with a_def b_def have "ab = a @ [r1] @ b" by (auto dest: splitAt_ram) with vs show ?thesis by (simp add: before_def) auto qed (* usefule simplifier rules *) lemma before_dist_eq_fst: "distinct vs \ r2 \ set vs \ r1 \ set (fst (splitAt r2 vs)) = before vs r1 r2" by (auto intro: before_fst before_dist_r1) lemma before_dist_eq_snd: "distinct vs \ r2 \ set (snd (splitAt r1 vs)) = before vs r1 r2" by (auto intro: before_snd before_dist_r2) lemma before_dist_not1: "distinct vs \ before vs ram1 ram2 \ \ before vs ram2 ram1" proof assume d: "distinct vs" and b1: "before vs ram2 ram1" and b2: "before vs ram1 ram2" from b2 have r1: "ram1 \ set vs" by (drule_tac before_r1) from d b1 have r2: "ram2 \ set (fst (splitAt ram1 vs))" by (rule before_dist_r1) from d b2 have r2':"ram2 \ set (snd (splitAt ram1 vs))" by (rule before_dist_r2) from d r1 r2 r2' show "False" by (drule_tac splitAt_distinct_fst_snd) auto qed lemma before_dist_not2: "distinct vs \ ram1 \ set vs \ ram2 \ set vs \ ram1 \ ram2 \ \ (before vs ram1 ram2) \ before vs ram2 ram1" proof - assume "distinct vs" "ram1 \ set vs " "ram2 \ set vs" "ram1 \ ram2" "\ before vs ram1 ram2" then show "before vs ram2 ram1" apply (frule_tac before_or) by auto qed lemma before_dist_eq: "distinct vs \ ram1 \ set vs \ ram2 \ set vs \ ram1 \ ram2 \ ( \ (before vs ram1 ram2)) = before vs ram2 ram1" by (auto intro: before_dist_not2 dest: before_dist_not1) lemma before_vs: "distinct vs \ before vs ram1 ram2 \ vs = fst (splitAt ram1 vs) @ ram1 # fst (splitAt ram2 (snd (splitAt ram1 vs))) @ ram2 # snd (splitAt ram2 vs)" proof - assume d: "distinct vs" and b: "before vs ram1 ram2" define s where "s = snd (splitAt ram1 vs)" from b have "ram1 \ set vs" by (auto simp: before_def) with s_def have vs: "vs = fst (splitAt ram1 vs) @ [ram1] @ s" by (auto dest: splitAt_ram) from d b s_def have "ram2 \ set s" by (auto intro: before_dist_r2) then have snd: "s = fst (splitAt ram2 s) @ [ram2] @ snd (splitAt ram2 s)" by (auto dest: splitAt_ram) with vs have "vs = fst (splitAt ram1 vs) @ [ram1] @ fst (splitAt ram2 s) @ [ram2] @ snd (splitAt ram2 s)" by auto with d b s_def show ?thesis by auto qed (************************ between lemmas *************************************) subsection \@{const between}\ definition pre_between :: "'a list \ 'a \ 'a \ bool" where "pre_between vs ram1 ram2 \ distinct vs \ ram1 \ set vs \ ram2 \ set vs \ ram1 \ ram2" declare pre_between_def [simp] lemma pre_between_dist[intro]: "pre_between vs ram1 ram2 \ distinct vs" by (auto simp: pre_between_def) lemma pre_between_r1[intro]: "pre_between vs ram1 ram2 \ ram1 \ set vs" by auto lemma pre_between_r2[intro]: "pre_between vs ram1 ram2 \ ram2 \ set vs" by auto lemma pre_between_r12[intro]: "pre_between vs ram1 ram2 \ ram1 \ ram2" by auto lemma pre_between_symI: "pre_between vs ram1 ram2 \ pre_between vs ram2 ram1" by auto lemma pre_between_before[dest]: "pre_between vs ram1 ram2 \ before vs ram1 ram2 \ before vs ram2 ram1" by (rule_tac before_or) auto lemma pre_between_rotate1[intro]: "pre_between vs ram1 ram2 \ pre_between (rotate1 vs) ram1 ram2" by auto lemma pre_between_rotate[intro]: "pre_between vs ram1 ram2 \ pre_between (rotate n vs) ram1 ram2" by auto lemma(*<*) before_xor: (*>*) "pre_between vs ram1 ram2 \ (\ before vs ram1 ram2) = before vs ram2 ram1" by (simp add: before_dist_eq) declare pre_between_def [simp del] lemma between_simp1[simp]: "before vs ram1 ram2 \ pre_between vs ram1 ram2 \ between vs ram1 ram2 = fst (splitAt ram2 (snd (splitAt ram1 vs)))" by (simp add: pre_between_def between_def split_def before_dist_eq_snd) lemma between_simp2[simp]: "before vs ram1 ram2 \ pre_between vs ram1 ram2 \ between vs ram2 ram1 = snd (splitAt ram2 vs) @ fst (splitAt ram1 vs)" proof - assume b: "before vs ram1 ram2" and p: "pre_between vs ram1 ram2" from p b have b2: "\ before vs ram2 ram1" apply (simp add: pre_between_def) by (auto dest: before_dist_not1) with p have "ram2 \ set (fst (splitAt ram1 vs))" by (simp add: pre_between_def before_dist_eq_fst) then have "fst (splitAt ram1 vs) = fst (splitAt ram2 (fst (splitAt ram1 vs)))" by (auto dest: splitAt_no_ram) then have "fst (splitAt ram2 (fst (splitAt ram1 vs))) = fst (splitAt ram1 vs)" by auto with b2 b p show ?thesis apply (simp add: pre_between_def between_def split_def) by (auto dest: before_dist_not_r1) qed lemma between_not_r1[intro]: "distinct vs \ ram1 \ set (between vs ram1 ram2)" proof (cases "pre_between vs ram1 ram2") assume d: "distinct vs" case True then have p: "pre_between vs ram1 ram2" by auto then show "ram1 \ set (between vs ram1 ram2)" proof (cases "before vs ram1 ram2") case True with d p show ?thesis by (auto del: notI) next from p have p2: "pre_between vs ram2 ram1" by (auto intro: pre_between_symI) case False with p have "before vs ram2 ram1" by auto with d p2 show ?thesis by (auto del: notI) qed next assume d:"distinct vs" case False then have p: "\ pre_between vs ram1 ram2" by auto then show ?thesis proof (cases "ram1 = ram2") case True with d have h1:"ram2 \ set (snd (splitAt ram2 vs))" by (auto del: notI) from True d have h2: "ram2 \ set (fst (splitAt ram2 (fst (splitAt ram2 vs))))" by (auto del: notI) with True d h1 show ?thesis by (auto simp: between_def split_def) next case False then have neq: "ram1 \ ram2" by auto then show ?thesis proof (cases "ram1 \ set vs") case True with d show ?thesis by (auto dest: splitAt_no_ram splitAt_in_fst simp: between_def split_def) next case False then have r1in: "ram1 \ set vs" by auto then show ?thesis proof (cases "ram2 \ set vs") from d have h1: "ram1 \ set (fst (splitAt ram1 vs))" by (auto del: notI) case True with d h1 show ?thesis by (auto dest: splitAt_not1 splitAt_in_fst splitAt_ram splitAt_no_ram simp: between_def split_def del: notI) next case False then have r2in: "ram2 \ set vs" by auto with d neq r1in have "pre_between vs ram1 ram2" by (auto simp: pre_between_def) with p show ?thesis by auto qed qed qed qed lemma between_not_r2[intro]: "distinct vs \ ram2 \ set (between vs ram1 ram2)" proof (cases "pre_between vs ram1 ram2") assume d: "distinct vs" case True then have p: "pre_between vs ram1 ram2" by auto then show "ram2 \ set (between vs ram1 ram2)" proof (cases "before vs ram1 ram2") from d have "ram2 \ set (fst (splitAt ram2 vs))" by (auto del: notI) then have h1: "ram2 \ set (snd (splitAt ram1 (fst (splitAt ram2 vs))))" by (auto dest: splitAt_in_fst) case True with d p h1 show ?thesis by (auto del: notI) next from p have p2: "pre_between vs ram2 ram1" by (auto intro: pre_between_symI) case False with p have "before vs ram2 ram1" by auto with d p2 show ?thesis by (auto del: notI) qed next assume d:"distinct vs" case False then have p: "\ pre_between vs ram1 ram2" by auto then show ?thesis proof (cases "ram1 = ram2") case True with d have h1:"ram2 \ set (snd (splitAt ram2 vs))" by (auto del: notI) from True d have h2: "ram2 \ set (fst (splitAt ram2 (fst (splitAt ram2 vs))))" by (auto del: notI) with True d h1 show ?thesis by (auto simp: between_def split_def) next case False then have neq: "ram1 \ ram2" by auto then show ?thesis proof (cases "ram2 \ set vs") case True with d show ?thesis by (auto dest: splitAt_no_ram splitAt_in_fst splitAt_in_fst simp: between_def split_def) next case False then have r1in: "ram2 \ set vs" by auto then show ?thesis proof (cases "ram1 \ set vs") from d have h1: "ram1 \ set (fst (splitAt ram1 vs))" by (auto del: notI) case True with d h1 show ?thesis by (auto dest: splitAt_ram splitAt_no_ram simp: between_def split_def del: notI) next case False then have r2in: "ram1 \ set vs" by auto with d neq r1in have "pre_between vs ram1 ram2" by (auto simp: pre_between_def) with p show ?thesis by auto qed qed qed qed lemma between_distinct[intro]: "distinct vs \ distinct (between vs ram1 ram2)" proof - assume vs: "distinct vs" define a where "a = fst (splitAt ram1 vs)" define b where "b = snd (splitAt ram1 vs)" from a_def b_def have ab: "(a,b) = splitAt ram1 vs" by auto with vs have ab_disj:"set a \ set b = {}" by (drule_tac splitAt_distinct_ab) auto define c where "c = fst (splitAt ram2 a)" define d where "d = snd (splitAt ram2 a)" from c_def d_def have c_d: "(c,d) = splitAt ram2 a" by auto with ab_disj have "set c \ set b = {}" by (drule_tac splitAt_subset_ab) auto with vs a_def b_def c_def show ?thesis by (auto simp: between_def split_def splitAt_no_ram dest: splitAt_ram intro: splitAt_distinct_fst splitAt_distinct_snd) qed lemma between_distinct_r12: "distinct vs \ ram1 \ ram2 \ distinct (ram1 # between vs ram1 ram2 @ [ram2])" by (auto del: notI) lemma between_vs: "before vs ram1 ram2 \ pre_between vs ram1 ram2 \ vs = fst (splitAt ram1 vs) @ ram1 # (between vs ram1 ram2) @ ram2 # snd (splitAt ram2 vs)" apply (simp) apply (frule pre_between_dist) apply (drule before_vs) by auto lemma between_in: "before vs ram1 ram2 \ pre_between vs ram1 ram2 \ x \ set vs \ x = ram1 \ x \ set (between vs ram1 ram2) \ x = ram2 \ x \ set (between vs ram2 ram1)" proof - assume b: "before vs ram1 ram2" and p: "pre_between vs ram1 ram2" and xin: "x \ set vs" define a where "a = fst (splitAt ram1 vs)" define b where "b = between vs ram1 ram2" define c where "c = snd (splitAt ram2 vs)" from p have "distinct vs" by auto from p b a_def b_def c_def have "vs = a @ ram1 # b @ ram2 # c" apply (drule_tac between_vs) by auto with xin have "x \ set (a @ ram1 # b @ ram2 # c)" by auto then have "x \ set (a) \ x \ set (ram1 #b) \ x \ set (ram2 # c)" by auto then have "x \ set (a) \ x = ram1 \ x \ set b \ x = ram2 \ x \ set c" by auto then have "x \ set c \ x \ set (a) \ x = ram1 \ x \ set b \ x = ram2" by auto then have "x \ set (c @ a) \ x = ram1 \ x \ set b \ x = ram2" by auto with b p a_def b_def c_def show ?thesis by auto qed lemma "before vs ram1 ram2 \ pre_between vs ram1 ram2 \ hd vs \ ram1 \ (a,b) = splitAt (hd vs) (between vs ram2 ram1) \ vs = [hd vs] @ b @ [ram1] @ (between vs ram1 ram2) @ [ram2] @ a" proof - assume b: "before vs ram1 ram2" and p: "pre_between vs ram1 ram2" and vs: "hd vs \ ram1" and ab: "(a,b) = splitAt (hd vs) (between vs ram2 ram1)" from p have dist_b: "distinct (between vs ram2 ram1)" by (auto intro: between_distinct simp: pre_between_def) with ab have "distinct a \ distinct b" by (auto intro: splitAt_distinct_a splitAt_distinct_b) define r where "r = snd (splitAt ram1 vs)" define btw where "btw = between vs ram2 ram1" from p r_def have vs2: "vs = fst (splitAt ram1 vs) @ [ram1] @ r" by (auto dest: splitAt_ram simp: pre_between_def) then have "fst (splitAt ram1 vs) = [] \ hd vs = ram1" by auto with vs have neq: "fst (splitAt ram1 vs) \ []" by auto with vs2 have vs_fst: "hd vs = hd (fst (splitAt ram1 vs))" by (induct ("fst (splitAt ram1 vs)")) auto with neq have "hd vs \ set (fst (splitAt ram1 vs))" by auto with b p have "hd vs \ set (between vs ram2 ram1)" by auto with btw_def have help1: "btw = fst (splitAt (hd vs) btw) @ [hd vs] @ snd (splitAt (hd vs) btw)" by (auto dest: splitAt_ram) from p b btw_def have "btw = snd (splitAt ram2 vs) @ fst (splitAt ram1 vs)" by auto with neq have "btw = snd (splitAt ram2 vs) @ hd (fst (splitAt ram1 vs)) # tl (fst (splitAt ram1 vs))" by auto with vs_fst have "btw = snd (splitAt ram2 vs) @ [hd vs] @ tl (fst (splitAt ram1 vs))" by auto with help1 have eq: "snd (splitAt ram2 vs) @ [hd vs] @ tl (fst (splitAt ram1 vs)) = fst (splitAt (hd vs) btw) @ [hd vs] @ snd (splitAt (hd vs) btw)" by auto from dist_b btw_def help1 have "distinct (fst (splitAt (hd vs) btw) @ [hd vs] @ snd (splitAt (hd vs) btw))" by auto with eq have eq2: "snd (splitAt ram2 vs) = fst (splitAt (hd vs) btw) \ tl (fst (splitAt ram1 vs)) = snd (splitAt (hd vs) btw)" apply (rule_tac dist_at) by auto with btw_def ab have a: "a = snd (splitAt ram2 vs)" by (auto dest: pairD) from eq2 vs_fst have "hd (fst (splitAt ram1 vs)) # tl (fst (splitAt ram1 vs)) = hd vs # snd (splitAt (hd vs) btw)" by auto with ab btw_def neq have hdb: "hd vs # b = fst (splitAt ram1 vs)" by (auto dest: pairD) from b p have "vs = fst (splitAt ram1 vs) @ [ram1] @ fst (splitAt ram2 (snd (splitAt ram1 vs))) @ [ram2] @ snd (splitAt ram2 vs)" apply simp apply (rule_tac before_vs) by (auto simp: pre_between_def) with hdb have "vs = (hd vs # b) @ [ram1] @ fst (splitAt ram2 (snd (splitAt ram1 vs))) @ [ram2] @ snd (splitAt ram2 vs)" by auto with a b p show ?thesis by (simp) qed lemma between_congs: "pre_between vs ram1 ram2 \ vs \ vs' \ between vs ram1 ram2 = between vs' ram1 ram2" proof - have "\ us. pre_between us ram1 ram2 \ before us ram1 ram2 \ between us ram1 ram2 = between (rotate1 us) ram1 ram2" proof - fix us assume vors: "pre_between us ram1 ram2" "before us ram1 ram2" then have pb2: "pre_between (rotate1 us) ram1 ram2" by auto with vors show "between us ram1 ram2 = between (rotate1 us) ram1 ram2" proof (cases "us") case Nil then show ?thesis by auto next case (Cons u' us') with vors pb2 show ?thesis apply (auto simp: before_def) apply (case_tac "a") apply auto by (simp_all add: between_def split_def pre_between_def) qed qed moreover have "\ us. pre_between us ram1 ram2 \ before us ram2 ram1 \ between us ram1 ram2 = between (rotate1 us) ram1 ram2" proof - fix us assume vors: " pre_between us ram1 ram2" "before us ram2 ram1" then have pb2: "pre_between (rotate1 us) ram1 ram2" by auto with vors show "between us ram1 ram2 = between (rotate1 us) ram1 ram2" proof (cases "us") case Nil then show ?thesis by auto next case (Cons u' us') with vors pb2 show ?thesis apply (auto simp: before_def) apply (case_tac "a") apply auto by (simp_all add: between_def split_def pre_between_def) qed qed ultimately have "help": "\ us. pre_between us ram1 ram2 \ between us ram1 ram2 = between (rotate1 us) ram1 ram2" apply (subgoal_tac "before us ram1 ram2 \ before us ram2 ram1") by auto assume "vs \ vs'" and pre_b: "pre_between vs ram1 ram2" then obtain n where vs': "vs' = rotate n vs" by (auto simp: congs_def) have "between vs ram1 ram2 = between (rotate n vs) ram1 ram2" proof (induct n) case 0 then show ?case by auto next case (Suc m) then show ?case apply simp apply (subgoal_tac " between (rotate1 (rotate m vs)) ram1 ram2 = between (rotate m vs) ram1 ram2") by (auto intro: "help" [symmetric] pre_b) qed with vs' show ?thesis by auto qed lemma between_inter_empty: "pre_between vs ram1 ram2 \ set (between vs ram1 ram2) \ set (between vs ram2 ram1) = {}" apply (case_tac "before vs ram1 ram2") apply (simp add: pre_between_def) apply (elim conjE) apply (frule (1) before_vs) apply (subgoal_tac "distinct (fst (splitAt ram1 vs) @ ram1 # fst (splitAt ram2 (snd (splitAt ram1 vs))) @ ram2 # snd (splitAt ram2 vs))") apply (thin_tac "vs = fst (splitAt ram1 vs) @ ram1 # fst (splitAt ram2 (snd (splitAt ram1 vs))) @ ram2 # snd (splitAt ram2 vs)") apply (frule (1) before_dist_fst_snd) apply(simp) apply blast apply (simp only:) apply (simp add: before_xor) apply (subgoal_tac "pre_between vs ram2 ram1") apply (simp add: pre_between_def) apply (elim conjE) apply (frule (1) before_vs) apply (subgoal_tac "distinct (fst (splitAt ram2 vs) @ ram2 # fst (splitAt ram1 (snd (splitAt ram2 vs))) @ ram1 # snd (splitAt ram1 vs))") apply (thin_tac "vs = fst (splitAt ram2 vs) @ ram2 # fst (splitAt ram1 (snd (splitAt ram2 vs))) @ ram1 # snd (splitAt ram1 vs)") apply simp apply blast apply (simp only:) by (rule pre_between_symI) (*********************** between - is_nextElem *************************) subsubsection \\between is_nextElem\\ lemma is_nextElem_or1: "pre_between vs ram1 ram2 \ is_nextElem vs x y \ before vs ram1 ram2 \ is_sublist [x,y] (ram1 # between vs ram1 ram2 @ [ram2]) \ is_sublist [x,y] (ram2 # between vs ram2 ram1 @ [ram1])" proof - assume p: "pre_between vs ram1 ram2" and is_nextElem: "is_nextElem vs x y" and b: "before vs ram1 ram2" from p have r1: "ram1 \ set vs" by (auto simp: pre_between_def) define bs where "bs = [ram1] @ (between vs ram1 ram2) @ [ram2]" have rule1: "x \ set (ram1 # (between vs ram1 ram2)) \ is_sublist [x,y] bs" proof - assume xin:"x \ set (ram1 # (between vs ram1 ram2))" with bs_def have xin2: "x \ set bs" by auto define s where "s = snd (splitAt ram1 vs)" from r1 s_def have sub1:"is_sublist (ram1 # s) vs" by (auto intro: splitAt_is_sublist2R) from b p s_def have "ram2 \ set s" by (auto intro!: before_dist_r2 simp: pre_between_def) then have "is_prefix (fst (splitAt ram2 s) @ [ram2]) s" by (auto intro!: splitAt_is_prefix) then have "is_prefix ([ram1] @ ((fst (splitAt ram2 s)) @ [ram2])) ([ram1] @ s)" by (rule_tac is_prefix_add) auto with sub1 have "is_sublist (ram1 # (fst (splitAt ram2 s)) @ [ram2]) vs" apply (rule_tac is_sublist_trans) apply (rule is_prefix_sublist) by simp_all with p b s_def bs_def have subl: "is_sublist bs vs" by (auto) with p have db: "distinct bs" by (auto simp: pre_between_def) with xin bs_def have xnlb:"x \ last bs" by auto with p is_nextElem subl xin2 show "is_sublist [x,y] bs" apply (rule_tac is_sublist_is_nextElem) by (auto simp: pre_between_def) qed define bs2 where "bs2 = [ram2] @ (between vs ram2 ram1) @ [ram1]" have rule2: "x \ set (ram2 # (between vs ram2 ram1)) \ is_sublist [x,y] bs2" proof - assume xin:"x \ set (ram2 # (between vs ram2 ram1))" with bs2_def have xin2: "x \ set bs2" by auto define cs1 where "cs1 = ram2 # snd (splitAt ram2 vs)" then have cs1ne: "cs1 \ []" by auto define cs2 where "cs2 = fst (splitAt ram1 vs)" define cs2ram1 where "cs2ram1 = cs2 @ [ram1]" from cs1_def cs2_def bs2_def p b have bs2: "bs2 = cs1 @ cs2 @ [ram1]" by (auto simp: pre_between_def) have "x \ set cs1 \ x \ last cs1 \ is_sublist [x,y] cs1" proof- assume xin2: "x \ set cs1" and xnlcs1: "x \ last cs1" from cs1_def p have "is_sublist cs1 vs" by (simp add: pre_between_def) with p is_nextElem xnlcs1 xin2 show ?thesis apply (rule_tac is_sublist_is_nextElem) by (auto simp: pre_between_def) qed with bs2 have incs1nl: "x \ set cs1 \ x \ last cs1 \ is_sublist [x,y] bs2" apply (auto simp: is_sublist_def) apply (intro exI) apply (subgoal_tac "as @ x # y # bs @ cs2 @ [ram1] = as @ x # y # (bs @ cs2 @ [ram1])") apply assumption by auto have "x = last cs1 \ y = hd (cs2 @ [ram1])" proof - assume xl: "x = last cs1" from p have "distinct vs" by auto with p have "vs = fst (splitAt ram2 vs) @ ram2 # snd (splitAt ram2 vs)" by (auto intro: splitAt_ram) with cs1_def have "last vs = last (fst (splitAt ram2 vs) @ cs1)" by auto with cs1ne have "last vs = last cs1" by auto with xl have "x = last vs" by auto with p is_nextElem have yhd: "y = hd vs" by auto from p have "vs = fst (splitAt ram1 vs) @ ram1 # snd (splitAt ram1 vs)" by (auto intro: splitAt_ram) with yhd cs2ram1_def cs2_def have yhd2: "y = hd (cs2ram1 @ snd (splitAt ram1 vs))" by auto from cs2ram1_def have "cs2ram1 \ []" by auto then have "hd (cs2ram1 @ snd(splitAt ram1 vs)) = hd (cs2ram1)" by auto with yhd2 cs2ram1_def show ?thesis by auto qed with bs2 cs1ne have "x = last cs1 \ is_sublist [x,y] bs2" apply (auto simp: is_sublist_def) apply (intro exI) apply (subgoal_tac "cs1 @ cs2 @ [ram1] = butlast cs1 @ last cs1 # hd (cs2 @ [ram1]) # tl (cs2 @ [ram1])") apply assumption by auto with incs1nl have incs1: "x \ set cs1 \ is_sublist [x,y] bs2" by auto have "x \ set cs2 \ is_sublist [x,y] (cs2 @ [ram1])" proof- assume xin2': "x \ set cs2" then have xin2: "x \ set (cs2 @ [ram1])" by auto define fr2 where "fr2 = snd (splitAt ram1 vs)" from p have "vs = fst (splitAt ram1 vs) @ ram1 # snd (splitAt ram1 vs)" by (auto intro: splitAt_ram) with fr2_def cs2_def have "vs = cs2 @ [ram1] @ fr2" by simp with p xin2' have "x \ ram1" by (auto simp: pre_between_def) then have xnlcs2: "x \ last (cs2 @ [ram1])" by auto from cs2_def p have "is_sublist (cs2 @ [ram1]) vs" by (simp add: pre_between_def) with p is_nextElem xnlcs2 xin2 show ?thesis apply (rule_tac is_sublist_is_nextElem) by (auto simp: pre_between_def) qed with bs2 have "x \ set cs2 \ is_sublist [x,y] bs2" apply (auto simp: is_sublist_def) apply (intro exI) apply (subgoal_tac "cs1 @ as @ x # y # bs = (cs1 @ as) @ x # y # bs") apply assumption by auto with p b cs1_def cs2_def incs1 xin show ?thesis by auto qed from is_nextElem have "x \ set vs" by auto with b p have "x = ram1 \ x \ set (between vs ram1 ram2) \ x = ram2 \ x \ set (between vs ram2 ram1)" by (rule_tac between_in) auto then have "x \ set (ram1 # (between vs ram1 ram2)) \ x \ set (ram2 # (between vs ram2 ram1))" by auto with rule1 rule2 bs_def bs2_def show ?thesis by auto qed lemma is_nextElem_or: "pre_between vs ram1 ram2 \ is_nextElem vs x y \ is_sublist [x,y] (ram1 # between vs ram1 ram2 @ [ram2]) \ is_sublist [x,y] (ram2 # between vs ram2 ram1 @ [ram1])" proof (cases "before vs ram1 ram2") case True assume "pre_between vs ram1 ram2" "is_nextElem vs x y" with True show ?thesis by (rule_tac is_nextElem_or1) next assume p: "pre_between vs ram1 ram2" and is_nextElem: "is_nextElem vs x y" from p have p': "pre_between vs ram2 ram1" by (auto intro: pre_between_symI) case False with p have "before vs ram2 ram1" by auto with p' is_nextElem show ?thesis apply (drule_tac is_nextElem_or1) apply assumption+ by auto qed lemma(*<*) between_eq2: (*>*) "pre_between vs ram1 ram2 \ before vs ram2 ram1 \ \as bs cs. between vs ram1 ram2 = cs @ as \ vs = as @[ram2] @ bs @ [ram1] @ cs" apply (subgoal_tac "pre_between vs ram2 ram1") apply auto apply (intro exI conjI) apply simp apply (simp add: pre_between_def) apply auto apply (frule_tac before_vs) apply auto by (auto simp: pre_between_def) lemma is_sublist_same_len[simp]: "length xs = length ys \ is_sublist xs ys = (xs = ys)" apply(cases xs) apply simp apply(rename_tac a as) apply(cases ys) apply simp apply(rename_tac b bs) apply(case_tac "a = b") apply(subst is_sublist_rec) apply simp apply simp done lemma is_nextElem_between_empty[simp]: "distinct vs \ is_nextElem vs a b \ between vs a b = []" apply (simp add: is_nextElem_def between_def split_def) apply (cases "vs") apply simp+ apply (simp split: if_split_asm) apply (case_tac "b = aa") apply (simp add: is_sublist_def) apply (erule disjE) apply (elim exE) apply (case_tac "as") apply simp apply simp apply simp apply (case_tac "list" rule: rev_exhaust) apply simp apply simp apply simp apply (subgoal_tac "aa # list = vs") apply (thin_tac "vs = aa # list") apply simp apply (subgoal_tac "distinct vs") apply (simp add: is_sublist_def) apply (elim exE) by auto lemma is_nextElem_between_empty': "between vs a b = [] \ distinct vs \ a \ set vs \ b \ set vs \ a \ b \ is_nextElem vs a b" apply (simp add: is_nextElem_def between_def split_def split: if_split_asm) apply (case_tac vs) apply simp apply simp apply (rule conjI) apply (rule impI) apply simp apply (case_tac "aa = b") apply simp apply (rule impI) apply (case_tac "list" rule: rev_exhaust) apply simp apply simp apply (case_tac "a = y") apply simp apply simp apply (elim conjE) apply (drule split_list_first) apply (elim exE) apply simp apply (rule impI) apply (subgoal_tac "b \ aa") apply simp apply (case_tac "a = aa") apply simp apply (case_tac "list") apply simp apply simp apply (case_tac "aaa = b") apply (simp add: is_sublist_def) apply force apply simp apply simp apply (drule split_list_first) apply (elim exE) apply simp apply (case_tac "zs") apply simp apply simp apply (case_tac "aaa = b") apply simp apply (simp add: is_sublist_def) apply force apply simp apply force apply (case_tac vs) apply simp apply simp apply (rule conjI) apply (rule impI) apply simp apply (rule impI) apply (case_tac "b = aa") apply simp apply (case_tac "list" rule: rev_exhaust) apply simp apply simp apply (case_tac "a = y") apply simp apply simp apply (drule split_list_first) apply (elim exE) apply simp apply simp apply (case_tac "a = aa") by auto lemma between_nextElem: "pre_between vs u v \ between vs u (nextElem vs (hd vs) v) = between vs u v @ [v]" apply(subgoal_tac "pre_between vs v u") prefer 2 apply (clarsimp simp add:pre_between_def) apply(case_tac "before vs u v") apply(drule (1) between_eq2) apply(clarsimp simp:pre_between_def hd_append split:list.split) apply(simp add:between_def split_def) apply(fastforce simp:neq_Nil_conv) apply (simp only:before_xor) apply(drule (1) between_eq2) apply(clarsimp simp:pre_between_def hd_append split:list.split) apply (simp add: append_eq_Cons_conv) apply(fastforce simp:between_def split_def) done (******************** section split_face ********************************) lemma nextVertices_in_face[simp]: "v \ \ f \ f\<^bsup>n\<^esup> \ v \ \ f" proof - assume v: "v \ \ f" then have f: "vertices f \ []" by auto show ?thesis apply (simp add: nextVertices_def) apply (induct n) by (auto simp: f v) qed subsubsection \\is_nextElem edges\ equivalence\ lemma is_nextElem_edges1: "distinct (vertices f) \ (a,b) \ edges (f::face) \ is_nextElem (vertices f) a b" apply (simp add: edges_face_def nextVertex_def) apply (rule is_nextElem1) by auto lemma is_nextElem_edges2: "distinct (vertices f) \ is_nextElem (vertices f) a b \ (a,b) \ edges (f::face)" apply (auto simp add: edges_face_def nextVertex_def) apply (rule sym) apply (rule is_nextElem2) by (auto intro: is_nextElem_a) lemma is_nextElem_edges_eq[simp]: "distinct (vertices (f::face)) \ (a,b) \ edges f = is_nextElem (vertices f) a b" by (auto intro: is_nextElem_edges1 is_nextElem_edges2) (*********************** nextVertex *****************************) subsubsection \@{const nextVertex}\ lemma nextElem_suc2: "distinct (vertices f) \ last (vertices f) = v \ v \ set (vertices f) \ f \ v = hd (vertices f)" proof - assume dist: "distinct (vertices f)" and last: "last (vertices f) = v" and v: "v \ set (vertices f)" define ls where "ls = vertices f" have ind: "\ c. distinct ls \ last ls = v \ v \ set ls \ nextElem ls c v = c" proof (induct ls) case Nil then show ?case by auto next case (Cons m ms) then show ?case proof (cases "m = v") case True with Cons have "ms = []" apply (cases ms rule: rev_exhaust) by auto then show ?thesis by auto next case False with Cons have a1: "v \ set ms" by auto then have ms: "ms \ []" by auto with False Cons ms have "nextElem ms c v = c" apply (rule_tac Cons) by simp_all with False ms show ?thesis by simp qed qed from dist ls_def last v have "nextElem ls (hd ls) v = hd ls" apply (rule_tac ind) by auto with ls_def show ?thesis by (simp add: nextVertex_def) qed (*********************** split_face ****************************) subsection \@{const split_face}\ definition pre_split_face :: "face \ nat \ nat \ nat list \ bool" where "pre_split_face oldF ram1 ram2 newVertexList \ distinct (vertices oldF) \ distinct (newVertexList) \ \ oldF \ set newVertexList = {} \ ram1 \ \ oldF \ ram2 \ \ oldF \ ram1 \ ram2" declare pre_split_face_def [simp] lemma pre_split_face_p_between[intro]: "pre_split_face oldF ram1 ram2 newVertexList \ pre_between (vertices oldF) ram1 ram2" by (simp add: pre_between_def) lemma pre_split_face_symI: "pre_split_face oldF ram1 ram2 newVertexList \ pre_split_face oldF ram2 ram1 newVertexList" by auto lemma pre_split_face_rev[intro]: "pre_split_face oldF ram1 ram2 newVertexList \ pre_split_face oldF ram1 ram2 (rev newVertexList)" by auto lemma split_face_distinct1: "(f12, f21) = split_face oldF ram1 ram2 newVertexList \ pre_split_face oldF ram1 ram2 newVertexList \ distinct (vertices f12)" proof - assume split: "(f12, f21) = split_face oldF ram1 ram2 newVertexList" and p: "pre_split_face oldF ram1 ram2 newVertexList" define old_vs where "old_vs = vertices oldF" with p have d_old_vs: "distinct old_vs" by auto from p have p2: "pre_between (vertices oldF) ram1 ram2" by auto have rule1: "before old_vs ram1 ram2 \ distinct (vertices f12)" proof - assume b: "before old_vs ram1 ram2" with split p have "f12 = Face (rev newVertexList @ ram1 # between (vertices oldF) ram1 ram2 @ [ram2]) Nonfinal" by (simp add: split_face_def) then have h1:"vertices f12 = rev newVertexList @ ram1 # between (vertices oldF) ram1 ram2 @ [ram2]" by auto from p have d1: "distinct(ram1 # between (vertices oldF) ram1 ram2 @ [ram2])" by (auto del: notI) from b p p2 old_vs_def have d2: "set (ram1 # between (vertices oldF) ram1 ram2 @ [ram2]) \ set newVertexList = {}" by (auto dest!: splitAt_in_fst splitAt_in_snd) with h1 d1 p show ?thesis by auto qed have rule2: "before old_vs ram2 ram1 \ distinct (vertices f12)" proof - assume b: "before old_vs ram2 ram1" from p have p3: "pre_split_face oldF ram1 ram2 newVertexList" by (auto intro: pre_split_face_symI) then have p4: "pre_between (vertices oldF) ram2 ram1" by auto with split p have "f12 = Face (rev newVertexList @ ram1 # between (vertices oldF) ram1 ram2 @ [ram2]) Nonfinal" by (simp add: split_face_def) then have h1:"vertices f12 = rev newVertexList @ ram1 # between (vertices oldF) ram1 ram2 @ [ram2]" by auto from p3 have d1: "distinct(ram1 # between (vertices oldF) ram1 ram2 @ [ram2])" by (auto del: notI) from b p3 p4 old_vs_def have d2: "set (ram1 # between (vertices oldF) ram1 ram2 @ [ram2]) \ set newVertexList = {}" by auto with h1 d1 p show ?thesis by auto qed from p2 rule1 rule2 old_vs_def show ?thesis by auto qed lemma split_face_distinct1'[intro]: "pre_split_face oldF ram1 ram2 newVertexList \ distinct (vertices (fst(split_face oldF ram1 ram2 newVertexList)))" apply (rule_tac split_face_distinct1) by (auto simp del: pre_split_face_def simp: split_face_def) lemma split_face_distinct2: "(f12, f21) = split_face oldF ram1 ram2 newVertexList \ pre_split_face oldF ram1 ram2 newVertexList \ distinct (vertices f21)" proof - assume split: "(f12, f21) = split_face oldF ram1 ram2 newVertexList" and p: "pre_split_face oldF ram1 ram2 newVertexList" define old_vs where "old_vs = vertices oldF" with p have d_old_vs: "distinct old_vs" by auto with p have p1: "pre_split_face oldF ram1 ram2 newVertexList" by auto from p have p2: "pre_between (vertices oldF) ram1 ram2" by auto have rule1: "before old_vs ram1 ram2 \ distinct (vertices f21)" proof - assume b: "before old_vs ram1 ram2" with split p have "f21 = Face (ram2 # between (vertices oldF) ram2 ram1 @ [ram1] @ newVertexList) Nonfinal" by (simp add: split_face_def) then have h1:"vertices f21 = ram2 # between (vertices oldF) ram2 ram1 @ [ram1] @ newVertexList" by auto from p have d1: "distinct(ram1 # between (vertices oldF) ram2 ram1 @ [ram2])" by (auto del: notI) from b p1 p2 old_vs_def have d2: "set (ram2 # between (vertices oldF) ram2 ram1 @ [ram1]) \ set newVertexList = {}" by auto with h1 d1 p1 show ?thesis by auto qed have rule2: "before old_vs ram2 ram1 \ distinct (vertices f21)" proof - assume b: "before old_vs ram2 ram1" from p have p3: "pre_split_face oldF ram1 ram2 newVertexList" by (auto intro: pre_split_face_symI) then have p4: "pre_between (vertices oldF) ram2 ram1" by auto with split p have "f21 = Face (ram2 # between (vertices oldF) ram2 ram1 @ [ram1] @ newVertexList) Nonfinal" by (simp add: split_face_def) then have h1:"vertices f21 = ram2 # between (vertices oldF) ram2 ram1 @ [ram1] @ newVertexList" by auto from p3 have d1: "distinct(ram2 # between (vertices oldF) ram2 ram1 @ [ram1])" by (auto del: notI) from b p3 p4 old_vs_def have d2: "set (ram2 # between (vertices oldF) ram2 ram1 @ [ram1]) \ set newVertexList = {}" by auto with h1 d1 p1 show ?thesis by auto qed from p2 rule1 rule2 old_vs_def show ?thesis by auto qed lemma split_face_distinct2'[intro]: "pre_split_face oldF ram1 ram2 newVertexList \ distinct (vertices (snd(split_face oldF ram1 ram2 newVertexList)))" apply (rule_tac split_face_distinct2) by (auto simp del: pre_split_face_def simp: split_face_def) declare pre_split_face_def [simp del] lemma split_face_edges_or: "(f12, f21) = split_face oldF ram1 ram2 newVertexList \ pre_split_face oldF ram1 ram2 newVertexList \ (a, b) \ edges oldF \ (a,b) \ edges f12 \ (a,b) \ edges f21" proof - assume nf: "(f12, f21) = split_face oldF ram1 ram2 newVertexList" and p: "pre_split_face oldF ram1 ram2 newVertexList" and old:"(a, b) \ edges oldF" from p have d1:"distinct (vertices oldF)" by auto from nf p have d2: "distinct (vertices f12)" by (auto dest: pairD) from nf p have d3: "distinct (vertices f21)" by (auto dest: pairD) from p have p': "pre_between (vertices oldF) ram1 ram2" by auto from old d1 have is_nextElem: "is_nextElem (vertices oldF) a b" by simp with p have "is_sublist [a,b] (ram1 # (between (vertices oldF) ram1 ram2) @ [ram2]) \ is_sublist [a,b] (ram2 # (between (vertices oldF) ram2 ram1) @ [ram1])" apply (rule_tac is_nextElem_or) by auto then have "is_nextElem (rev newVertexList @ ram1 # between (vertices oldF) ram1 ram2 @ [ram2]) a b \ is_nextElem (ram2 # between (vertices oldF) ram2 ram1 @ ram1 # newVertexList) a b" proof (cases "is_sublist [a,b] (ram1 # (between (vertices oldF) ram1 ram2) @ [ram2])") case True then show ?thesis by (auto dest: is_sublist_add intro!: is_nextElem_sublistI) next case False assume "is_sublist [a,b] (ram1 # (between (vertices oldF) ram1 ram2) @ [ram2]) \ is_sublist [a,b] (ram2 # (between (vertices oldF) ram2 ram1) @ [ram1])" with False have "is_sublist [a,b] (ram2 # (between (vertices oldF) ram2 ram1) @ [ram1])" by auto then show ?thesis apply (rule_tac disjI2) apply (rule_tac is_nextElem_sublistI) apply (subgoal_tac "is_sublist [a, b] ([] @ (ram2 # between (vertices oldF) ram2 ram1 @ [ram1]) @ newVertexList)") apply force by (frule is_sublist_add) qed with nf d1 d2 d3 show ?thesis by (simp add: split_face_def) qed subsection \\verticesFrom\\ definition verticesFrom :: "face \ vertex \ vertex list" where "verticesFrom f \ rotate_to (vertices f)" lemmas verticesFrom_Def = verticesFrom_def rotate_to_def lemma len_vFrom[simp]: "v \ \ f \ |verticesFrom f v| = |vertices f|" apply(drule split_list_first) apply(clarsimp simp: verticesFrom_Def) done lemma verticesFrom_empty[simp]: "v \ \ f \ (verticesFrom f v = []) = (vertices f = [])" by(clarsimp simp: verticesFrom_Def) lemma verticesFrom_congs: "v \ \ f \ (vertices f) \ (verticesFrom f v)" by(simp add:verticesFrom_def cong_rotate_to) lemma verticesFrom_eq_if_vertices_cong: "\distinct(vertices f); distinct(vertices f'); vertices f \ vertices f'; x \ \ f \ \ verticesFrom f x = verticesFrom f' x" by(clarsimp simp:congs_def verticesFrom_Def splitAt_rotate_pair_conv) lemma verticesFrom_in[intro]: "v \ \ f \ a \ \ f \ a \ set (verticesFrom f v)" by (auto dest: verticesFrom_congs congs_pres_nodes) lemma verticesFrom_in': "a \ set (verticesFrom f v) \ a \ v \ a \ \ f" apply (cases "v \ \ f") apply (auto dest: verticesFrom_congs congs_pres_nodes) by (simp add: verticesFrom_Def) lemma set_verticesFrom: "v \ \ f \ set (verticesFrom f v) = \ f" apply(auto) apply (auto simp add: verticesFrom_Def) done lemma verticesFrom_hd: "hd (verticesFrom f v) = v" by (simp add: verticesFrom_Def) lemma verticesFrom_distinct[simp]: "distinct (vertices f) \ v \ \ f \ distinct (verticesFrom f v)" apply (frule_tac verticesFrom_congs) by (auto simp: congs_distinct) lemma verticesFrom_nextElem_eq: "distinct (vertices f) \ v \ \ f \ u \ \ f \ nextElem (verticesFrom f v) (hd (verticesFrom f v)) u = nextElem (vertices f) (hd (vertices f)) u" apply (subgoal_tac "(verticesFrom f v) \ (vertices f)") apply (rule nextElem_congs_eq) apply (auto simp: verticesFrom_congs congs_pres_nodes) apply (rule congs_sym) by (simp add: verticesFrom_congs) lemma nextElem_vFrom_suc1: "distinct (vertices f) \ v \ \ f \ i < length (vertices f) \ last (verticesFrom f v) \ u \ (verticesFrom f v)!i = u \ f \ u = (verticesFrom f v)!(Suc i)" proof - assume dist: "distinct (vertices f)" and ith: "(verticesFrom f v)!i = u" and small_i: "i < length (vertices f)" and notlast: "last (verticesFrom f v) \ u" and v: "v \ \ f" hence eq: "(vertices f) \ (verticesFrom f v)" by (auto simp: verticesFrom_congs) from ith eq small_i have "u \ set (verticesFrom f v)" by (auto simp: congs_length) with eq have u: "u \ \ f" by (auto simp: congs_pres_nodes) define ls where "ls = verticesFrom f v" with dist ith have "ls!i = u" by auto have ind: "\ c i. i < length ls \ distinct ls \ last ls \ u \ ls!i = u \ u \ set ls \ nextElem ls c u = ls ! Suc i" proof (induct ls) case Nil then show ?case by auto next case (Cons m ms) then show ?case proof (cases "m = u") case True with Cons have "u \ set ms" by auto with Cons True have i: "i = 0" apply (induct i) by auto with Cons show ?thesis apply (simp split: if_split_asm) apply (cases ms) by simp_all next case False with Cons have a1: "u \ set ms" by auto then have ms: "ms \ []" by auto from False Cons have i: "0 < i" by auto define i' where "i' = i - 1" with i have i': "i = Suc i'" by auto with False Cons i' ms have "nextElem ms c u = ms ! Suc i'" apply (rule_tac Cons) by simp_all with False Cons i' ms show ?thesis by simp qed qed from eq dist ith ls_def small_i notlast v have "nextElem ls (hd ls) u = ls ! Suc i" apply (rule_tac ind) by (auto simp: congs_length) with dist u v ls_def show ?thesis by (simp add: nextVertex_def verticesFrom_nextElem_eq) qed lemma verticesFrom_nth: "distinct (vertices f) \ d < length (vertices f) \ v \ \ f \ (verticesFrom f v)!d = f\<^bsup>d\<^esup> \ v" proof (induct d) case 0 then show ?case by (simp add: verticesFrom_Def nextVertices_def) next case (Suc n) then have dist2: "distinct (verticesFrom f v)" apply (frule_tac verticesFrom_congs) by (auto simp: congs_distinct) from Suc have in2: "v \ set (verticesFrom f v)" apply (frule_tac verticesFrom_congs) by (auto dest: congs_pres_nodes) then have vFrom: "(verticesFrom f v) = butlast (verticesFrom f v) @ [last (verticesFrom f v)]" apply (cases "(verticesFrom f v)" rule: rev_exhaust) by auto from Suc show ?case proof (cases "last (verticesFrom f v) = f\<^bsup>n\<^esup> \ v") case True with Suc have "verticesFrom f v ! n = f\<^bsup>n\<^esup> \ v" by (rule_tac Suc) auto with True have "last (verticesFrom f v) = verticesFrom f v ! n" by auto with Suc dist2 in2 have "Suc n = length (verticesFrom f v)" apply (rule_tac nth_last_Suc_n) by auto with Suc show ?thesis by auto next case False with Suc show ?thesis apply (simp add: nextVertices_def) apply (rule sym) apply (rule_tac nextElem_vFrom_suc1) by simp_all qed qed lemma verticesFrom_length: "distinct (vertices f) \ v \ set (vertices f) \ length (verticesFrom f v) = length (vertices f)" by (auto intro: congs_length verticesFrom_congs) lemma verticesFrom_between: "v' \ \ f \ pre_between (vertices f) u v \ between (vertices f) u v = between (verticesFrom f v') u v" by (auto intro!: between_congs verticesFrom_congs) lemma verticesFrom_is_nextElem: "v \ \ f \ is_nextElem (vertices f) a b = is_nextElem (verticesFrom f v) a b" apply (rule is_nextElem_congs_eq) by (rule verticesFrom_congs) lemma verticesFrom_is_nextElem_last: "v' \ \ f \ distinct (vertices f) \ is_nextElem (verticesFrom f v') (last (verticesFrom f v')) v \ v = v'" apply (subgoal_tac "distinct (verticesFrom f v')") apply (subgoal_tac "last (verticesFrom f v') \ set (verticesFrom f v')") apply (simp add: nextElem_is_nextElem) apply (simp add: verticesFrom_hd) apply (cases "(verticesFrom f v')" rule: rev_exhaust) apply (simp add: verticesFrom_Def) by auto lemma verticesFrom_is_nextElem_hd: "v' \ \ f \ distinct (vertices f) \ is_nextElem (verticesFrom f v') u v' \ u = last (verticesFrom f v')" apply (subgoal_tac "distinct (verticesFrom f v')") apply (thin_tac "distinct (vertices f)") apply (auto simp: is_nextElem_def) apply (drule is_sublist_y_hd) apply (simp add: verticesFrom_hd) by auto lemma verticesFrom_pres_nodes1: "v \ \ f \ \ f = set(verticesFrom f v)" proof - assume "v \ \ f" then have "fst (splitAt v (vertices f)) @ [v] @ snd (splitAt v (vertices f)) = vertices f" apply (drule_tac splitAt_ram) by simp moreover have "set (fst (splitAt v (vertices f)) @ [v] @ snd (splitAt v (vertices f))) = set (verticesFrom f v)" by (auto simp: verticesFrom_Def) ultimately show ?thesis by simp qed lemma verticesFrom_pres_nodes: "v \ \ f \ w \ \ f \ w \ set (verticesFrom f v)" by (auto dest: verticesFrom_pres_nodes1) lemma before_verticesFrom: "distinct (vertices f) \ v \ \ f \ w \ \ f \ v \ w \ before (verticesFrom f v) v w" proof - assume v: "v \ \ f" and w: "w \ \ f" and neq: "v \ w" have "hd (verticesFrom f v) = v" by (rule verticesFrom_hd) with v have vf:"verticesFrom f v = v # tl (verticesFrom f v)" apply (frule_tac verticesFrom_pres_nodes1) apply (cases "verticesFrom f v") by auto moreover with v w have "w \ set (verticesFrom f v)" by (auto simp: verticesFrom_pres_nodes) ultimately have "w \ set (v # tl (verticesFrom f v))" by auto with neq have "w \ set (tl (verticesFrom f v))" by auto with vf have "verticesFrom f v = [] @ v # fst (splitAt w (tl (verticesFrom f v))) @ w # snd (splitAt w (tl (verticesFrom f v)))" by (auto dest: splitAt_ram) then show ?thesis apply (unfold before_def) by (intro exI) qed lemma last_vFrom: "\ distinct(vertices f); x \ \ f \ \ last(verticesFrom f x) = f\<^bsup>-1\<^esup> \ x" apply(frule split_list) apply(clarsimp simp:prevVertex_def verticesFrom_Def split:list.split) done lemma rotate_before_vFrom: "\ distinct(vertices f); r \ \ f; r\u \ \ before (verticesFrom f r) u v \ before (verticesFrom f v) r u" using [[linarith_neq_limit = 1]] apply(frule split_list) apply(clarsimp simp:verticesFrom_Def) apply(rename_tac as bs) apply(clarsimp simp:before_def) apply(rename_tac xs ys zs) apply(subst (asm) Cons_eq_append_conv) apply clarsimp apply(rename_tac bs') apply(subst (asm) append_eq_append_conv2) apply clarsimp apply(rename_tac as') apply(erule disjE) defer apply clarsimp apply(rule_tac x = "v#zs" in exI) apply(rule_tac x = "bs@as'" in exI) apply simp apply clarsimp apply(subst (asm) append_eq_Cons_conv) apply(erule disjE) apply clarsimp apply(rule_tac x = "v#zs" in exI) apply simp apply blast apply clarsimp apply(rename_tac ys') apply(subst (asm) append_eq_append_conv2) apply clarsimp apply(rename_tac vs') apply(erule disjE) apply clarsimp apply(subst (asm) append_eq_Cons_conv) apply(erule disjE) apply clarsimp apply(rule_tac x = "v#zs" in exI) apply simp apply blast apply clarsimp apply(rule_tac x = "v#ys'@as" in exI) apply simp apply blast apply clarsimp apply(rule_tac x = "v#zs" in exI) apply simp apply blast done lemma before_between: "\ before(verticesFrom f x) y z; distinct(vertices f); x \ \ f; x \ y \ \ y \ set(between (vertices f) x z)" apply(induct f) apply(clarsimp simp:verticesFrom_Def before_def) apply(frule split_list) apply(clarsimp simp:Cons_eq_append_conv) apply(subst (asm) append_eq_append_conv2) apply clarsimp apply(erule disjE) apply(clarsimp simp:append_eq_Cons_conv) prefer 2 apply(clarsimp simp add:between_def split_def) apply(erule disjE) apply (clarsimp simp:between_def split_def) apply clarsimp apply(subst (asm) append_eq_append_conv2) apply clarsimp apply(erule disjE) prefer 2 apply(clarsimp simp add:between_def split_def) apply(clarsimp simp:append_eq_Cons_conv) apply(fastforce simp:between_def split_def) done lemma before_between2: "\ before (verticesFrom f u) v w; distinct(vertices f); u \ \ f \ \ u = v \ u \ set (between (vertices f) w v)" apply(subgoal_tac "pre_between (vertices f) v w") apply(subst verticesFrom_between) apply assumption apply(erule pre_between_symI) apply(frule pre_between_r1) apply(drule (1) verticesFrom_distinct) using verticesFrom_hd[of f u] apply(clarsimp simp add:before_def between_def split_def hd_append split:if_split_asm) apply(frule (1) verticesFrom_distinct) apply(clarsimp simp:pre_between_def before_def simp del:verticesFrom_distinct) apply(rule conjI) apply(cases "v = u") apply simp apply(rule verticesFrom_in'[of v f u])apply simp apply assumption apply(cases "w = u") apply simp apply(rule verticesFrom_in'[of w f u])apply simp apply assumption done (************************** splitFace ********************************) subsection \@{const splitFace}\ definition pre_splitFace :: "graph \ vertex \ vertex \ face \ vertex list \ bool" where "pre_splitFace g ram1 ram2 oldF nvs \ oldF \ \ g \ \ final oldF \ distinct (vertices oldF) \ distinct nvs \ \ g \ set nvs = {} \ \ oldF \ set nvs = {} \ ram1 \ \ oldF \ ram2 \ \ oldF \ ram1 \ ram2 \ (((ram1,ram2) \ edges oldF \ (ram2,ram1) \ edges oldF \ (ram1, ram2) \ edges g \ (ram2, ram1) \ edges g) \ nvs \ [])" declare pre_splitFace_def [simp] lemma pre_splitFace_pre_split_face[simp]: "pre_splitFace g ram1 ram2 oldF nvs \ pre_split_face oldF ram1 ram2 nvs" by (simp add: pre_splitFace_def pre_split_face_def) lemma pre_splitFace_oldF[simp]: "pre_splitFace g ram1 ram2 oldF nvs \ oldF \ \ g" apply (unfold pre_splitFace_def) by force declare pre_splitFace_def [simp del] lemma splitFace_split_face: "oldF \ \ g \ (f\<^sub>1, f\<^sub>2, newGraph) = splitFace g ram\<^sub>1 ram\<^sub>2 oldF newVs \ (f\<^sub>1, f\<^sub>2) = split_face oldF ram\<^sub>1 ram\<^sub>2 newVs" by (simp add: splitFace_def split_def) (* split_face *) lemma split_face_empty_ram2_ram1_in_f12: "pre_split_face oldF ram1 ram2 [] \ (f12, f21) = split_face oldF ram1 ram2 [] \ (ram2, ram1) \ edges f12" proof - assume split: "(f12, f21) = split_face oldF ram1 ram2 []" "pre_split_face oldF ram1 ram2 []" then have "ram2 \ \ f12" by (simp add: split_face_def) moreover with split have "f12 \ ram2 = hd (vertices f12)" apply (rule_tac nextElem_suc2) apply (simp add: pre_split_face_def split_face_distinct1) by (simp add: split_face_def) with split have "ram1 = f12 \ ram2" by (simp add: split_face_def) ultimately show ?thesis by (simp add: edges_face_def) qed lemma split_face_empty_ram2_ram1_in_f12': "pre_split_face oldF ram1 ram2 [] \ (ram2, ram1) \ edges (fst (split_face oldF ram1 ram2 []))" proof - assume split: "pre_split_face oldF ram1 ram2 []" define f12 where "f12 = fst (split_face oldF ram1 ram2 [])" define f21 where "f21 = snd (split_face oldF ram1 ram2 [])" then have "(f12, f21) = split_face oldF ram1 ram2 []" by (simp add: f12_def f21_def) with split have "(ram2, ram1) \ edges f12" by (rule split_face_empty_ram2_ram1_in_f12) with f12_def show ?thesis by simp qed lemma splitFace_empty_ram2_ram1_in_f12: "pre_splitFace g ram1 ram2 oldF [] \ (f12, f21, newGraph) = splitFace g ram1 ram2 oldF [] \ (ram2, ram1) \ edges f12" proof - assume pre: "pre_splitFace g ram1 ram2 oldF []" then have oldF: "oldF \ \ g" by (unfold pre_splitFace_def) simp assume sp: "(f12, f21, newGraph) = splitFace g ram1 ram2 oldF []" with oldF have "(f12, f21) = split_face oldF ram1 ram2 []" by (rule splitFace_split_face) with pre sp show ?thesis apply (unfold splitFace_def pre_splitFace_def) apply (simp add: split_def) apply (rule split_face_empty_ram2_ram1_in_f12') apply (rule pre_splitFace_pre_split_face) apply (rule pre) done qed lemma splitFace_f12_new_vertices: "(f12, f21, newGraph) = splitFace g ram1 ram2 oldF newVs \ v \ set newVs \ v \ \ f12" apply (unfold splitFace_def) apply (simp add: split_def) apply (unfold split_face_def Let_def) by simp lemma splitFace_add_vertices_direct[simp]: "vertices (snd (snd (splitFace g ram1 ram2 oldF [countVertices g ..< countVertices g + n]))) = vertices g @ [countVertices g ..< countVertices g + n]" apply (auto simp: splitFace_def split_def) apply (cases g) apply auto apply (induct n) by auto lemma splitFace_delete_oldF: " (f12, f21, newGraph) = splitFace g ram1 ram2 oldF newVertexList \ oldF \ f12 \ oldF \ f21 \ distinct (faces g) \ oldF \ \ newGraph" by (simp add: splitFace_def split_def distinct_set_replace) lemma splitFace_faces_1: "(f12, f21, newGraph) = splitFace g ram1 ram2 oldF newVertexList \ oldF \ \ g \ set (faces newGraph) \ {oldF} = {f12, f21} \ set (faces g)" (is "?oldF \ ?C \ ?A = ?B") proof (intro equalityI subsetI) fix x assume "x \ ?A" and "?C" and "?oldF" then show "x \ ?B" apply (simp add: splitFace_def split_def) by (auto dest: replace1) next fix x assume "x \ ?B" and "?C" and "?oldF" then show "x \ ?A" apply (simp add: splitFace_def split_def) apply (cases "x = oldF") apply simp_all apply (cases "x = f12") apply simp_all apply (cases "x = f21") by (auto intro: replace3 replace4) qed lemma splitFace_distinct1[intro]:"pre_splitFace g ram1 ram2 oldF newVertexList \ distinct (vertices (fst (snd (splitFace g ram1 ram2 oldF newVertexList))))" apply (unfold splitFace_def split_def) by (auto simp add: split_def) lemma splitFace_distinct2[intro]:"pre_splitFace g ram1 ram2 oldF newVertexList \ distinct (vertices (fst (splitFace g ram1 ram2 oldF newVertexList)))" apply (unfold splitFace_def split_def) by (auto simp add: split_def) lemma splitFace_add_f21': "f' \ \ g' \ fst (snd (splitFace g' v a f' nvl)) \ \ (snd (snd (splitFace g' v a f' nvl)))" apply (simp add: splitFace_def split_def) apply (rule disjI2) apply (rule replace3) by simp_all lemma split_face_help[simp]: "Suc 0 < |vertices (fst (split_face f' v a nvl))|" by (simp add: split_face_def) lemma split_face_help'[simp]: "Suc 0 < |vertices (snd (split_face f' v a nvl))|" by (simp add: split_face_def) lemma splitFace_split: "f \ \ (snd (snd (splitFace g v a f' nvl))) \ f \ \ g \ f = fst (splitFace g v a f' nvl) \ f = (fst (snd (splitFace g v a f' nvl)))" apply (simp add: splitFace_def split_def) apply safe by (frule replace5) simp lemma pre_FaceDiv_between1: "pre_splitFace g' ram1 ram2 f [] \ \ between (vertices f) ram1 ram2 = []" proof - assume pre_f: "pre_splitFace g' ram1 ram2 f []" then have pre_bet: "pre_between (vertices f) ram1 ram2" apply (unfold pre_splitFace_def) by (simp add: pre_between_def) then have pre_bet': "pre_between (verticesFrom f ram1) ram1 ram2" by (simp add: pre_between_def set_verticesFrom) from pre_f have dist': "distinct (verticesFrom f ram1)" apply (unfold pre_splitFace_def) by simp from pre_f have ram2: "ram2 \ \ f" apply (unfold pre_splitFace_def) by simp from pre_f have "\ is_nextElem (vertices f) ram1 ram2" apply (unfold pre_splitFace_def) by auto with pre_f have "\ is_nextElem (verticesFrom f ram1) ram1 ram2" apply (unfold pre_splitFace_def) by (simp add: verticesFrom_is_nextElem [symmetric]) moreover from pre_f have "ram2 \ set (verticesFrom f ram1)" apply (unfold pre_splitFace_def) by auto moreover from pre_f have "ram2 \ ram1" apply (unfold pre_splitFace_def) by auto ultimately have ram2_not: "ram2 \ hd (snd (splitAt ram1 (vertices f)) @ fst (splitAt ram1 (vertices f)))" apply (simp add: is_nextElem_def verticesFrom_Def) apply (cases "snd (splitAt ram1 (vertices f)) @ fst (splitAt ram1 (vertices f))") apply simp apply simp apply (simp add: is_sublist_def) by auto from pre_f have before: "before (verticesFrom f ram1) ram1 ram2" apply (unfold pre_splitFace_def) apply safe apply (rule before_verticesFrom) by auto have "fst (splitAt ram2 (snd (splitAt ram1 (verticesFrom f ram1)))) = [] \ False" proof - assume "fst (splitAt ram2 (snd (splitAt ram1 (verticesFrom f ram1)))) = []" moreover from ram2 pre_f have "ram2 \ set (verticesFrom f ram1)"apply (unfold pre_splitFace_def) by auto with pre_f have "ram2 \ set (snd (splitAt ram1 (vertices f)) @ fst (splitAt ram1 (vertices f)))" apply (simp add: verticesFrom_Def) apply (unfold pre_splitFace_def) by auto moreover note dist' ultimately have "ram2 = hd (snd (splitAt ram1 (vertices f)) @ fst (splitAt ram1 (vertices f)))" apply (rule_tac ccontr) apply (cases "(snd (splitAt ram1 (vertices f)) @ fst (splitAt ram1 (vertices f)))") apply simp apply simp by (simp add: verticesFrom_Def del: distinct_append) with ram2_not show ?thesis by auto qed with before pre_bet' have "between (verticesFrom f ram1) ram1 ram2 \ []" by auto with pre_f pre_bet show ?thesis apply (unfold pre_splitFace_def) apply safe by (simp add: verticesFrom_between) qed lemma pre_FaceDiv_between2: "pre_splitFace g' ram1 ram2 f [] \ \ between (vertices f) ram2 ram1 = []" proof - assume pre_f: "pre_splitFace g' ram1 ram2 f []" then have "pre_splitFace g' ram2 ram1 f []" apply (unfold pre_splitFace_def) by auto then show ?thesis by (rule pre_FaceDiv_between1) qed (********************** Edges *******************) definition Edges :: "vertex list \ (vertex \ vertex) set" where "Edges vs \ {(a,b). is_sublist [a,b] vs}" lemma Edges_Nil[simp]: "Edges [] = {}" by(simp add:Edges_def is_sublist_def) lemma Edges_rev: "Edges (rev (zs::vertex list)) = {(b,a). (a,b) \ Edges zs}" by (auto simp add: Edges_def is_sublist_rev) lemma in_Edges_rev[simp]: "((a,b) : Edges (rev (zs::vertex list))) = ((b,a) \ Edges zs)" by (simp add: Edges_rev) lemma notinset_notinEdge1: "x \ set xs \ (x,y) \ Edges xs" by(unfold Edges_def)(blast intro:is_sublist_in) lemma notinset_notinEdge2: "y \ set xs \ (x,y) \ Edges xs" by(unfold Edges_def)(blast intro:is_sublist_in1) lemma in_Edges_in_set: "(x,y) : Edges vs \ x \ set vs \ y \ set vs" by(unfold Edges_def)(blast intro:is_sublist_in is_sublist_in1) lemma edges_conv_Edges: "distinct(vertices(f::face)) \ \ f = Edges (vertices f) \ (if vertices f = [] then {} else {(last(vertices f), hd(vertices f))})" by(induct f) (auto simp: Edges_def is_nextElem_def) lemma Edges_Cons: "Edges(x#xs) = (if xs = [] then {} else Edges xs \ {(x,hd xs)})" apply(auto simp:Edges_def) apply(rule ccontr) apply(simp) apply(erule thin_rl) apply(erule contrapos_np) apply(clarsimp simp:is_sublist_def Cons_eq_append_conv) apply(rename_tac as bs) apply(erule disjE) apply simp apply(clarsimp) apply(rename_tac cs) apply(rule_tac x = cs in exI) apply(rule_tac x = bs in exI) apply(rule HOL.refl) apply(clarsimp simp:neq_Nil_conv) apply(subst is_sublist_rec) apply simp apply(simp add:is_sublist_def) apply(erule exE)+ apply(rename_tac as bs) apply simp apply(rule_tac x = "x#as" in exI) apply(rule_tac x = bs in exI) apply simp done lemma Edges_append: "Edges(xs @ ys) = (if xs = [] then Edges ys else if ys = [] then Edges xs else Edges xs \ Edges ys \ {(last xs, hd ys)})" apply(induct xs) apply simp apply (simp add:Edges_Cons) apply blast done lemma Edges_rev_disj: "distinct xs \ Edges(rev xs) \ Edges(xs) = {}" apply(induct xs) apply simp apply(auto simp:Edges_Cons Edges_append last_rev notinset_notinEdge1 notinset_notinEdge2) done lemma disj_sets_disj_Edges: "set xs \ set ys = {} \ Edges xs \ Edges ys = {}" by(unfold Edges_def)(blast intro:is_sublist_in is_sublist_in1) lemma disj_sets_disj_Edges2: "set ys \ set xs = {} \ Edges xs \ Edges ys = {}" by(blast intro!:disj_sets_disj_Edges) lemma finite_Edges[iff]: "finite(Edges xs)" by(induct xs)(simp_all add:Edges_Cons) lemma Edges_compl: "\ distinct vs; x \ set vs; y \ set vs; x \ y \ \ Edges(x # between vs x y @ [y]) \ Edges(y # between vs y x @ [x]) = {}" using [[linarith_neq_limit = 1]] apply(subgoal_tac "\xs (ys::vertex list). xs \ [] \ set xs \ set ys = {} \ hd xs \ set ys") prefer 2 apply(drule hd_in_set) apply(blast) apply(frule split_list[of x]) apply clarsimp apply(erule disjE) apply(frule split_list[of y]) apply(clarsimp simp add:between_def split_def) apply (clarsimp simp add:Edges_Cons Edges_append notinset_notinEdge1 notinset_notinEdge2 disj_sets_disj_Edges disj_sets_disj_Edges2 Int_Un_distrib Int_Un_distrib2) apply(fastforce) apply(frule split_list[of y]) apply(clarsimp simp add:between_def split_def) apply (clarsimp simp add:Edges_Cons Edges_append notinset_notinEdge1 notinset_notinEdge2 disj_sets_disj_Edges disj_sets_disj_Edges2 Int_Un_distrib Int_Un_distrib2) apply fastforce done lemma Edges_disj: "\ distinct vs; x \ set vs; z \ set vs; x \ y; y \ z; y \ set(between vs x z) \ \ Edges(x # between vs x y @ [y]) \ Edges(y # between vs y z @ [z]) = {}" apply(subgoal_tac "\xs (ys::vertex list). xs \ [] \ set xs \ set ys = {} \ hd xs \ set ys") prefer 2 apply(drule hd_in_set) apply(blast) apply(frule split_list[of x]) apply clarsimp apply(erule disjE) apply simp apply(drule inbetween_inset) apply(rule Edges_compl) apply simp apply simp apply simp apply simp apply(erule disjE) apply(frule split_list[of z]) apply(clarsimp simp add:between_def split_def) apply(erule disjE) apply(frule split_list[of y]) apply clarsimp apply (clarsimp simp add:Edges_Cons Edges_append notinset_notinEdge1 notinset_notinEdge2 disj_sets_disj_Edges disj_sets_disj_Edges2 Int_Un_distrib Int_Un_distrib2) apply fastforce apply(frule split_list[of y]) apply clarsimp apply (clarsimp simp add:Edges_Cons Edges_append notinset_notinEdge1 notinset_notinEdge2 disj_sets_disj_Edges disj_sets_disj_Edges2 Int_Un_distrib Int_Un_distrib2) apply fastforce apply(frule split_list[of z]) apply(clarsimp simp add:between_def split_def) apply(frule split_list[of y]) apply clarsimp apply (clarsimp simp add:Edges_Cons Edges_append notinset_notinEdge1 notinset_notinEdge2 disj_sets_disj_Edges disj_sets_disj_Edges2 Int_Un_distrib Int_Un_distrib2) apply fastforce done lemma edges_conv_Un_Edges: "\ distinct(vertices(f::face)); x \ \ f; y \ \ f; x \ y \ \ \ f = Edges(x # between (vertices f) x y @ [y]) \ Edges(y # between (vertices f) y x @ [x])" apply(simp add:edges_conv_Edges) apply(rule conjI) apply clarsimp apply clarsimp apply(frule split_list[of x]) apply clarsimp apply(erule disjE) apply(frule split_list[of y]) apply(clarsimp simp add:between_def split_def) apply (clarsimp simp add:Edges_Cons Edges_append notinset_notinEdge1 notinset_notinEdge2 disj_sets_disj_Edges disj_sets_disj_Edges2 Int_Un_distrib Int_Un_distrib2) apply(fastforce) apply(frule split_list[of y]) apply(clarsimp simp add:between_def split_def) apply (clarsimp simp add:Edges_Cons Edges_append notinset_notinEdge1 notinset_notinEdge2 disj_sets_disj_Edges disj_sets_disj_Edges2 Int_Un_distrib Int_Un_distrib2) apply(fastforce) done lemma Edges_between_edges: "\ (a,b) \ Edges (u # between (vertices(f::face)) u v @ [v]); pre_split_face f u v vs \ \ (a,b) \ \ f" apply(simp add:pre_split_face_def) apply(induct f) apply(simp add:edges_conv_Edges Edges_Cons) apply clarify apply(rename_tac list) apply(case_tac "between list u v = []") apply simp apply(drule (4) is_nextElem_between_empty') apply(simp add:Edges_def) apply(subgoal_tac "pre_between list u v") prefer 2 apply (simp add:pre_between_def) apply(subgoal_tac "pre_between list v u") prefer 2 apply (simp add:pre_between_def) apply(case_tac "before list u v") apply(drule (1) between_eq2) apply clarsimp apply(erule disjE) apply (clarsimp simp:neq_Nil_conv) apply(rule is_nextElem_sublistI) apply(simp (no_asm) add:is_sublist_def) apply blast apply(rule is_nextElem_sublistI) apply(clarsimp simp add:Edges_def is_sublist_def) apply(rename_tac as bs cs xs ys) apply(rule_tac x = "as @ u # xs" in exI) apply(rule_tac x = "ys @ cs" in exI) apply simp apply (simp only:before_xor) apply(drule (1) between_eq2) apply clarsimp apply(rename_tac as bs cs) apply(erule disjE) apply (clarsimp simp:neq_Nil_conv) apply(case_tac cs) apply clarsimp apply(simp add:is_nextElem_def) apply simp apply(rule is_nextElem_sublistI) apply(simp (no_asm) add:is_sublist_def) apply(rule_tac x = "as @ v # bs" in exI) apply simp apply(rule_tac m1 = "|as|+1" in is_nextElem_rotate_eq[THEN iffD1]) apply simp apply(simp add:rotate_drop_take) apply(rule is_nextElem_sublistI) apply(clarsimp simp add:Edges_def is_sublist_def) apply(rename_tac xs ys) apply(rule_tac x = "bs @ u # xs" in exI) apply simp done (******************** split_face_edges ********************************) (* split_face *) lemma edges_split_face1: "pre_split_face f u v vs \ \(fst(split_face f u v vs)) = Edges(v # rev vs @ [u]) \ Edges(u # between (vertices f) u v @ [v])" apply(simp add: edges_conv_Edges split_face_distinct1') apply(auto simp add:split_face_def Edges_Cons Edges_append) done lemma edges_split_face2: "pre_split_face f u v vs \ \(snd(split_face f u v vs)) = Edges(u # vs @ [v]) \ Edges(v # between (vertices f) v u @ [u])" apply(simp add: edges_conv_Edges split_face_distinct2') apply(auto simp add:split_face_def Edges_Cons Edges_append) done lemma split_face_empty_ram1_ram2_in_f21: "pre_split_face oldF ram1 ram2 [] \ (f12, f21) = split_face oldF ram1 ram2 [] \ (ram1, ram2) \ edges f21" proof - assume split: "(f12, f21) = split_face oldF ram1 ram2 []" "pre_split_face oldF ram1 ram2 []" then have "ram1 \ \ f21" by (simp add: split_face_def) moreover with split have "f21 \ ram1 = hd (vertices f21)" apply (rule_tac nextElem_suc2) apply (simp add: pre_split_face_def split_face_distinct2) by (simp add: split_face_def) with split have "ram2 = f21 \ ram1" by (simp add: split_face_def) ultimately show ?thesis by (simp add: edges_face_def) qed lemma split_face_empty_ram1_ram2_in_f21': "pre_split_face oldF ram1 ram2 [] \ (ram1, ram2) \ edges (snd (split_face oldF ram1 ram2 []))" proof - assume split: "pre_split_face oldF ram1 ram2 []" define f12 where "f12 = fst (split_face oldF ram1 ram2 [])" define f21 where "f21 = snd (split_face oldF ram1 ram2 [])" then have "(f12, f21) = split_face oldF ram1 ram2 []" by (simp add: f12_def f21_def) with split have "(ram1, ram2) \ edges f21" by (rule split_face_empty_ram1_ram2_in_f21) with f21_def show ?thesis by simp qed lemma splitFace_empty_ram1_ram2_in_f21: "pre_splitFace g ram1 ram2 oldF [] \ (f12, f21, newGraph) = splitFace g ram1 ram2 oldF [] \ (ram1, ram2) \ edges f21" proof - assume pre: "pre_splitFace g ram1 ram2 oldF []" then have oldF: "oldF \ \ g" by (unfold pre_splitFace_def) simp assume sp: "(f12, f21, newGraph) = splitFace g ram1 ram2 oldF []" with oldF have "(f12, f21) = split_face oldF ram1 ram2 []" by (rule splitFace_split_face) with pre sp show ?thesis apply (unfold splitFace_def pre_splitFace_def) apply (simp add: split_def) apply (rule split_face_empty_ram1_ram2_in_f21') apply (rule pre_splitFace_pre_split_face) apply (rule pre) done qed lemma splitFace_f21_new_vertices: "(f12, f21, newGraph) = splitFace g ram1 ram2 oldF newVs \ v \ set newVs \ v \ \ f21" apply (unfold splitFace_def) apply (simp add: split_def) apply (unfold split_face_def) by simp lemma split_face_edges_f12: assumes vors: "pre_split_face f ram1 ram2 vs" "(f12, f21) = split_face f ram1 ram2 vs" "vs \ []" "vs1 = between (vertices f) ram1 ram2" "vs1 \ []" shows "edges f12 = {(hd vs, ram1) , (ram1, hd vs1), (last vs1, ram2), (ram2, last vs)} \ Edges(rev vs) \ Edges vs1" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto from x vors show "x \ ?rhs" apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f12) apply (case_tac "c = ram2 \ d = last vs") apply simp apply simp apply (elim exE) apply (case_tac "c = ram1") apply simp apply (subgoal_tac "between (vertices f) ram1 ram2 @ [ram2] = d # bs") apply (case_tac "between (vertices f) ram1 ram2") apply simp apply simp apply (rule dist_at2) apply (rule dist_f12) apply (rule sym) apply simp apply simp (* c \ ram1 *) apply (case_tac "c \ set(rev vs)") apply (subgoal_tac "distinct(rev vs)") apply (simp only: in_set_conv_decomp) apply (elim exE) apply simp apply (case_tac "zs") apply simp apply (subgoal_tac "ys = as") apply(drule last_rev) apply (simp) apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp apply (simp add:rev_swap) apply (subgoal_tac "ys = as") apply (clarsimp simp add: Edges_def is_sublist_def) apply (rule conjI) apply (subgoal_tac "\as bs. rev list @ [d, c] = as @ d # c # bs") apply simp apply (intro exI) apply simp apply (subgoal_tac "\asa bs. rev list @ d # c # rev as = asa @ d # c # bs") apply simp apply (intro exI) apply simp apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp apply (simp add: pre_split_face_def) (* c \ set vs *) apply (case_tac "c \ set (between (vertices f) ram1 ram2)") apply (subgoal_tac "distinct (between (vertices f) ram1 ram2)") apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp apply (case_tac zs) apply simp apply (subgoal_tac "rev vs @ ram1 # ys = as") apply force apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp apply simp apply (subgoal_tac "rev vs @ ram1 # ys = as") apply (simp add: Edges_def is_sublist_def) apply (subgoal_tac "(rev vs @ ram1 # ys) @ c # a # list @ [ram2] = as @ c # d # bs") apply simp apply (rule conjI) apply (rule impI) apply (rule disjI2)+ apply (rule exI) apply force apply (rule impI) apply (rule disjI2)+ apply force apply force apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp apply (thin_tac "rev vs @ ram1 # between (vertices f) ram1 ram2 @ [ram2] = as @ c # d # bs") apply (subgoal_tac "distinct (vertices f12)") apply simp apply (rule dist_f12) (* c \ set (between (vertices f) ram1 ram2) *) apply (subgoal_tac "c = ram2") apply simp apply (subgoal_tac "rev vs @ ram1 # between (vertices f) ram1 ram2 = as") apply force apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp (* c \ ram2 *) apply (subgoal_tac "c \ set (rev vs @ ram1 # between (vertices f) ram1 ram2 @ [ram2])") apply (thin_tac "rev vs @ ram1 # between (vertices f) ram1 ram2 @ [ram2] = as @ c # d # bs") apply simp by simp next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto from x vors show "x \ ?lhs" apply (simp add: dist_f12 is_nextElem_def is_sublist_def) apply (simp add: split_face_def) apply (case_tac "c = ram2 \ d = last vs") apply simp apply (rule disjCI) apply simp apply (case_tac "c = hd vs \ d = ram1") apply (case_tac "vs") apply simp apply force apply simp apply (case_tac "c = ram1 \ d = hd (between (vertices f) ram1 ram2)") apply (case_tac "between (vertices f) ram1 ram2") apply simp apply force apply simp apply (case_tac "c = last (between (vertices f) ram1 ram2) \ d = ram2") apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp apply simp apply (intro exI) apply (subgoal_tac "rev vs @ ram1 # ys @ [y, ram2] = (rev vs @ ram1 # ys) @ y # ram2 # []") apply assumption apply simp apply simp apply (case_tac "(d,c) \ Edges vs") apply (simp add: Edges_def is_sublist_def) apply (elim exE) apply simp apply (intro exI) apply simp apply (simp add: Edges_def is_sublist_def) apply (elim exE) apply simp apply (intro exI) apply (subgoal_tac "rev vs @ ram1 # as @ c # d # bs @ [ram2] = (rev vs @ ram1 # as) @ c # d # bs @ [ram2]") apply assumption by simp qed lemma split_face_edges_f12_vs: assumes vors: "pre_split_face f ram1 ram2 []" "(f12, f21) = split_face f ram1 ram2 []" "vs1 = between (vertices f) ram1 ram2" "vs1 \ []" shows "edges f12 = {(ram2, ram1) , (ram1, hd vs1), (last vs1, ram2)} \ Edges vs1" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto from x vors show "x \ ?rhs" apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f12) apply (case_tac " c = ram2 \ d = ram1") apply simp apply simp apply (elim exE) apply (case_tac "c = ram1") apply simp apply (subgoal_tac "as = []") apply simp apply (case_tac "between (vertices f) ram1 ram2") apply simp apply simp apply (rule dist_at1) apply (rule dist_f12) apply force apply (rule sym) apply simp (* a \ ram1 *) apply (case_tac "c \ set (between (vertices f) ram1 ram2)") apply (subgoal_tac "distinct (between (vertices f) ram1 ram2)") apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp apply (case_tac zs) apply simp apply (subgoal_tac "ram1 # ys = as") apply force apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp apply simp apply (subgoal_tac "ram1 # ys = as") apply (simp add: Edges_def is_sublist_def) apply (subgoal_tac "(ram1 # ys) @ c # a # list @ [ram2] = as @ c # d # bs") apply simp apply (rule conjI) apply (rule impI) apply (rule disjI2)+ apply (rule exI) apply force apply (rule impI) apply (rule disjI2)+ apply force apply force apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp apply (thin_tac "ram1 # between (vertices f) ram1 ram2 @ [ram2] = as @ c # d # bs") apply (subgoal_tac "distinct (vertices f12)") apply simp apply (rule dist_f12) (* c \ set (between (vertices f) ram1 ram2) *) apply (subgoal_tac "c = ram2") apply simp apply (subgoal_tac "ram1 # between (vertices f) ram1 ram2 = as") apply force apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp (* c \ ram2 *) apply (subgoal_tac "c \ set (ram1 # between (vertices f) ram1 ram2 @ [ram2])") apply (thin_tac "ram1 # between (vertices f) ram1 ram2 @ [ram2] = as @ c # d # bs") apply simp by simp next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto from x vors show "x \ ?lhs" apply (simp add: dist_f12 is_nextElem_def is_sublist_def) apply (simp add: split_face_def) apply (case_tac "c = ram2 \ d = ram1") apply simp apply (rule disjCI) apply simp apply (case_tac "c = ram1 \ d = hd (between (vertices f) ram1 ram2)") apply (case_tac "between (vertices f) ram1 ram2") apply simp apply force apply simp apply (case_tac "c = last (between (vertices f) ram1 ram2) \ d = ram2") apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp apply simp apply (intro exI) apply (subgoal_tac "ram1 # ys @ [y, ram2] = (ram1 # ys) @ y # ram2 # []") apply assumption apply simp apply simp apply (case_tac "(c, d) \ Edges vs") apply (simp add: Edges_def is_sublist_def) apply (elim exE) apply simp apply (intro exI) apply (subgoal_tac "ram1 # as @ c # d # bs @ [ram2] = (ram1 # as) @ c # d # (bs @ [ram2])") apply assumption apply simp apply (simp add: Edges_def is_sublist_def) apply (elim exE) apply simp apply (intro exI) apply (subgoal_tac "ram1 # as @ c # d # bs @ [ram2] = (ram1 # as) @ c # d # bs @ [ram2]") apply assumption by simp qed lemma split_face_edges_f12_bet: assumes vors: "pre_split_face f ram1 ram2 vs" "(f12, f21) = split_face f ram1 ram2 vs" "vs \ []" "between (vertices f) ram1 ram2 = []" shows "edges f12 = {(hd vs, ram1) , (ram1, ram2), (ram2, last vs)} \ Edges(rev vs)" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto from x vors show "x \ ?rhs" apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f12) apply (case_tac " c = ram2 \ d = last vs") apply simp apply simp apply (elim exE) apply (case_tac "c = ram1") apply simp apply (subgoal_tac "rev vs = as") apply simp apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp (* a \ ram1 *) apply (case_tac "c \ set(rev vs)") apply (subgoal_tac "distinct(rev vs)") apply (simp only: in_set_conv_decomp) apply (elim exE) apply simp apply (case_tac "zs") apply simp apply (subgoal_tac "ys = as") apply (simp add:rev_swap) apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp apply simp apply (subgoal_tac "ys = as") apply (simp add: Edges_def is_sublist_def rev_swap) apply (rule conjI) apply (subgoal_tac "\as bs. rev list @ [d, c] = as @ d # c # bs") apply simp apply (intro exI) apply simp apply (subgoal_tac "\asa bs. rev list @ d # c # rev as = asa @ d # c # bs") apply simp apply (intro exI) apply simp apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp apply (simp add: pre_split_face_def) (* c \ set vs *) apply (subgoal_tac "c = ram2") apply simp apply (subgoal_tac "rev vs @ [ram1] = as") apply force apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp (* c \ ram2 *) apply (subgoal_tac "c \ set (rev vs @ ram1 # [ram2])") apply (thin_tac "rev vs @ ram1 # [ram2] = as @ c # d # bs") apply simp by simp next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto from x vors show "x \ ?lhs" apply (simp add: dist_f12 is_nextElem_def is_sublist_def) apply (simp add: split_face_def) apply (case_tac "c = hd vs \ d = ram1") apply (case_tac "vs") apply simp apply force apply simp apply (case_tac "c = ram1 \ d = ram2") apply force apply simp apply (case_tac "c = ram2 \ d = last vs") apply (case_tac "vs" rule:rev_exhaust) apply simp apply simp apply (simp add: Edges_def is_sublist_def) apply (elim exE) apply simp apply (rule conjI) apply (rule impI) apply (rule disjI1) apply (intro exI) apply (subgoal_tac "c # d # rev as @ [ram1, ram2] = [] @ c # d # rev as @ [ram1,ram2]") apply assumption apply simp apply (rule impI) apply (rule disjI1) apply (intro exI) by simp qed lemma split_face_edges_f12_bet_vs: assumes vors: "pre_split_face f ram1 ram2 []" "(f12, f21) = split_face f ram1 ram2 []" "between (vertices f) ram1 ram2 = []" shows "edges f12 = {(ram2, ram1) , (ram1, ram2)}" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto from x vors show "x \ ?rhs" apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f12) apply (case_tac " c = ram2 \ d = ram1") apply force apply simp apply (elim exE) apply (case_tac "c = ram1") apply simp apply (case_tac "as") apply simp apply (case_tac "list") apply simp apply simp apply (case_tac "as") apply simp apply (case_tac "list") apply simp by simp next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto from x vors show "x \ ?lhs" apply (simp add: dist_f12 is_nextElem_def is_sublist_def) apply (simp add: split_face_def) by auto qed lemma split_face_edges_f12_subset: "pre_split_face f ram1 ram2 vs \ (f12, f21) = split_face f ram1 ram2 vs \ vs \ [] \ {(hd vs, ram1), (ram2, last vs)} \ Edges(rev vs) \ edges f12" apply (case_tac "between (vertices f) ram1 ram2") apply (frule split_face_edges_f12_bet) apply simp apply simp apply simp apply force apply (frule split_face_edges_f12) apply simp+ by force (*declare in_Edges_rev [simp del] rev_eq_Cons_iff [simp del]*) lemma split_face_edges_f21: assumes vors: "pre_split_face f ram1 ram2 vs" "(f12, f21) = split_face f ram1 ram2 vs" "vs \ []" "vs2 = between (vertices f) ram2 ram1" "vs2 \ []" shows "edges f21 = {(last vs2, ram1) , (ram1, hd vs), (last vs, ram2), (ram2, hd vs2)} \ Edges vs \ Edges vs2" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto from x vors show "x \ ?rhs" apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f21) apply (case_tac " c = last vs \ d = ram2") apply simp apply simp apply (elim exE) apply (case_tac "c = ram1") apply simp apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 = as") apply clarsimp apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp (* a \ ram1 *) apply (case_tac "c \ set vs") apply (subgoal_tac "distinct vs") apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp apply (case_tac "zs") apply simp apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # ys = as") apply force apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp apply simp apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # ys = as") apply (subgoal_tac "(ram2 # between (vertices f) ram2 ram1 @ ram1 # ys) @ c # a # list = as @ c # d # bs") apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # ys @ c # a # list = as @ c # d # bs") apply (simp add: Edges_def is_sublist_def) apply(rule conjI) apply (subgoal_tac "\as bs. ys @ [c, d] = as @ c # d # bs") apply simp apply force apply force apply force apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp apply (simp add: pre_split_face_def) (* c \ set (rev vs) *) apply (case_tac "c \ set (between (vertices f) ram2 ram1)") apply (subgoal_tac "distinct (between (vertices f) ram2 ram1)") apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp apply (case_tac zs) apply simp apply (subgoal_tac "ram2 # ys = as") apply force apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp apply simp apply (subgoal_tac "ram2 # ys = as") apply (simp add: Edges_def is_sublist_def) apply (subgoal_tac "(ram2 # ys) @ c # a # list @ ram1 # vs = as @ c # d # bs") apply (thin_tac "ram2 # ys @ c # a # list @ ram1 # vs = as @ c # d # bs") apply (rule conjI) apply (rule impI) apply (rule disjI2)+ apply force apply (rule impI) apply (rule disjI2)+ apply force apply force apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp apply (subgoal_tac "distinct (vertices f21)") apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # vs = as @ c # d # bs") apply simp apply (rule dist_f21) (* c \ set (between (vertices f) ram2 ram1) *) apply (subgoal_tac "c = ram2") apply simp apply (subgoal_tac "[] = as") apply simp apply (case_tac "between (vertices f) ram2 ram1") apply simp apply simp apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp (* c \ ram2 *) apply (subgoal_tac "c \ set (ram2 # between (vertices f) ram2 ram1 @ ram1 # vs)") apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # vs = as @ c # d # bs") apply simp by simp next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto from x vors show "x \ ?lhs" apply (simp add: dist_f21 is_nextElem_def is_sublist_def) apply (simp add: split_face_def) apply (case_tac "c = ram2 \ d = hd (between (vertices f) ram2 ram1)") apply simp apply (rule disjI1) apply (intro exI) apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # vs = [] @ ram2 # hd (between (vertices f) ram2 ram1) # tl (between (vertices f) ram2 ram1) @ ram1 # vs") apply assumption apply simp apply (case_tac "c = ram1 \ d = hd vs") apply (rule disjI1) apply (case_tac "vs") apply simp apply simp apply (intro exI) apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # a # list = (ram2 # between (vertices f) ram2 ram1) @ ram1 # a # list") apply assumption apply simp apply (case_tac "c = last vs \ d = ram2") apply (case_tac vs rule:rev_exhaust) apply simp apply simp apply simp apply (case_tac "c = last (between (vertices f) ram2 ram1) \ d = ram1") apply (rule disjI1) apply (case_tac "between (vertices f) ram2 ram1" rule: rev_exhaust) apply simp apply (intro exI) apply simp apply (subgoal_tac "ram2 # ys @ y # ram1 # vs = (ram2 # ys) @ y # ram1 # vs") apply assumption apply simp apply simp apply (case_tac "(c, d) \ Edges vs") apply (simp add: Edges_def is_sublist_def) apply (elim exE) apply simp apply (rule conjI) apply (rule impI) apply (rule disjI1) apply (intro exI) apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # as @ [c, d] = (ram2 # between (vertices f) ram2 ram1 @ ram1 # as) @ c # d # []") apply assumption apply simp apply (rule impI) apply (rule disjI1) apply (intro exI) apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # as @ c # d # bs = (ram2 # between (vertices f) ram2 ram1 @ ram1 # as) @ c # d # bs") apply assumption apply simp apply (simp add: Edges_def is_sublist_def) apply (elim exE) apply simp apply (rule disjI1) apply (intro exI) apply (subgoal_tac "ram2 # as @ c # d # bs @ ram1 # vs = (ram2 # as) @ c # d # (bs @ ram1 # vs)") apply assumption by simp qed lemma split_face_edges_f21_vs: assumes vors: "pre_split_face f ram1 ram2 []" "(f12, f21) = split_face f ram1 ram2 []" "vs2 = between (vertices f) ram2 ram1" "vs2 \ []" shows "edges f21 = {(last vs2, ram1) , (ram1, ram2), (ram2, hd vs2)} \ Edges vs2" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto from x vors show "x \ ?rhs" apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f21) apply (case_tac " c = ram1 \ d = ram2") apply simp apply simp apply (elim exE) apply (case_tac "c = ram1") apply simp apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 = as") apply (subgoal_tac "(ram2 # between (vertices f) ram2 ram1) @ [ram1] = as @ ram1 # d # bs") apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ [ram1] = as @ ram1 # d # bs") apply simp apply force apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp (* a \ ram1 *) apply (case_tac "c \ set (between (vertices f) ram2 ram1)") apply (subgoal_tac "distinct (between (vertices f) ram2 ram1)") apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp apply (case_tac zs) apply simp apply (subgoal_tac "ram2 # ys = as") apply force apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp apply simp apply (subgoal_tac "ram2 # ys = as") apply (simp add: Edges_def is_sublist_def) apply (subgoal_tac "(ram2 # ys) @ c # a # list @ [ram1] = as @ c # d # bs") apply (thin_tac "ram2 # ys @ c # a # list @ [ram1] = as @ c # d # bs") apply (rule conjI) apply (rule impI) apply (rule disjI2)+ apply force apply (rule impI) apply (rule disjI2)+ apply force apply force apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp apply (subgoal_tac "distinct (vertices f21)") apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ [ram1] = as @ c # d # bs") apply simp apply (rule dist_f21) (* c \ set (between (vertices f) ram2 ram1) *) apply (subgoal_tac "c = ram2") apply simp apply (subgoal_tac "[] = as") apply simp apply (case_tac "between (vertices f) ram2 ram1") apply simp apply simp apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp (* c \ ram2 *) apply (subgoal_tac "c \ set (ram2 # between (vertices f) ram2 ram1 @ [ram1])") apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ [ram1] = as @ c # d # bs") apply simp by simp next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto from x vors show "x \ ?lhs" apply (simp add: dist_f21 is_nextElem_def is_sublist_def) apply (simp add: split_face_def) apply (case_tac "c = ram2 \ d = hd (between (vertices f) ram2 ram1)") apply simp apply (rule disjI1) apply (intro exI) apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ [ram1] = [] @ ram2 # hd (between (vertices f) ram2 ram1) # tl (between (vertices f) ram2 ram1) @ [ram1]") apply assumption apply simp apply (case_tac "c = ram1 \ d = ram2") apply (rule disjI2) apply simp apply simp apply (case_tac "c = last (between (vertices f) ram2 ram1) \ d = ram1") apply (rule disjI1) apply (case_tac "between (vertices f) ram2 ram1" rule: rev_exhaust) apply simp apply (intro exI) apply simp apply (subgoal_tac "ram2 # ys @ y # [ram1] = (ram2 # ys) @ y # [ram1]") apply assumption apply simp apply simp apply (simp add: Edges_def is_sublist_def) apply (elim exE) apply simp apply (rule disjI1) apply (intro exI) apply (subgoal_tac "ram2 # as @ c # d # bs @ [ram1] = (ram2 # as) @ c # d # (bs @ [ram1])") apply assumption by simp qed lemma split_face_edges_f21_bet: assumes vors: "pre_split_face f ram1 ram2 vs" "(f12, f21) = split_face f ram1 ram2 vs" "vs \ []" "between (vertices f) ram2 ram1 = []" shows "edges f21 = {(ram1, hd vs), (last vs, ram2), (ram2, ram1)} \ Edges vs" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto from x vors show "x \ ?rhs" apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f21) apply (case_tac " c = last vs \ d = ram2") apply simp apply simp apply (elim exE) apply (case_tac "c = ram1") apply simp apply (subgoal_tac "[ram2] = as") apply clarsimp apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp (* a \ ram1 *) apply (case_tac "c \ set vs") apply (subgoal_tac "distinct vs") apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp apply (case_tac "zs") apply simp apply (subgoal_tac "ram2 # ram1 # ys = as") apply force apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp apply simp apply (subgoal_tac "ram2 # ram1 # ys = as") apply (subgoal_tac "(ram2 # ram1 # ys) @ c # a # list = as @ c # d # bs") apply (thin_tac "ram2 # ram1 # ys @ c # a # list = as @ c # d # bs") apply (simp add: Edges_def is_sublist_def) apply force apply force apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp apply (simp add: pre_split_face_def) (* c \ set (rev vs) *) apply (subgoal_tac "c = ram2") apply simp apply (subgoal_tac "[] = as") apply simp apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp (* c \ ram2 *) apply (subgoal_tac "c \ set (ram2 # ram1 # vs)") apply (thin_tac "ram2 # ram1 # vs = as @ c # d # bs") apply simp by simp next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto from x vors show "x \ ?lhs" apply (simp add: dist_f21 is_nextElem_def is_sublist_def) apply (simp add: split_face_def) apply (case_tac "c = ram2 \ d = ram1") apply simp apply (rule disjI1) apply force apply (case_tac "c = ram1 \ d = hd vs") apply (rule disjI1) apply (case_tac "vs") apply simp apply simp apply (intro exI) apply (subgoal_tac "ram2 # ram1 # a # list = [ram2] @ ram1 # a # list") apply assumption apply simp apply (case_tac "c = last vs \ d = ram2") apply (case_tac vs rule: rev_exhaust) apply simp apply simp apply (simp add: Edges_def is_sublist_def) apply (elim exE) apply simp apply (rule conjI) apply (rule impI) apply (rule disjI1) apply (intro exI) apply (subgoal_tac "ram2 # ram1 # as @ [c, d] = (ram2 # ram1 # as) @ c # d # []") apply assumption apply simp apply (rule impI) apply (rule disjI1) apply (intro exI) apply (subgoal_tac "ram2 # ram1 # as @ c # d # bs = (ram2 # ram1 # as) @ c # d # bs") apply assumption by simp qed lemma split_face_edges_f21_bet_vs: assumes vors: "pre_split_face f ram1 ram2 []" "(f12, f21) = split_face f ram1 ram2 []" "between (vertices f) ram2 ram1 = []" shows "edges f21 = {(ram1, ram2), (ram2, ram1)}" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto from x vors show "x \ ?rhs" apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f21) apply (case_tac " c = ram1 \ d = ram2") apply simp apply simp apply (elim exE) apply (case_tac "as") apply simp apply (case_tac "list") by auto next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto from x vors show "x \ ?lhs" apply (simp add: dist_f21 is_nextElem_def is_sublist_def) apply (simp add: split_face_def) by auto qed lemma split_face_edges_f21_subset: "pre_split_face f ram1 ram2 vs \ (f12, f21) = split_face f ram1 ram2 vs \ vs \ [] \ {(last vs, ram2), (ram1, hd vs)} \ Edges vs \ edges f21" apply (case_tac "between (vertices f) ram2 ram1") apply (frule split_face_edges_f21_bet) apply simp apply simp apply simp apply force apply (frule split_face_edges_f21) apply simp+ by force lemma verticesFrom_ram1: "pre_split_face f ram1 ram2 vs \ verticesFrom f ram1 = ram1 # between (vertices f) ram1 ram2 @ ram2 # between (vertices f) ram2 ram1" apply (subgoal_tac "pre_between (vertices f) ram1 ram2") apply (subgoal_tac "distinct (vertices f)") apply (case_tac "before (vertices f) ram1 ram2") apply (simp add: verticesFrom_Def) apply (subgoal_tac "ram2 \ set (snd (splitAt ram1 (vertices f)))") apply (drule splitAt_ram) apply (subgoal_tac "snd (splitAt ram2 (snd (splitAt ram1 (vertices f)))) = snd (splitAt ram2 (vertices f))") apply simp apply (thin_tac "snd (splitAt ram1 (vertices f)) = fst (splitAt ram2 (snd (splitAt ram1 (vertices f)))) @ ram2 # snd (splitAt ram2 (snd (splitAt ram1 (vertices f))))") apply simp apply (rule before_dist_r2) apply simp apply simp apply (subgoal_tac "before (vertices f) ram2 ram1") apply (subgoal_tac "pre_between (vertices f) ram2 ram1") apply (simp add: verticesFrom_Def) apply (subgoal_tac "ram2 \ set (fst (splitAt ram1 (vertices f)))") apply (drule splitAt_ram) apply (subgoal_tac "fst (splitAt ram2 (fst (splitAt ram1 (vertices f)))) = fst (splitAt ram2 (vertices f))") apply simp apply (thin_tac "fst (splitAt ram1 (vertices f)) = fst (splitAt ram2 (fst (splitAt ram1 (vertices f)))) @ ram2 # snd (splitAt ram2 (fst (splitAt ram1 (vertices f))))") apply simp apply (rule before_dist_r1) apply simp apply simp apply (simp add: pre_between_def) apply force apply (simp add: pre_split_face_def) by (rule pre_split_face_p_between) lemma split_face_edges_f_vs1_vs2: assumes vors: "pre_split_face f ram1 ram2 vs" "between (vertices f) ram1 ram2 = []" "between (vertices f) ram2 ram1 = []" shows "edges f = {(ram2, ram1) , (ram1, ram2)}" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def) from x vors show "x \ ?rhs" apply (simp add: dist_f) apply (subgoal_tac "pre_between (vertices f) ram1 ram2") apply (drule is_nextElem_or) apply assumption apply (simp add: Edges_def) apply (case_tac "is_sublist [c, d] [ram1, ram2]") apply (simp) apply (simp) apply blast by (rule pre_split_face_p_between) next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def) from x vors show "x \ ?lhs" apply (simp add: dist_f) apply (subgoal_tac "ram1 \ \ f") apply (simp add: verticesFrom_is_nextElem verticesFrom_ram1) apply (simp add: is_nextElem_def) apply blast by (simp add: pre_split_face_def) qed lemma split_face_edges_f_vs1: assumes vors: "pre_split_face f ram1 ram2 vs" "between (vertices f) ram1 ram2 = []" "vs2 = between (vertices f) ram2 ram1" "vs2 \ []" shows "edges f = {(last vs2, ram1) , (ram1, ram2), (ram2, hd vs2)} \ Edges vs2" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def) from vors have dist_vs2: "distinct (ram2 # vs2 @ [ram1])" apply (simp only:) apply (rule between_distinct_r12) apply (rule dist_f) apply (rule not_sym) by (simp add: pre_split_face_def) from x vors show "x \ ?rhs" apply (simp add: dist_f) apply (subgoal_tac "pre_between (vertices f) ram1 ram2") apply (drule is_nextElem_or) apply assumption apply (simp add: Edges_def) apply (case_tac "is_sublist [c, d] [ram1, ram2]") apply simp apply simp apply(erule disjE) apply blast apply (case_tac "c = ram2") apply (case_tac "between (vertices f) ram2 ram1") apply simp apply simp apply (drule is_sublist_distinct_prefix) apply (subgoal_tac "distinct (ram2 # vs2 @ [ram1])") apply simp apply (rule dist_vs2) apply simp apply (case_tac "c = ram1") apply (subgoal_tac "\ is_sublist [c, d] (ram2 # vs2 @ [ram1])") apply simp apply (rule is_sublist_notlast) apply (rule_tac dist_vs2) apply simp apply simp apply (simp add: is_sublist_def) apply (elim exE) apply (case_tac "between (vertices f) ram2 ram1" rule: rev_exhaust) apply simp apply simp apply (case_tac "bs" rule: rev_exhaust) apply simp apply simp apply (rule disjI2) apply (intro exI) apply simp apply (rule pre_split_face_p_between) by simp next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def) from vors have dist_vs2: "distinct (ram2 # vs2 @ [ram1])" apply (simp only:) apply (rule between_distinct_r12) apply (rule dist_f) apply (rule not_sym) by (simp add: pre_split_face_def) from x vors show "x \ ?lhs" apply (simp add: dist_f) apply (subgoal_tac "ram1 \ set (vertices f)") apply (simp add: verticesFrom_is_nextElem verticesFrom_ram1) apply (simp add: is_nextElem_def) apply (case_tac "c = last (between (vertices f) ram2 ram1) \ d = ram1") apply simp apply simp apply (rule disjI1) apply (case_tac "c = ram1 \ d = ram2") apply (simp add: is_sublist_def) apply force apply simp apply (case_tac "c = ram2 \ d = hd (between (vertices f) ram2 ram1)") apply (case_tac "between (vertices f) ram2 ram1") apply simp apply (simp add: is_sublist_def) apply (intro exI) apply (subgoal_tac "ram1 # ram2 # a # list = [ram1] @ ram2 # a # (list)") apply assumption apply simp apply simp apply (subgoal_tac "is_sublist [c, d] ((ram1 # [ram2]) @ between (vertices f) ram2 ram1 @ [])") apply simp apply (rule is_sublist_add) apply (simp add: Edges_def) by (simp add: pre_split_face_def) qed lemma split_face_edges_f_vs2: assumes vors: "pre_split_face f ram1 ram2 vs" "vs1 = between (vertices f) ram1 ram2" "vs1 \ []" "between (vertices f) ram2 ram1 = []" shows "edges f = {(ram2, ram1) , (ram1, hd vs1), (last vs1, ram2)} \ Edges vs1" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def) from vors have dist_vs1: "distinct (ram1 # vs1 @ [ram2])" apply (simp only:) apply (rule between_distinct_r12) apply (rule dist_f) by (simp add: pre_split_face_def) from x vors show "x \ ?rhs" apply (simp add: dist_f) apply (subgoal_tac "pre_between (vertices f) ram1 ram2") apply (drule is_nextElem_or) apply assumption apply (simp add: Edges_def) apply (case_tac "is_sublist [c, d] (ram1 # between (vertices f) ram1 ram2 @ [ram2])") apply simp apply (case_tac "c = ram1") apply (case_tac "between (vertices f) ram1 ram2") apply simp apply simp apply (drule is_sublist_distinct_prefix) apply (subgoal_tac "distinct (ram1 # vs1 @ [ram2])") apply simp apply (rule dist_vs1) apply simp apply (case_tac "c = ram2") apply (subgoal_tac "\ is_sublist [c, d] (ram1 # vs1 @ [ram2])") apply simp apply (rule is_sublist_notlast) apply (rule_tac dist_vs1) apply simp apply simp apply (simp add: is_sublist_def) apply (elim exE) apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp apply simp apply (case_tac "bs" rule: rev_exhaust) apply simp apply simp apply (rule disjI2) apply (intro exI) apply simp apply simp apply (rule pre_split_face_p_between) by simp next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def) from vors have dist_vs1: "distinct (ram1 # vs1 @ [ram2])" apply (simp only:) apply (rule between_distinct_r12) apply (rule dist_f) by (simp add: pre_split_face_def) from x vors show "x \ ?lhs" apply (simp add: dist_f) apply (subgoal_tac "ram1 \ \ f") apply (simp add: verticesFrom_is_nextElem verticesFrom_ram1) apply (simp add: is_nextElem_def) apply (case_tac "c = ram2 \ d = ram1") apply simp apply simp apply (rule disjI1) apply (case_tac "c = ram1 \ d = hd (between (vertices f) ram1 ram2)") apply (case_tac "between (vertices f) ram1 ram2") apply simp apply (force simp: is_sublist_def) apply simp apply (case_tac "c = last (between (vertices f) ram1 ram2) \ d = ram2") apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp apply (simp add: is_sublist_def) apply (intro exI) apply (subgoal_tac "ram1 # ys @ [y, ram2] = (ram1 # ys) @ y # ram2 # []") apply assumption apply simp apply simp apply (simp add: Edges_def) apply (subgoal_tac "is_sublist [c, d] ([ram1] @ between (vertices f) ram1 ram2 @ [ram2])") apply simp apply (rule is_sublist_add) apply simp by (simp add: pre_split_face_def) qed lemma split_face_edges_f: assumes vors: "pre_split_face f ram1 ram2 vs" "vs1 = between (vertices f) ram1 ram2" "vs1 \ []" "vs2 = between (vertices f) ram2 ram1" "vs2 \ []" shows "edges f = {(last vs2, ram1) , (ram1, hd vs1), (last vs1, ram2), (ram2, hd vs2)} \ Edges vs1 \ Edges vs2" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x: "x \ ?lhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def) from vors have dist_vs1: "distinct (ram1 # vs1 @ [ram2])" apply (simp only:) apply (rule between_distinct_r12) apply (rule dist_f) by (simp add: pre_split_face_def) from vors have dist_vs2: "distinct (ram2 # vs2 @ [ram1])" apply (simp only:) apply (rule between_distinct_r12) apply (rule dist_f) apply (rule not_sym) by (simp add: pre_split_face_def) from x vors show "x \ ?rhs" apply (simp add: dist_f) apply (subgoal_tac "pre_between (vertices f) ram1 ram2") apply (drule is_nextElem_or) apply assumption apply (simp add: Edges_def) apply (case_tac "is_sublist [c, d] (ram1 # between (vertices f) ram1 ram2 @ [ram2])") apply simp apply (case_tac "c = ram1") apply (case_tac "between (vertices f) ram1 ram2") apply simp apply simp apply (drule is_sublist_distinct_prefix) apply (subgoal_tac "distinct (ram1 # vs1 @ [ram2])") apply simp apply (rule dist_vs1) apply simp apply (case_tac "c = ram2") apply (subgoal_tac "\ is_sublist [c, d] (ram1 # vs1 @ [ram2])") apply simp apply (rule is_sublist_notlast) apply (rule_tac dist_vs1) apply simp apply simp apply (simp add: is_sublist_def) apply (elim exE) apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp apply simp apply (case_tac "bs" rule: rev_exhaust) apply simp apply simp apply (rule disjI2) apply (rule disjI2) apply (rule disjI1) apply (intro exI) apply simp apply simp apply (case_tac "c = ram2") apply (case_tac "between (vertices f) ram2 ram1") apply simp apply simp apply (drule is_sublist_distinct_prefix) apply (subgoal_tac "distinct (ram2 # vs2 @ [ram1])") apply simp apply (rule dist_vs2) apply simp apply (case_tac "c = ram1") apply (subgoal_tac "\ is_sublist [c, d] (ram2 # vs2 @ [ram1])") apply simp apply (rule is_sublist_notlast) apply (rule_tac dist_vs2) apply simp apply simp apply (simp add: is_sublist_def) apply (elim exE) apply (case_tac "between (vertices f) ram2 ram1" rule: rev_exhaust) apply simp apply simp apply (case_tac "bs" rule: rev_exhaust) apply simp apply simp apply (rule disjI2) apply (rule disjI2) apply (rule disjI2) apply (intro exI) apply simp apply (rule pre_split_face_p_between) by simp next fix x assume x: "x \ ?rhs" define c where "c = fst x" define d where "d = snd x" with c_def have [simp]: "x = (c,d)" by simp from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def) from vors have dist_vs1: "distinct (ram1 # vs1 @ [ram2])" apply (simp only:) apply (rule between_distinct_r12) apply (rule dist_f) by (simp add: pre_split_face_def) from vors have dist_vs2: "distinct (ram2 # vs2 @ [ram1])" apply (simp only:) apply (rule between_distinct_r12) apply (rule dist_f) apply (rule not_sym) by (simp add: pre_split_face_def) from x vors show "x \ ?lhs" apply (simp add: dist_f) apply (subgoal_tac "ram1 \ \ f") apply (simp add: verticesFrom_is_nextElem verticesFrom_ram1) apply (simp add: is_nextElem_def) apply (case_tac "c = last (between (vertices f) ram2 ram1) \ d = ram1") apply simp apply simp apply (rule disjI1) apply (case_tac "c = ram1 \ d = hd (between (vertices f) ram1 ram2)") apply (case_tac "between (vertices f) ram1 ram2") apply simp apply (force simp: is_sublist_def) apply simp apply (case_tac "c = last (between (vertices f) ram1 ram2) \ d = ram2") apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp apply (simp add: is_sublist_def) apply (intro exI) apply (subgoal_tac "ram1 # ys @ y # ram2 # between (vertices f) ram2 ram1 = (ram1 # ys) @ y # ram2 # (between (vertices f) ram2 ram1)") apply assumption apply simp apply simp apply (case_tac "c = ram2 \ d = hd (between (vertices f) ram2 ram1)") apply (case_tac "between (vertices f) ram2 ram1") apply simp apply (simp add: is_sublist_def) apply (intro exI) apply (subgoal_tac "ram1 # between (vertices f) ram1 ram2 @ ram2 # a # list = (ram1 # between (vertices f) ram1 ram2) @ ram2 # a # (list)") apply assumption apply simp apply simp apply (case_tac "(c, d) \ Edges (between (vertices f) ram1 ram2)") apply (simp add: Edges_def) apply (subgoal_tac "is_sublist [c, d] ([ram1] @ between (vertices f) ram1 ram2 @ (ram2 # between (vertices f) ram2 ram1))") apply simp apply (rule is_sublist_add) apply simp apply simp apply (subgoal_tac "is_sublist [c, d] ((ram1 # between (vertices f) ram1 ram2 @ [ram2]) @ between (vertices f) ram2 ram1 @ [])") apply simp apply (rule is_sublist_add) apply (simp add: Edges_def) by (simp add: pre_split_face_def) qed lemma split_face_edges_f12_f21: "pre_split_face f ram1 ram2 vs \ (f12, f21) = split_face f ram1 ram2 vs \ vs \ [] \ edges f12 \ edges f21 = edges f \ {(hd vs, ram1), (ram1, hd vs), (last vs, ram2), (ram2, last vs)} \ Edges vs \ Edges (rev vs)" apply (case_tac "between (vertices f) ram1 ram2 = []") apply (case_tac "between (vertices f) ram2 ram1 = []") apply (simp add: split_face_edges_f12_bet split_face_edges_f21_bet split_face_edges_f_vs1_vs2) apply force apply (simp add: split_face_edges_f12_bet split_face_edges_f21 split_face_edges_f_vs1) apply force apply (case_tac "between (vertices f) ram2 ram1 = []") apply (simp add: split_face_edges_f21_bet split_face_edges_f12 split_face_edges_f_vs2) apply force apply (simp add: split_face_edges_f21 split_face_edges_f12 split_face_edges_f) by force lemma split_face_edges_f12_f21_vs: "pre_split_face f ram1 ram2 [] \ (f12, f21) = split_face f ram1 ram2 [] \ edges f12 \ edges f21 = edges f \ {(ram2, ram1), (ram1, ram2)}" apply (case_tac "between (vertices f) ram1 ram2 = []") apply (case_tac "between (vertices f) ram2 ram1 = []") apply (simp add: split_face_edges_f12_bet_vs split_face_edges_f21_bet_vs split_face_edges_f_vs1_vs2) apply force apply (simp add: split_face_edges_f12_bet_vs split_face_edges_f21_vs split_face_edges_f_vs1) apply force apply (case_tac "between (vertices f) ram2 ram1 = []") apply (simp add: split_face_edges_f21_bet_vs split_face_edges_f12_vs split_face_edges_f_vs2) apply force apply (simp add: split_face_edges_f21_vs split_face_edges_f12_vs split_face_edges_f) by force lemma split_face_edges_f12_f21_sym: "f \ \ g \ pre_split_face f ram1 ram2 vs \ (f12, f21) = split_face f ram1 ram2 vs \ ((a,b) \ edges f12 \ (a,b) \ edges f21) = ((a,b) \ edges f \ (((b,a) \ edges f12 \ (b,a) \ edges f21) \ ((a,b) \ edges f12 \ (a,b) \ edges f21)))" apply (subgoal_tac "((a,b) \ edges f12 \ edges f21) = ((a,b) \ edges f \ ((b,a) \ edges f12 \ edges f21) \ (a,b) \ edges f12 \ edges f21)") apply force apply (case_tac "vs = []") apply (subgoal_tac "pre_split_face f ram1 ram2 []") apply (drule split_face_edges_f12_f21_vs) apply simp apply simp apply force apply simp apply (drule split_face_edges_f12_f21) apply simp apply simp apply simp by force lemma splitFace_edges_g'_help: "pre_splitFace g ram1 ram2 f vs \ (f12, f21, g') = splitFace g ram1 ram2 f vs \ vs \ [] \ edges g' = edges g \ edges f \ Edges vs \ Edges(rev vs) \ {(ram2, last vs), (hd vs, ram1), (ram1, hd vs), (last vs, ram2)}" proof - assume pre: "pre_splitFace g ram1 ram2 f vs" and fdg: "(f12, f21, g') = splitFace g ram1 ram2 f vs" and vs_neq: "vs \ []" from pre fdg have split: "(f12, f21) = split_face f ram1 ram2 vs" apply (unfold pre_splitFace_def) apply (elim conjE) by (simp add: splitFace_split_face) from fdg pre have "edges g' = (\\<^bsub>a\set (replace f [f21] (faces g))\<^esub> edges a) \ edges (f12)" by(auto simp: splitFace_def split_def edges_graph_def) with pre vs_neq show ?thesis apply (simp add: UNION_eq) apply (rule equalityI) apply simp apply (rule conjI) apply (rule subsetI) apply simp apply (erule bexE) apply (drule replace5) apply (case_tac "xa \ \ g") apply simp apply (subgoal_tac "x \ edges g") apply simp apply (simp add: edges_graph_def) apply force apply simp apply (subgoal_tac "pre_split_face f ram1 ram2 vs") apply (case_tac "between (vertices f) ram2 ram1 = []") apply (frule split_face_edges_f21_bet) apply (rule split) apply simp apply simp apply (case_tac "between (vertices f) ram1 ram2 = []") apply (frule split_face_edges_f_vs1_vs2) apply simp apply simp apply simp apply force apply (frule split_face_edges_f_vs2) apply simp apply simp apply simp apply force apply (frule split_face_edges_f21) apply (rule split) apply simp apply simp apply simp apply (case_tac "between (vertices f) ram1 ram2 = []") apply (frule split_face_edges_f_vs1) apply simp apply simp apply simp apply simp apply force apply (frule split_face_edges_f) apply simp apply simp apply simp apply simp apply force apply simp apply (subgoal_tac "pre_split_face f ram1 ram2 vs") apply (case_tac "between (vertices f) ram1 ram2 = []") apply (frule split_face_edges_f12_bet) apply (rule split) apply simp apply simp apply (case_tac "between (vertices f) ram2 ram1 = []") apply (frule split_face_edges_f_vs1_vs2) apply simp apply simp apply simp apply force apply (frule split_face_edges_f_vs1) apply simp apply simp apply simp apply force apply (frule split_face_edges_f12) apply (rule split) apply simp apply simp apply simp apply (case_tac "between (vertices f) ram2 ram1 = []") apply (frule split_face_edges_f_vs2) apply simp apply simp apply simp apply simp apply force apply (frule split_face_edges_f) apply simp apply simp apply simp apply simp apply force apply simp apply simp apply (subgoal_tac "pre_split_face f ram1 ram2 vs") apply (subgoal_tac "(ram2, last vs) \ edges f12 \ (hd vs, ram1) \ edges f12") apply (rule conjI) apply simp apply (rule conjI) apply simp apply (subgoal_tac "(ram1, hd vs) \ edges f21 \ (last vs, ram2) \ edges f21") apply (rule conjI) apply (rule disjI1) apply (rule bexI) apply (elim conjE) apply simp apply (rule replace3) apply(erule pre_splitFace_oldF) apply simp apply (rule conjI) apply (rule disjI1) apply (rule bexI) apply (elim conjE) apply simp apply (rule replace3) apply(erule pre_splitFace_oldF) apply simp apply (subgoal_tac "edges f \ {y. \x\set (replace f [f21] (faces g)). y \ edges x} \ edges f12") apply (subgoal_tac "edges g \ {y. \x\set (replace f [f21] (faces g)). y \ edges x} \ edges f12") apply (rule conjI) apply simp apply (rule conjI) apply simp apply (subgoal_tac "Edges(rev vs) \ edges f12") apply (rule conjI) prefer 2 apply blast apply (subgoal_tac "Edges vs \ edges f21") apply (subgoal_tac "Edges vs \ {y. \x\set (replace f [f21] (faces g)). y \ edges x}") apply blast apply (rule subset_trans) apply assumption apply (rule subsetI) apply simp apply (rule bexI) apply simp apply (rule replace3) apply(erule pre_splitFace_oldF) apply simp (* jetzt alle 7 subgoals aufloesen *) apply (frule split_face_edges_f21_subset) apply (rule split) apply simp apply simp apply (frule split_face_edges_f12_subset) apply (rule split) apply simp apply simp apply (simp add: edges_graph_def) apply (rule subsetI) apply simp apply (elim bexE) apply (case_tac "xa = f") apply simp apply blast apply (rule disjI1) apply (rule bexI) apply simp apply (rule replace4) apply simp apply force apply (rule subsetI) apply (subgoal_tac "\ u v. x = (u,v)") apply (elim exE conjE) apply (frule split_face_edges_or [OF split]) apply simp apply (case_tac "(u, v) \ edges f12") apply simp apply simp apply (rule bexI) apply (thin_tac "(u, v) \ edges f") apply assumption apply (rule replace3) apply(erule pre_splitFace_oldF) apply simp apply simp apply (frule split_face_edges_f21_subset) apply (rule split) apply simp apply simp apply (frule split_face_edges_f12_subset) apply (rule split) apply simp apply simp by simp qed lemma pre_splitFace_edges_f_in_g: "pre_splitFace g ram1 ram2 f vs \ edges f \ edges g" apply (simp add: edges_graph_def) by (force) lemma pre_splitFace_edges_f_in_g2: "pre_splitFace g ram1 ram2 f vs \ x \ edges f \ x \ edges g" apply (simp add: edges_graph_def) by (force) lemma splitFace_edges_g': "pre_splitFace g ram1 ram2 f vs \ (f12, f21, g') = splitFace g ram1 ram2 f vs \ vs \ [] \ edges g' = edges g \ Edges vs \ Edges(rev vs) \ {(ram2, last vs), (hd vs, ram1), (ram1, hd vs), (last vs, ram2)}" apply (subgoal_tac "edges g \ edges f = edges g") apply (frule splitFace_edges_g'_help) apply simp apply simp apply simp apply (frule pre_splitFace_edges_f_in_g) by blast lemma splitFace_edges_g'_vs: "pre_splitFace g ram1 ram2 f [] \ (f12, f21, g') = splitFace g ram1 ram2 f [] \ edges g' = edges g \ {(ram1, ram2), (ram2, ram1)}" proof - assume pre: "pre_splitFace g ram1 ram2 f []" and fdg: "(f12, f21, g') = splitFace g ram1 ram2 f []" from pre fdg have split: "(f12, f21) = split_face f ram1 ram2 []" apply (unfold pre_splitFace_def) apply (elim conjE) by (simp add: splitFace_split_face) from fdg pre have "edges g' = (\\<^bsub>a\set (replace f [f21] (faces g))\<^esub> edges a) \ edges (f12)" by (auto simp: splitFace_def split_def edges_graph_def) with pre show ?thesis apply (simp add: UNION_eq) apply (rule equalityI) apply simp apply (rule conjI) apply (rule subsetI) apply simp apply (erule bexE) apply (drule replace5) apply (case_tac "xa \ \ g") apply simp apply (subgoal_tac "x \ edges g") apply simp apply (simp add: edges_graph_def) apply force apply simp apply (subgoal_tac "pre_split_face f ram1 ram2 []") apply (case_tac "between (vertices f) ram2 ram1 = []") apply (simp add: pre_FaceDiv_between2) apply (frule split_face_edges_f21_vs) apply (rule split) apply simp apply simp apply simp apply (case_tac "x = (ram1, ram2)") apply simp apply simp apply (rule disjI2) apply (rule pre_splitFace_edges_f_in_g2) apply simp apply (subgoal_tac "pre_split_face f ram1 ram2 []") apply (frule split_face_edges_f) apply simp apply simp apply (rule pre_FaceDiv_between1) apply simp apply simp apply simp apply force apply simp apply simp apply (rule subsetI) apply simp apply (subgoal_tac "pre_split_face f ram1 ram2 []") apply (case_tac "between (vertices f) ram1 ram2 = []") apply (simp add: pre_FaceDiv_between1) apply (frule split_face_edges_f12_vs) apply (rule split) apply simp apply simp apply simp apply (case_tac "x = (ram2, ram1)") apply simp apply simp apply (rule disjI2) apply (rule pre_splitFace_edges_f_in_g2) apply simp apply (subgoal_tac "pre_split_face f ram1 ram2 []") apply (frule split_face_edges_f) apply simp apply simp apply simp apply (rule pre_FaceDiv_between2) apply simp apply simp apply force apply simp apply simp apply simp apply (subgoal_tac "pre_split_face f ram1 ram2 []") apply (subgoal_tac "(ram1, ram2) \ edges f21") apply (rule conjI) apply (rule disjI1) apply (rule bexI) apply simp apply (force) apply (subgoal_tac "(ram2, ram1) \ edges f12") apply (rule conjI) apply force apply (rule subsetI) apply (simp add: edges_graph_def) apply (elim bexE) apply (case_tac "xa = f") apply simp apply (subgoal_tac "\ u v. x = (u,v)") apply (elim exE conjE) apply (subgoal_tac "pre_split_face f ram1 ram2 []") apply (frule split_face_edges_or [OF split]) apply simp apply (case_tac "(u, v) \ edges f12") apply simp apply simp apply force apply simp apply simp apply (rule disjI1) apply (rule bexI) apply simp apply (rule replace4) apply simp apply force apply (frule split_face_edges_f12_vs) apply simp apply (rule split) apply simp apply (rule pre_FaceDiv_between1) apply simp apply simp apply (frule split_face_edges_f21_vs) apply simp apply (rule split) apply simp apply (rule pre_FaceDiv_between2) apply simp apply simp by simp qed lemma splitFace_edges_incr: "pre_splitFace g ram1 ram2 f vs \ (f\<^sub>1, f\<^sub>2, g') = splitFace g ram1 ram2 f vs \ edges g \ edges g'" apply(cases vs) apply(simp add:splitFace_edges_g'_vs) apply blast apply(simp add:splitFace_edges_g') apply blast done lemma snd_snd_splitFace_edges_incr: "pre_splitFace g v\<^sub>1 v\<^sub>2 f vs \ edges g \ edges(snd(snd(splitFace g v\<^sub>1 v\<^sub>2 f vs)))" apply(erule splitFace_edges_incr [where f\<^sub>1 = "fst(splitFace g v\<^sub>1 v\<^sub>2 f vs)" and f\<^sub>2 = "fst(snd(splitFace g v\<^sub>1 v\<^sub>2 f vs))"]) apply(auto) done (********************* Vorbereitung für subdivFace *****************) (**** computes the list of ram vertices **********) subsection \\removeNones\\ definition removeNones :: "'a option list \ 'a list" where "removeNones vOptionList \ [the x. x \ vOptionList, x \ None]" (* "removeNones vOptionList \ map the [x \ vOptionList. x \ None]" *) declare removeNones_def [simp] lemma removeNones_inI[intro]: "Some a \ set ls \ a \ set (removeNones ls)" by (induct ls) auto lemma removeNones_hd[simp]: "removeNones ( Some a # ls) = a # removeNones ls" by auto lemma removeNones_last[simp]: "removeNones (ls @ [Some a]) = removeNones ls @ [a]" by auto lemma removeNones_in[simp]: "removeNones (as @ Some a # bs) = removeNones as @ a # removeNones bs" by auto lemma removeNones_none_hd[simp]: "removeNones ( None # ls) = removeNones ls" by auto lemma removeNones_none_last[simp]: "removeNones (ls @ [None]) = removeNones ls" by auto lemma removeNones_none_in[simp]: "removeNones (as @ None # bs) = removeNones (as @ bs)" by auto lemma removeNones_empty[simp]: "removeNones [] = []" by auto declare removeNones_def [simp del] (************* natToVertexList ************) subsection\\natToVertexList\\ (* Hilfskonstrukt natToVertexList *) primrec natToVertexListRec :: "nat \ vertex \ face \ nat list \ vertex option list" where "natToVertexListRec old v f [] = []" | "natToVertexListRec old v f (i#is) = (if i = old then None#natToVertexListRec i v f is else Some (f\<^bsup>i\<^esup> \ v) # natToVertexListRec i v f is)" primrec natToVertexList :: "vertex \ face \ nat list \ vertex option list" where "natToVertexList v f [] = []" | "natToVertexList v f (i#is) = (if i = 0 then (Some v)#(natToVertexListRec i v f is) else [])" subsection \@{const indexToVertexList}\ lemma nextVertex_inj: "distinct (vertices f) \ v \ \ f \ i < length (vertices (f::face)) \ a < length (vertices f) \ f\<^bsup>a\<^esup>\v = f\<^bsup>i\<^esup>\v \ i = a" proof - assume d: "distinct (vertices f)" and v: "v \ \ f" and i: "i < length (vertices (f::face))" and a: "a < length (vertices f)" and eq:" f\<^bsup>a\<^esup>\v = f\<^bsup>i\<^esup>\v" then have eq: "(verticesFrom f v)!a = (verticesFrom f v)!i " by (simp add: verticesFrom_nth) define xs where "xs = verticesFrom f v" with eq have eq: "xs!a = xs!i" by auto from d v have z: "distinct (verticesFrom f v)" by auto moreover from xs_def a v d have "(verticesFrom f v) = take a xs @ xs ! a # drop (Suc a) xs" by (auto intro: id_take_nth_drop simp: verticesFrom_length) with eq have "(verticesFrom f v) = take a xs @ xs ! i # drop (Suc a) xs" by simp moreover from xs_def i v d have "(verticesFrom f v) = take i xs @ xs ! i # drop (Suc i) xs" by (auto intro: id_take_nth_drop simp: verticesFrom_length) ultimately have "take a xs = take i xs" by (rule dist_at1) moreover from v d have vertFrom[simp]: "length (vertices f) = length (verticesFrom f v)" by (auto simp: verticesFrom_length) from xs_def a i have "a < length xs" "i < length xs" by auto moreover have "\ a i. a < length xs \ i < length xs \ take a xs = take i xs \ a = i" proof (induct xs) case Nil then show ?case by auto next case (Cons x xs) then show ?case apply (cases a) apply auto apply (cases i) apply auto apply (cases i) by auto qed ultimately show ?thesis by simp qed lemma a: "distinct (vertices f) \ v \ \ f \ (\i \ set is. i < length (vertices f)) \ (\a. a < length (vertices f) \ hideDupsRec ((f \ ^^ a) v) [(f \ ^^ k) v. k \ is] = natToVertexListRec a v f is)" proof (induct "is") case Nil then show ?case by simp next case (Cons i "is") then show ?case by (auto simp: nextVertices_def intro: nextVertex_inj) qed lemma indexToVertexList_natToVertexList_eq: "distinct (vertices f) \ v \ \ f \ (\i \ set is. i < length (vertices f)) \ is \ [] \ hd is = 0 \ indexToVertexList f v is = natToVertexList v f is" apply (cases "is") by (auto simp: a [where a = 0, simplified] indexToVertexList_def nextVertices_def) (********************* Einschub Ende ***************************************) (******** natToVertexListRec ************) lemma nvlr_length: "\ old. (length (natToVertexListRec old v f ls)) = length ls" apply (induct ls) by auto lemma nvl_length[simp]: "hd e = 0 \ length (natToVertexList v f e) = length e" apply (cases "e") by (auto intro: nvlr_length) lemma natToVertexListRec_length[simp]: "\ e f. length (natToVertexListRec e v f es) = length es" by (induct es) auto lemma natToVertexList_length[simp]: "incrIndexList es (length es) (length (vertices f)) \ length (natToVertexList v f es) = length es" apply (case_tac es) by simp_all lemma natToVertexList_nth_Suc: "incrIndexList es (length es) (length (vertices f)) \ Suc n < length es \ (natToVertexList v f es)!(Suc n) = (if (es!n = es!(Suc n)) then None else Some (f\<^bsup>(es!Suc n)\<^esup> \ v))" proof - assume incr: "incrIndexList es (length es) (length (vertices f))" and n: "Suc n < length es" have rec: "\ old n. Suc n < length es \ (natToVertexListRec old v f es)!(Suc n) = (if (es!n = es!(Suc n)) then None else Some (f\<^bsup>(es!Suc n)\<^esup> \ v))" proof (induct es) case Nil then show ?case by auto next case (Cons e es) note cons1 = this then show ?case proof (cases es) case Nil with cons1 show ?thesis by simp next case (Cons e' es') with cons1 show ?thesis proof (cases n) case 0 with Cons cons1 show ?thesis by simp next case (Suc m) with Cons cons1 have "\ old. natToVertexListRec old v f es ! Suc m = (if es ! m = es ! Suc m then None else Some (f\<^bsup>es ! Suc m\<^esup> \ v))" by (rule_tac cons1) auto then show ?thesis apply (cases "e = old") by (simp_all add: Suc) qed qed qed with n have "natToVertexListRec 0 v f es ! Suc n = (if es ! n = es ! Suc n then None else Some (f\<^bsup>es ! Suc n\<^esup> \ v))" by (rule_tac rec) auto with incr show ?thesis by (cases es) auto qed lemma natToVertexList_nth_0: "incrIndexList es (length es) (length (vertices f)) \ 0 < length es \ (natToVertexList v f es)!0 = Some (f\<^bsup>(es!0)\<^esup> \ v)" apply (cases es) apply (simp_all add: nextVertices_def) by (subgoal_tac "a = 0") auto lemma natToVertexList_hd[simp]: "incrIndexList es (length es) (length (vertices f)) \ hd (natToVertexList v f es) = Some v" apply (cases es) by (simp_all add: nextVertices_def) lemma nth_last[intro]: "Suc i = length xs \ xs!i = last xs" by (cases xs rule: rev_exhaust) auto declare incrIndexList_help4 [simp del] lemma natToVertexList_last[simp]: "distinct (vertices f) \ v \ \ f \ incrIndexList es (length es) (length (vertices f)) \ last (natToVertexList v f es) = Some (last (verticesFrom f v))" proof - assume vors: "distinct (vertices f)" "v \ \ f" and incr: "incrIndexList es (length es) (length (vertices f))" define n' where "n' = length es - 2" from incr have "1 < length es" by auto with n'_def have n'l: "Suc (Suc n') = length es" by arith from incr n'l have last_ntvl: "(natToVertexList v f es)!(Suc n') = last (natToVertexList v f es)" by auto from n'l have last_es: "es!(Suc n') = last es" by auto from n'l have "es!n' = last (butlast es)" apply (cases es rule: rev_exhaust) by (auto simp: nth_append) with last_es incr have less: "es!n' < es!(Suc n')" by auto from n'l have "Suc n' < length es" by arith with incr less have "(natToVertexList v f es)!(Suc n') = (Some (f\<^bsup>(es!Suc n')\<^esup> \ v))" by (auto dest: natToVertexList_nth_Suc) with incr last_ntvl last_es have rule1: "last (natToVertexList v f es) = Some (f\<^bsup>((length (vertices f)) - (Suc 0))\<^esup> \ v)" by auto from incr have lvf: "1 < length (vertices f)" by auto with vors have rule2: "verticesFrom f v ! ((length (vertices f)) - (Suc 0)) = f\<^bsup>((length (vertices f)) - (Suc 0))\<^esup> \ v" by (auto intro!: verticesFrom_nth) from vors lvf have "verticesFrom f v ! ((length (vertices f)) - (Suc 0)) = last (verticesFrom f v)" apply (rule_tac nth_last) by (auto simp: verticesFrom_length) with rule1 rule2 show ?thesis by auto qed lemma indexToVertexList_last[simp]: "distinct (vertices f) \ v \ \ f \ incrIndexList es (length es) (length (vertices f)) \ last (indexToVertexList f v es) = Some (last (verticesFrom f v))" apply (subgoal_tac "indexToVertexList f v es = natToVertexList v f es") apply simp apply (rule indexToVertexList_natToVertexList_eq) by auto lemma nths_take: "\ n iset. \ i \ iset. i < n \ nths (take n xs) iset = nths xs iset" proof (induct xs) case Nil then show ?case by simp next case (Cons x xs) then show ?case apply (simp add: nths_Cons) apply (cases n) apply simp apply (simp add: nths_Cons) apply (rule Cons) by auto qed lemma nths_reduceIndices: "\ iset. nths xs iset = nths xs {i. i < length xs \ i \ iset}" proof (induct xs) case Nil then show ?case by simp next case (Cons x xs) then have "nths xs {j. Suc j \ iset} = nths xs {i. i < length xs \ i \ {j. Suc j \ iset}}" by (rule_tac Cons) then show ?case by (simp add: nths_Cons) qed lemma natToVertexList_nths1: "distinct (vertices f) \ v \ \ f \ vs = verticesFrom f v \ incrIndexList es (length es) (length vs) \ n \ length es \ nths (take (Suc (es!(n - 1))) vs) (set (take n es)) = removeNones (take n (natToVertexList v f es))" proof (induct n) case 0 then show ?case by simp next case (Suc n) then have "nths (take (Suc (es ! (n - Suc 0))) (verticesFrom f v)) (set (take n es)) = removeNones (take n (natToVertexList v f es))" "distinct (vertices f)" "v \ \ f" "vs = verticesFrom f v" "incrIndexList es (length es) (length (verticesFrom f v))" "Suc n \ length es" by auto (* does this improve the performance? *) note suc1 = this then have lvs: "length vs = length (vertices f)" by (auto intro: verticesFrom_length) with suc1 have vsne: "vs \ []" by auto with suc1 show ?case proof (cases "natToVertexList v f es ! n") case None then show ?thesis proof (cases n) case 0 with None suc1 lvs show ?thesis by (simp add: take_Suc_conv_app_nth natToVertexList_nth_0) next case (Suc n') with None suc1 lvs have esn: "es!n = es!n'" by (simp add: natToVertexList_nth_Suc split: if_split_asm) from Suc have n': "n - Suc 0 = n'" by auto show ?thesis proof (cases "Suc n = length es") case True then have small_n: "n < length es" by auto from True have "take (Suc n) es = es" by auto with small_n have "take n es @ [es!n] = es" by (simp add: take_Suc_conv_app_nth) then have esn_simps: "take n es = butlast es \ es!n = last es" by (cases es rule: rev_exhaust) auto from True Suc have n'l: "Suc n' = length (butlast es)" by auto then have small_n': "n' < length (butlast es)" by auto from Suc small_n have take_n': "take (Suc n') (butlast es @ [last es]) = take (Suc n') (butlast es)" by auto from small_n have es_exh: "es = butlast es @ [last es]" by (cases es rule: rev_exhaust) auto from n'l have "take (Suc n') (butlast es @ [last es]) = butlast es" by auto with es_exh have "take (Suc n') es = butlast es" by auto with small_n Suc have "take n' es @ [es!n'] = (butlast es)" by (simp add: take_Suc_conv_app_nth) with small_n' have esn'_simps: "take n' es = butlast (butlast es) \ es!n' = last (butlast es)" by (cases "butlast es" rule: rev_exhaust) auto from suc1 have "last (butlast es) < last es" by auto with esn esn_simps esn'_simps have False by auto (* have "last (butlast es) = last es" by auto *) then show ?thesis by auto next case False with suc1 have le: "Suc n < length es" by auto from suc1 le have "es = take (Suc n) es @ es!(Suc n) # drop (Suc (Suc n)) es" by (auto intro: id_take_nth_drop) with suc1 have "increasing (take (Suc n) es @ es!(Suc n) # drop (Suc (Suc n)) es)" by auto then have "\ i \ (set (take (Suc n) es)). i \ es ! (Suc n)" by (auto intro: increasing2) with suc1 have "\ i \ (set (take n es)). i \ es ! (Suc n)" by (simp add: take_Suc_conv_app_nth) then have seq: "nths (take (Suc (es ! Suc n)) (verticesFrom f v)) (set (take n es)) = nths (verticesFrom f v) (set (take n es))" apply (rule_tac nths_take) by auto from suc1 have "es = take n es @ es!n # drop (Suc n) es" by (auto intro: id_take_nth_drop) with suc1 have "increasing (take n es @ es!n # drop (Suc n) es)" by auto then have "\ i \ (set (take n es)). i \ es ! n" by (auto intro: increasing2) with suc1 esn have "\ i \ (set (take n es)). i \ es ! n'" by (simp add: take_Suc_conv_app_nth) with Suc have seq2: "nths (take (Suc (es ! n')) (verticesFrom f v)) (set (take n es)) = nths (verticesFrom f v) (set (take n es))" apply (rule_tac nths_take) by auto from Suc suc1 have "(insert (es ! n') (set (take n es))) = set (take n es)" apply auto by (simp add: take_Suc_conv_app_nth) with esn None suc1 seq seq2 n' show ?thesis by (simp add: take_Suc_conv_app_nth) qed qed next case (Some v') then show ?thesis proof (cases n) case 0 from suc1 lvs have "verticesFrom f v \ []" by auto then have "verticesFrom f v = hd (verticesFrom f v) # tl (verticesFrom f v)" by auto then have "verticesFrom f v = v # tl (verticesFrom f v)" by (simp add: verticesFrom_hd) then obtain z where "verticesFrom f v = v # z" by auto then have sub: "nths (verticesFrom f v) {0} = [v]" by (auto simp: nths_Cons) from 0 suc1 have "es!0 = 0" by (cases es) auto with 0 Some suc1 lvs sub vsne show ?thesis by (simp add: take_Suc_conv_app_nth natToVertexList_nth_0 nextVertices_def take_Suc nths_Cons verticesFrom_hd del:verticesFrom_empty) next case (Suc n') with Some suc1 lvs have esn: "es!n \ es!n'" by (simp add: natToVertexList_nth_Suc split: if_split_asm) from suc1 Suc have "Suc n' < length es" by auto with suc1 lvs esn have "natToVertexList v f es !(Suc n') = Some (f\<^bsup>(es!(Suc n'))\<^esup> \ v)" apply (simp add: natToVertexList_nth_Suc) by (simp add: Suc) with Suc have "natToVertexList v f es ! n = Some (f\<^bsup>(es!n)\<^esup> \ v)" by auto with Some have v': "v' = f\<^bsup>(es!n)\<^esup> \ v" by simp from Suc have n': "n - Suc 0 = n'" by auto from suc1 Suc have "es = take (Suc n') es @ es!n # drop (Suc n) es" by (auto intro: id_take_nth_drop) with suc1 have "increasing (take (Suc n') es @ es!n # drop (Suc n) es)" by auto with suc1 Suc have "es!n' \ es!n" apply (auto intro!: increasing2) by (auto simp: take_Suc_conv_app_nth) with esn have smaller_n: "es!n' < es!n" by auto from suc1 lvs have smaller: "(es!n) < length vs" by auto from suc1 smaller lvs have "(verticesFrom f v)!(es!n) = f\<^bsup>(es!n)\<^esup> \ v" by (auto intro: verticesFrom_nth) with v' have "(verticesFrom f v)!(es!n) = v'" by auto then have sub1: "nths ([((verticesFrom f v)!(es!n))]) {j. j + (es!n) : (insert (es ! n) (set (take n es)))} = [v']" by auto from suc1 smaller lvs have len: "length (take (es ! n) (verticesFrom f v)) = es!n" by auto have "\x. x \ (set (take n es)) \ x < (es ! n)" proof - fix x assume x: "x \ set (take n es)" from suc1 Suc have "es = take n' es @ es!n' # drop (Suc n') es" by (auto intro: id_take_nth_drop) with suc1 have "increasing (take n' es @ es!n' # drop (Suc n') es)" by auto then have "\ x. x \ set (take n' es) \ x \ es!n'" by (auto intro!: increasing2) with x Suc suc1 have "x \ es!n'" by (auto simp: take_Suc_conv_app_nth) with smaller_n show "x < es!n" by auto qed then have "{i. i < es ! n \ i \ set (take n es)} = (set (take n es))" by auto then have elim_insert: "{i. i < es ! n \ i \ insert (es ! n) (set (take n es))} = (set (take n es))" by auto have "nths (take (es ! n) (verticesFrom f v)) (insert (es ! n) (set (take n es))) = nths (take (es ! n) (verticesFrom f v)) {i. i < length (take (es ! n) (verticesFrom f v)) \ i \ (insert (es ! n) (set (take n es)))}" by (rule nths_reduceIndices) with len have "nths (take (es ! n) (verticesFrom f v)) (insert (es ! n) (set (take n es))) = nths (take (es ! n) (verticesFrom f v)) {i. i < (es ! n) \ i \ (insert (es ! n) (set (take n es)))}" by simp with elim_insert have sub2: "nths (take (es ! n) (verticesFrom f v)) (insert (es ! n) (set (take n es))) = nths (take (es ! n) (verticesFrom f v)) (set (take n es))" by simp define m where "m = es!n - es!n'" with smaller_n have mgz: "0 < m" by auto with m_def have esn: "es!n = (es!n') + m" by auto have helper: "\x. x \ (set (take n es)) \ x \ (es ! n')" proof - fix x assume x: "x \ set (take n es)" from suc1 Suc have "es = take n' es @ es!n' # drop (Suc n') es" by (auto intro: id_take_nth_drop) with suc1 have "increasing (take n' es @ es!n' # drop (Suc n') es)" by auto then have "\ x. x \ set (take n' es) \ x \ es!n'" by (auto intro!: increasing2) with x Suc suc1 show "x \ es!n'" by (auto simp: take_Suc_conv_app_nth) qed define m' where "m' = m - 1" define Suc_es_n' where "Suc_es_n' = Suc (es!n')" from smaller smaller_n have "Suc (es!n') < length vs" by auto then have "min (length vs) (Suc (es ! n')) = Suc (es!n')" by arith with Suc_es_n'_def have empty: "{j. j + length (take Suc_es_n' vs) \ set (take n es)} = {}" apply auto apply (frule helper) by arith from Suc_es_n'_def mgz esn m'_def have esn': "es!n = Suc_es_n' + m'" by auto with smaller have "(take (Suc_es_n' + m') vs) = take (Suc_es_n') vs @ take m' (drop (Suc_es_n') vs)" by (auto intro: take_add) with esn' have "nths (take (es ! n) vs) (set (take n es)) = nths (take (Suc_es_n') vs @ take m' (drop (Suc_es_n') vs)) (set (take n es))" by auto then have "nths (take (es ! n) vs) (set (take n es)) = nths (take (Suc_es_n') vs) (set (take n es)) @ nths (take m' (drop (Suc_es_n') vs)) {j. j + length (take (Suc_es_n') vs) : (set (take n es))}" by (simp add: nths_append) with empty Suc_es_n'_def have "nths (take (es ! n) vs) (set (take n es)) = nths (take (Suc (es!n')) vs) (set (take n es))" by simp with suc1 sub2 have sub3: "nths (take (es ! n) (verticesFrom f v)) (insert (es ! n) (set (take n es))) = nths (take (Suc (es!n')) (verticesFrom f v)) (set (take n es))" by simp from smaller suc1 have "take (Suc (es ! n)) (verticesFrom f v) = take (es ! n) (verticesFrom f v) @ [((verticesFrom f v)!(es!n))]" by (auto simp: take_Suc_conv_app_nth) with suc1 smaller have "nths (take (Suc (es ! n)) (verticesFrom f v)) (insert (es ! n) (set (take n es))) = nths (take (es ! n) (verticesFrom f v)) (insert (es ! n) (set (take n es))) @ nths ([((verticesFrom f v)!(es!n))]) {j. j + (es!n) : (insert (es ! n) (set (take n es)))}" by (auto simp: nths_append) with sub1 sub3 have "nths (take (Suc (es ! n)) (verticesFrom f v)) (insert (es ! n) (set (take n es))) = nths (take (Suc (es ! n')) (verticesFrom f v)) (set (take n es)) @ [v']" by auto with Some suc1 lvs n' show ?thesis by (simp add: take_Suc_conv_app_nth) qed qed qed lemma natToVertexList_nths: "distinct (vertices f) \ v \ \ f \ incrIndexList es (length es) (length (vertices f)) \ nths (verticesFrom f v) (set es) = removeNones (natToVertexList v f es)" proof - assume vors1: "distinct (vertices f)" "v \ \ f" "incrIndexList es (length es) (length (vertices f))" define vs where "vs = verticesFrom f v" with vors1 have lvs: "length vs = length (vertices f)" by (auto intro: verticesFrom_length) with vors1 vs_def have vors: "distinct (vertices f)" "v \ \ f" "vs = verticesFrom f v" "incrIndexList es (length es) (length vs)" by auto with lvs have vsne: "vs \ []" by auto define n where "n = length es" then have "es!(n - 1) = last es" proof (cases n) case 0 with n_def vors show ?thesis by (cases es) auto next case (Suc n') with n_def have small_n': "n' < length es" by arith from Suc n_def have "take (Suc n') es = es" by auto with small_n' have "take n' es @ [es!n'] = es" by (simp add: take_Suc_conv_app_nth) then have "es!n' = last es" by (cases es rule: rev_exhaust) auto with Suc show ?thesis by auto qed with vors have "es!(n - 1) = (length vs) - 1" by auto with vsne have "Suc (es! (n - 1)) = (length vs)" by auto then have take_vs: "take (Suc (es!(n - 1))) vs = vs" by auto from n_def vors have "n = length (natToVertexList v f es)" by auto then have take_nTVL: "take n (natToVertexList v f es) = natToVertexList v f es" by auto from n_def have take_es: "take n es = es" by auto from n_def have "n \ length es" by auto with vors have "nths (take (Suc (es!(n - 1))) vs) (set (take n es)) = removeNones (take n (natToVertexList v f es))" by (rule natToVertexList_nths1) with take_vs take_nTVL take_es vs_def show ?thesis by simp qed lemma filter_Cons2: "x \ set ys \ [y\ys. y = x \ P y] = [y\ys. P y]" by (induct ys) (auto) lemma natToVertexList_removeNones: "distinct (vertices f) \ v \ \ f \ incrIndexList es (length es) (length (vertices f)) \ [x\verticesFrom f v. x \ set (removeNones (natToVertexList v f es))] = removeNones (natToVertexList v f es)" proof - assume vors: "distinct (vertices f)" "v \ \ f" "incrIndexList es (length es) (length (vertices f))" then have dist: "distinct (verticesFrom f v)" by auto from vors have sub_eq: "nths (verticesFrom f v) (set es) = removeNones (natToVertexList v f es)" by (rule natToVertexList_nths) from dist have "[x \ verticesFrom f v. x \ set (nths (verticesFrom f v) (set es))] = removeNones (natToVertexList v f es)" apply (simp add: filter_in_nths) by (simp add: sub_eq) with sub_eq show ?thesis by simp qed (**************************** invalidVertexList ************) definition is_duplicateEdge :: "graph \ face \ vertex \ vertex \ bool" where "is_duplicateEdge g f a b \ ((a, b) \ edges g \ (a, b) \ edges f \ (b, a) \ edges f) \ ((b, a) \ edges g \ (b, a) \ edges f \ (a, b) \ edges f)" definition invalidVertexList :: "graph \ face \ vertex option list \ bool" where "invalidVertexList g f vs \ \i < |vs|- 1. case vs!i of None \ False | Some a \ case vs!(i+1) of None \ False | Some b \ is_duplicateEdge g f a b" (**************************** subdivFace **********************) subsection \\pre_subdivFace(')\\ definition pre_subdivFace_face :: "face \ vertex \ vertex option list \ bool" where "pre_subdivFace_face f v' vOptionList \ [v \ verticesFrom f v'. v \ set (removeNones vOptionList)] = (removeNones vOptionList) \ \ final f \ distinct (vertices f) \ hd (vOptionList) = Some v' \ v' \ \ f \ last (vOptionList) = Some (last (verticesFrom f v')) \ hd (tl (vOptionList)) \ last (vOptionList) \ 2 < | vOptionList | \ vOptionList \ [] \ tl (vOptionList) \ []" definition pre_subdivFace :: "graph \ face \ vertex \ vertex option list \ bool" where "pre_subdivFace g f v' vOptionList \ pre_subdivFace_face f v' vOptionList \ \ invalidVertexList g f vOptionList" (* zu teilende Fläche, ursprüngliches v, erster Ram-Punkt, Anzahl der überlaufenen NOnes, rest der vol *) definition pre_subdivFace' :: "graph \ face \ vertex \ vertex \ nat \ vertex option list \ bool" where "pre_subdivFace' g f v' ram1 n vOptionList \ \ final f \ v' \ \ f \ ram1 \ \ f \ v' \ set (removeNones vOptionList) \ distinct (vertices f) \ ( [v \ verticesFrom f v'. v \ set (removeNones vOptionList)] = (removeNones vOptionList) \ before (verticesFrom f v') ram1 (hd (removeNones vOptionList)) \ last (vOptionList) = Some (last (verticesFrom f v')) \ vOptionList \ [] \ ((v' = ram1 \ (0 < n)) \ ((v' = ram1 \ (hd (vOptionList) \ Some (last (verticesFrom f v')))) \ (v' \ ram1))) \ \ invalidVertexList g f vOptionList \ (n = 0 \ hd (vOptionList) \ None \ \ is_duplicateEdge g f ram1 (the (hd (vOptionList)))) \ (vOptionList = [] \ v' \ ram1) )" lemma pre_subdivFace_face_in_f[intro]: "pre_subdivFace_face f v ls \ Some a \ set ls \ a \ set (verticesFrom f v)" apply (subgoal_tac "a \ set (removeNones ls)") apply (auto simp: pre_subdivFace_face_def) apply (subgoal_tac "a \ set [v\verticesFrom f v . v \ set (removeNones ls)]") apply (thin_tac "[v\verticesFrom f v . v \ set (removeNones ls)] = removeNones ls") by auto lemma pre_subdivFace_in_f[intro]: "pre_subdivFace g f v ls \ Some a \ set ls \ a \ set (verticesFrom f v)" by (auto simp: pre_subdivFace_def) lemma pre_subdivFace_face_in_f'[intro]: "pre_subdivFace_face f v ls \ Some a \ set ls \ a \ \ f" apply (cases "a = v") apply (force simp: pre_subdivFace_face_def) apply (rule verticesFrom_in') apply (rule pre_subdivFace_face_in_f) by auto lemma filter_congs_shorten1: "distinct (verticesFrom f v) \ [v\verticesFrom f v . v = a \ v \ set vs] = (a # vs) \ [v\verticesFrom f v . v \ set vs] = vs" proof - assume dist: "distinct (verticesFrom f v)" and eq: "[v\verticesFrom f v . v = a \ v \ set vs] = (a # vs)" have rule1: "\ vs a ys. distinct vs \ [v\vs . v = a \ v \ set ys] = a # ys \ [v\vs. v \ set ys] = ys" proof - fix vs a ys assume dist: "distinct vs" and ays: "[v\vs . v = a \ v \ set ys] = a # ys" then have "distinct ([v\vs . v = a \ v \ set ys])" by (rule_tac distinct_filter) with ays have distys: "distinct (a # ys)" by simp from dist distys ays show "[v\vs. v \ set ys] = ys" apply (induct vs) by (auto split: if_split_asm simp: filter_Cons2) qed from dist eq show ?thesis by (rule_tac rule1) qed lemma ovl_shorten: "distinct (verticesFrom f v) \ [v\verticesFrom f v . v \ set (removeNones (va # vol))] = (removeNones (va # vol)) \ [v\verticesFrom f v . v \ set (removeNones (vol))] = (removeNones (vol))" proof - assume dist: "distinct (verticesFrom f v)" and vors: "[v\verticesFrom f v . v \ set (removeNones (va # vol))] = (removeNones (va # vol))" then show ?thesis proof (cases va) case None with vors Cons show ?thesis by auto next case (Some a) with vors dist show ?thesis by (auto intro!: filter_congs_shorten1) qed qed lemma pre_subdivFace_face_distinct: "pre_subdivFace_face f v vol \ distinct (removeNones vol)" apply (auto dest!: verticesFrom_distinct simp: pre_subdivFace_face_def) apply (subgoal_tac "distinct ([v\verticesFrom f v . v \ set (removeNones vol)])") apply simp apply (thin_tac "[v\verticesFrom f v . v \ set (removeNones vol)] = removeNones vol") by auto lemma invalidVertexList_shorten: "invalidVertexList g f vol \ invalidVertexList g f (v # vol)" apply (simp add: invalidVertexList_def) apply auto apply (rule exI) apply safe apply (subgoal_tac "(Suc i) < | vol |") apply assumption apply arith apply auto apply (case_tac "vol!i") by auto lemma pre_subdivFace_pre_subdivFace': "v \ \ f \ pre_subdivFace g f v (vo # vol) \ pre_subdivFace' g f v v 0 (vol)" proof - assume vors: "v \ \ f" "pre_subdivFace g f v (vo # vol)" then have vors': "v \ \ f" "pre_subdivFace_face f v (vo # vol)" "\ invalidVertexList g f (vo # vol)" by (auto simp: pre_subdivFace_def) then have r: "removeNones vol \ []" apply (cases "vol" rule: rev_exhaust) by (auto simp: pre_subdivFace_face_def) then have "Some (hd (removeNones vol)) \ set vol" apply (induct vol) apply auto apply (case_tac a) by auto then have "Some (hd (removeNones vol)) \ set (vo # vol)" by auto with vors' have hd: "hd (removeNones vol) \ \ f" by (rule_tac pre_subdivFace_face_in_f') from vors' have "Some v = vo" by (auto simp: pre_subdivFace_face_def) with vors' have "v \ set (tl (removeNones (vo # vol)))" apply (drule_tac pre_subdivFace_face_distinct) by auto with vors' r have ne: "v \ hd (removeNones vol)" by (cases "removeNones vol") (auto simp: pre_subdivFace_face_def) from vors' have dist: "distinct (removeNones (vo # vol))" apply (rule_tac pre_subdivFace_face_distinct) . from vors' have invalid: "\ invalidVertexList g f vol" by (auto simp: invalidVertexList_shorten) from ne hd vors' invalid dist show ?thesis apply (unfold pre_subdivFace'_def) apply (simp add: pre_subdivFace'_def pre_subdivFace_face_def) apply safe apply (rule ovl_shorten) apply (simp add: pre_subdivFace_face_def) apply assumption apply (rule before_verticesFrom) apply simp+ apply (simp add: invalidVertexList_def) apply (erule allE) apply (erule impE) apply (subgoal_tac "0 < |vol|") apply (thin_tac "Suc 0 < | vol |") apply assumption apply simp apply (simp) apply (case_tac "vol") apply simp by (simp add: is_duplicateEdge_def) qed lemma pre_subdivFace'_distinct: "pre_subdivFace' g f v' v n vol \ distinct (removeNones vol)" apply (unfold pre_subdivFace'_def) apply (cases vol) apply simp+ apply (elim conjE) apply (drule_tac verticesFrom_distinct) apply assumption apply (subgoal_tac "distinct [v\verticesFrom f v' . v \ set (removeNones (a # list))]") apply force apply (thin_tac "[v\verticesFrom f v' . v \ set (removeNones (a # list))] = removeNones (a # list)") by auto lemma natToVertexList_pre_subdivFace_face: "\ final f \ distinct (vertices f) \ v \ \ f \ 2 < |es| \ incrIndexList es (length es) (length (vertices f)) \ pre_subdivFace_face f v (natToVertexList v f es)" proof - assume vors: "\ final f" "distinct (vertices f)" "v \ \ f" "2 < |es|" "incrIndexList es (length es) (length (vertices f))" then have lastOvl: "last (natToVertexList v f es) = Some (last (verticesFrom f v))" by auto from vors have nvl_l: "2 < | natToVertexList v f es |" by auto from vors have "distinct [x\verticesFrom f v . x \ set (removeNones (natToVertexList v f es))]" by auto with vors have "distinct (removeNones (natToVertexList v f es))" by (simp add: natToVertexList_removeNones) with nvl_l lastOvl have hd_last: "hd (tl (natToVertexList v f es)) \ last (natToVertexList v f es)" apply auto apply (cases "natToVertexList v f es") apply simp apply (case_tac "list" rule: rev_exhaust) apply simp apply (case_tac "ys") apply simp apply (case_tac "a") apply simp by simp from vors lastOvl hd_last nvl_l show ?thesis apply (auto intro: natToVertexList_removeNones simp: pre_subdivFace_face_def) apply (cases es) apply auto apply (cases es) apply auto apply (subgoal_tac "0 < length list") apply (case_tac list) by (auto split: if_split_asm) qed lemma indexToVertexList_pre_subdivFace_face: "\ final f \ distinct (vertices f) \ v \ \ f \ 2 < |es| \ incrIndexList es (length es) (length (vertices f)) \ pre_subdivFace_face f v (indexToVertexList f v es)" apply (subgoal_tac "indexToVertexList f v es = natToVertexList v f es") apply simp apply (rule natToVertexList_pre_subdivFace_face) apply assumption+ apply (rule indexToVertexList_natToVertexList_eq) by auto lemma subdivFace_subdivFace'_eq: "pre_subdivFace g f v vol \ subdivFace g f vol = subdivFace' g f v 0 (tl vol)" by (simp_all add: subdivFace_def pre_subdivFace_def pre_subdivFace_face_def) lemma pre_subdivFace'_None: "pre_subdivFace' g f v' v n (None # vol) \ pre_subdivFace' g f v' v (Suc n) vol" by(auto simp: pre_subdivFace'_def dest:invalidVertexList_shorten split:if_split_asm) declare verticesFrom_between [simp del] (* zu zeigen: 1. vol \ [] \ [v\verticesFrom f21 v' . v \ set (removeNones vol)] = removeNones vol 2. vol \ [] \ before (verticesFrom f21 v') u (hd (removeNones vol)) 3. vol \ [] \ distinct (vertices f21) 4. vol \ [] \ last vol = Some (last (verticesFrom f21 v')) *) lemma verticesFrom_split: "v # tl (verticesFrom f v) = verticesFrom f v" by (auto simp: verticesFrom_Def) lemma verticesFrom_v: "distinct (vertices f) \ vertices f = a @ v # b \ verticesFrom f v = v # b @ a" by (simp add: verticesFrom_Def) lemma splitAt_fst[simp]: "distinct xs \ xs = a @ v # b \ fst (splitAt v xs) = a" by auto lemma splitAt_snd[simp]: "distinct xs \ xs = a @ v # b \ snd (splitAt v xs) = b" by auto lemma verticesFrom_splitAt_v_fst[simp]: "distinct (verticesFrom f v) \ fst (splitAt v (verticesFrom f v)) = []" by (simp add: verticesFrom_Def) lemma verticesFrom_splitAt_v_snd[simp]: "distinct (verticesFrom f v) \ snd (splitAt v (verticesFrom f v)) = tl (verticesFrom f v)" by (simp add: verticesFrom_Def) lemma filter_distinct_at: "distinct xs \ xs = (as @ u # bs) \ [v\xs. v = u \ P v] = u # us \ [v\bs. P v] = us \ [v\as. P v] = []" apply (subgoal_tac "filter P as @ u # filter P bs = [] @ u # us") apply (drule local_help') by (auto simp: filter_Cons2) lemma filter_distinct_at3: "distinct xs \ xs = (as @ u # bs) \ [v\xs. v = u \ P v] = u # us \ \ z \ set zs. z \ set as \ \ ( P z) \ [v\zs@bs. P v] = us" apply (drule filter_distinct_at) apply assumption+ apply simp by (induct zs) auto lemma filter_distinct_at4: "distinct xs \ xs = (as @ u # bs) \ [v\xs. v = u \ v \ set us] = u # us \ set zs \ set us \ {u} \ set as \ [v \ zs@bs. v \ set us] = us" proof - assume vors: "distinct xs" "xs = (as @ u # bs)" "[v\xs. v = u \ v \ set us] = u # us" "set zs \ set us \ {u} \ set as" then have "distinct ([v\xs. v = u \ v \ set us])" apply (rule_tac distinct_filter) by simp with vors have dist: "distinct (u # us)" by auto with vors show ?thesis apply (rule_tac filter_distinct_at3) by assumption+ auto qed lemma filter_distinct_at5: "distinct xs \ xs = (as @ u # bs) \ [v\xs. v = u \ v \ set us] = u # us \ set zs \ set xs \ {u} \ set as \ [v \ zs@bs. v \ set us] = us" proof - assume vors: "distinct xs" "xs = (as @ u # bs)" "[v\xs. v = u \ v \ set us] = u # us" "set zs \ set xs \ {u} \ set as" have "set ([v\xs. v = u \ v \ set us]) \ set xs" by auto with vors have "set (u # us) \ set xs" by simp then have "set us \ set xs" by simp with vors have "set zs \ set us \ set zs \ insert u (set as \ set bs)" by auto with vors show ?thesis apply (rule_tac filter_distinct_at4) apply assumption+ by auto qed lemma filter_distinct_at6: "distinct xs \ xs = (as @ u # bs) \ [v\xs. v = u \ v \ set us] = u # us \ set zs \ set xs \ {u} \ set as \ [v \ zs@bs. v \ set us] = us \ [v \ bs. v \ set us] = us" proof - assume vors: "distinct xs" "xs = (as @ u # bs)" "[v \ xs. v = u \ v \ set us] = u # us" "set zs \ set xs \ {u} \ set as" then show ?thesis apply (rule_tac conjI) apply (rule_tac filter_distinct_at5) apply assumption+ apply (drule filter_distinct_at) apply assumption+ by auto qed lemma filter_distinct_at_special: "distinct xs \ xs = (as @ u # bs) \ [v\xs. v = u \ v \ set us] = u # us \ set zs \ set xs \ {u} \ set as \ us = hd_us # tl_us \ [v \ zs@bs. v \ set us] = us \ hd_us \ set bs" proof - assume vors: "distinct xs" "xs = (as @ u # bs)" "[v\xs. v = u \ v \ set us] = u # us" "set zs \ set xs \ {u} \ set as" "us = hd_us # tl_us" then have "[v \ zs@bs. v \ set us] = us \ [v\bs. v \ set us] = us" by (rule_tac filter_distinct_at6) with vors show ?thesis apply (rule_tac conjI) apply safe apply simp apply (subgoal_tac "set (hd_us # tl_us) \ set bs") apply simp apply (subgoal_tac "set [v\bs . v = hd_us \ v \ set tl_us] \ set bs") apply simp by (rule_tac filter_is_subset) qed (* später ggf. pre_splitFace eliminieren *) (* nein, Elimination nicht sinnvoll *) lemma pre_subdivFace'_Some1': assumes pre_add: "pre_subdivFace' g f v' v n ((Some u) # vol)" and pre_fdg: "pre_splitFace g v u f ws" and fdg: "f21 = fst (snd (splitFace g v u f ws))" and g': "g' = snd (snd (splitFace g v u f ws))" shows "pre_subdivFace' g' f21 v' u 0 vol" proof (cases "vol = []") case True then show ?thesis using pre_add fdg pre_fdg apply(unfold pre_subdivFace'_def pre_splitFace_def) apply (simp add: splitFace_def split_face_def split_def del:distinct.simps) apply (rule conjI) apply(clarsimp) apply(rule before_between) apply(erule (5) rotate_before_vFrom) apply(erule not_sym) apply (clarsimp simp:between_distinct between_not_r1 between_not_r2) apply(blast dest:inbetween_inset) done next case False with pre_add have "removeNones vol \ []" apply (cases "vol" rule: rev_exhaust) by (auto simp: pre_subdivFace'_def) then have removeNones_split: "removeNones vol = hd (removeNones vol) # tl (removeNones vol)" by auto from pre_add have dist: "distinct (removeNones ((Some u) # vol))" by (rule_tac pre_subdivFace'_distinct) from pre_add have v': "v' \ \ f" by (auto simp: pre_subdivFace'_def) hence "(vertices f) \ (verticesFrom f v')" by (rule verticesFrom_congs) hence set_eq: "set (verticesFrom f v') = \ f" apply (rule_tac sym) by (rule congs_pres_nodes) from pre_fdg fdg have dist_f21: "distinct (vertices f21)" by auto from pre_add have pre_bet': "pre_between (verticesFrom f v') u v" apply (simp add: pre_between_def pre_subdivFace'_def) apply (elim conjE) apply (thin_tac "n = 0 \ \ is_duplicateEdge g f v u") apply (thin_tac "v' = v \ 0 < n \ v' = v \ u \ last (verticesFrom f v') \ v' \ v") apply (auto simp add: before_def) apply (subgoal_tac "distinct (verticesFrom f v')") apply simp apply (rule_tac verticesFrom_distinct) by auto with pre_add have pre_bet: "pre_between (vertices f) u v" apply (subgoal_tac "(vertices f) \ (verticesFrom f v')") apply (simp add: pre_between_def pre_subdivFace'_def) by (auto dest: congs_pres_nodes intro: verticesFrom_congs simp: pre_subdivFace'_def) from pre_bet pre_add have bet_eq[simp]: "between (vertices f) u v = between (verticesFrom f v') u v" by (auto intro: verticesFrom_between simp: pre_subdivFace'_def) from fdg have f21_split_face: "f21 = snd (split_face f v u ws)" by (simp add: splitFace_def split_def) then have f21: "f21 = Face (u # between (vertices f) u v @ v # ws) Nonfinal" by (simp add: split_face_def) with pre_add pre_bet' have vert_f21: "vertices f21 = u # snd (splitAt u (verticesFrom f v')) @ fst (splitAt v (verticesFrom f v')) @ v # ws" apply (drule_tac pre_between_symI) by (auto simp: pre_subdivFace'_def between_simp2 intro: pre_between_symI) moreover from pre_add have "v \ set (verticesFrom f v')" by (auto simp: pre_subdivFace'_def before_def) then have "verticesFrom f v' = fst (splitAt v (verticesFrom f v')) @ v # snd (splitAt v (verticesFrom f v'))" by (auto dest: splitAt_ram) then have m: "v' # tl (verticesFrom f v') = fst (splitAt v (verticesFrom f v')) @ v # snd (splitAt v (verticesFrom f v'))" by (simp add: verticesFrom_split) then have vv': "v \ v' \ fst (splitAt v (verticesFrom f v')) = v' # tl (fst (splitAt v (verticesFrom f v')))" by (cases "fst (splitAt v (verticesFrom f v'))") auto ultimately have "v \ v' \ vertices f21 = u # snd (splitAt u (verticesFrom f v')) @ v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws" by auto moreover with f21 have rule2: "v' \ \ f21" by auto with dist_f21 have dist_f21_v': "distinct (verticesFrom f21 v')" by auto ultimately have m1: "v \ v' \ verticesFrom f21 v' = v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws @ u # snd (splitAt u (verticesFrom f v'))" apply auto apply (subgoal_tac "snd (splitAt v' (vertices f21)) = tl (fst (splitAt v (verticesFrom f v'))) @ v # ws") apply (subgoal_tac "fst (splitAt v' (vertices f21)) = u # snd (splitAt u (verticesFrom f v'))") apply (subgoal_tac "verticesFrom f21 v' = v' # snd (splitAt v' (vertices f21)) @ fst (splitAt v' (vertices f21))") apply simp apply (intro verticesFrom_v dist_f21) apply force apply (subgoal_tac "distinct (vertices f21)") apply simp apply (rule_tac dist_f21) apply (subgoal_tac "distinct (vertices f21)") apply simp by (rule_tac dist_f21) from pre_add have dist_vf_v': "distinct (verticesFrom f v')" by (simp add: pre_subdivFace'_def) with vert_f21 have m2: "v = v' \ verticesFrom f21 v' = v' # ws @ u # snd (splitAt u (verticesFrom f v'))" apply auto apply (intro verticesFrom_v dist_f21) by simp from pre_add have u: "u \ set (verticesFrom f v')" by (fastforce simp: pre_subdivFace'_def before_def) then have split_u: "verticesFrom f v' = fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v'))" by (auto dest!: splitAt_ram) then have rule1': "[v \ snd (splitAt u (verticesFrom f v')) . v \ set (removeNones vol)] = removeNones vol" proof - from split_u have "v' # tl (verticesFrom f v') = fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v'))" by (simp add: verticesFrom_split) have "help": "set [] \ set (verticesFrom f v') \ {u} \ set (fst (splitAt u (verticesFrom f v')))" by auto from split_u dist_vf_v' pre_add have "[v \ [] @ snd (splitAt u (verticesFrom f v')) . v \ set (removeNones vol)] = removeNones vol" apply (rule_tac filter_distinct_at5) apply assumption+ apply (simp add: pre_subdivFace'_def) by (rule "help") then show ?thesis by auto qed then have inSnd_u: "\ x. x \ set (removeNones vol) \ x \ set (snd (splitAt u (verticesFrom f v')))" apply (subgoal_tac "x \ set [v \ snd (splitAt u (verticesFrom f v')) . v \ set (removeNones vol)] \ x \ set (snd (splitAt u (verticesFrom f v')))") apply force apply (thin_tac "[v \ snd (splitAt u (verticesFrom f v')) . v \ set (removeNones vol)] = removeNones vol") by simp from split_u dist_vf_v' have notinFst_u: "\ x. x \ set (removeNones vol) \ x \ set ((fst (splitAt u (verticesFrom f v'))) @ [u])" apply (drule_tac inSnd_u) apply (subgoal_tac "distinct ( fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v')))") apply (thin_tac "verticesFrom f v' = fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v'))") apply simp apply safe apply (subgoal_tac "x \ set (fst (splitAt u (verticesFrom f v'))) \ set (snd (splitAt u (verticesFrom f v')))") apply simp apply (thin_tac "set (fst (splitAt u (verticesFrom f v'))) \ set (snd (splitAt u (verticesFrom f v'))) = {}") apply simp by (simp only:) from rule2 v' have "\ a b. is_nextElem (vertices f) a b \ a \ set (removeNones vol) \ b \ set (removeNones vol) \ is_nextElem (vertices f21) a b" proof - fix a b assume vors: "is_nextElem (vertices f) a b \ a \ set (removeNones vol) \ b \ set (removeNones vol)" define vor_u where "vor_u = fst (splitAt u (verticesFrom f v'))" define nach_u where "nach_u = snd (splitAt u (verticesFrom f v'))" from vors v' have "is_nextElem (verticesFrom f v') a b" by (simp add: verticesFrom_is_nextElem) moreover from vors inSnd_u nach_u_def have "a \ set (nach_u)" by auto moreover from vors inSnd_u nach_u_def have "b \ set (nach_u)" by auto moreover from split_u vor_u_def nach_u_def have "verticesFrom f v' = vor_u @ u # nach_u" by auto moreover note dist_vf_v' ultimately have "is_sublist [a,b] (nach_u)" apply (simp add: is_nextElem_def split:if_split_asm) apply (subgoal_tac "b \ hd (vor_u @ u # nach_u)") apply simp apply (subgoal_tac "distinct (vor_u @ (u # nach_u))") apply (drule is_sublist_at5) apply simp apply simp apply (erule disjE) apply (drule is_sublist_in1)+ apply (subgoal_tac "b \ set vor_u \ set nach_u") apply simp apply (thin_tac "set vor_u \ set nach_u = {}") apply simp apply (erule disjE) apply (subgoal_tac "distinct ([u] @ nach_u)") apply (drule is_sublist_at5) apply simp apply simp apply (erule disjE) apply simp apply simp apply simp apply (subgoal_tac "distinct (vor_u @ (u # nach_u))") apply (drule is_sublist_at5) apply simp apply (erule disjE) apply (drule is_sublist_in1)+ apply simp apply (erule disjE) apply (drule is_sublist_in1)+ apply simp apply simp apply simp apply simp apply (cases "vor_u") by auto with nach_u_def have "is_sublist [a,b] (snd (splitAt u (verticesFrom f v')))" by auto then have "is_sublist [a,b] (verticesFrom f21 v')" apply (cases "v = v'") apply (simp_all add: m1 m2) apply (subgoal_tac "is_sublist [a, b] ((v' # ws @ [u]) @ snd (splitAt u (verticesFrom f v')) @ [])") apply simp apply (rule is_sublist_add) apply simp apply (subgoal_tac "is_sublist [a, b] ((v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws @ [u]) @ (snd (splitAt u (verticesFrom f v'))) @ [])") apply simp apply (rule is_sublist_add) by simp with rule2 show "is_nextElem (vertices f) a b \ a \ set (removeNones vol) \ b \ set (removeNones vol) \ is_nextElem (vertices f21) a b" apply (simp add: verticesFrom_is_nextElem) by (auto simp: is_nextElem_def) qed with pre_add dist_f21 have rule5': "\ a b. (a,b) \ edges f \ a \ set (removeNones vol) \ b \ set (removeNones vol) \ (a, b) \ edges f21" by (simp add: is_nextElem_edges_eq pre_subdivFace'_def) have rule1: "[v\verticesFrom f21 v' . v \ set (removeNones vol)] = removeNones vol \ hd (removeNones vol) \ set (snd (splitAt u (verticesFrom f v')))" proof (cases "v = v'") case True from split_u have "v' # tl (verticesFrom f v') = fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v'))" by (simp add: verticesFrom_split) then have "u \ v' \ fst (splitAt u (verticesFrom f v')) = v' # tl (fst (splitAt u (verticesFrom f v')))" by (cases "fst (splitAt u (verticesFrom f v'))") auto moreover have "v' \ set (v' # tl (fst (splitAt u (verticesFrom f v'))))" by simp ultimately have "u \ v' \ v' \ set (fst (splitAt u (verticesFrom f v')))" by simp moreover from pre_fdg have "set (v' # ws @ [u]) \ set (verticesFrom f v') \ {v', u}" apply (simp add: set_eq) by (unfold pre_splitFace_def) auto ultimately have "help": "set (v' # ws @ [u]) \ set (verticesFrom f v') \ {u} \ set (fst (splitAt u (verticesFrom f v')))" apply (rule_tac subset_trans) apply assumption apply (cases "u = v'") by simp_all from split_u dist_vf_v' pre_add pre_fdg removeNones_split have "[v \ (v' # ws @ [u]) @ snd (splitAt u (verticesFrom f v')) . v \ set (removeNones vol)] = removeNones vol \ hd (removeNones vol) \ set (snd (splitAt u (verticesFrom f v')))" apply (rule_tac filter_distinct_at_special) apply assumption+ apply (simp add: pre_subdivFace'_def) apply (rule "help") . with True m2 show ?thesis by auto next case False with m1 dist_f21_v' have ne_uv': "u \ v'" by auto define fst_u where "fst_u = fst (splitAt u (verticesFrom f v'))" define fst_v where "fst_v = fst (splitAt v (verticesFrom f v'))" from pre_add u dist_vf_v' have "v \ set (fst (splitAt u (verticesFrom f v')))" apply (rule_tac before_dist_r1) by (auto simp: pre_subdivFace'_def) with fst_u_def have "fst_u = fst (splitAt v (fst (splitAt u (verticesFrom f v')))) @ v # snd (splitAt v (fst (splitAt u (verticesFrom f v'))))" by (auto dest: splitAt_ram) with pre_add fst_v_def pre_bet' have fst_u':"fst_u = fst_v @ v # snd (splitAt v (fst (splitAt u (verticesFrom f v'))))" by (simp add: pre_subdivFace'_def) from pre_fdg have "set (v # ws @ [u]) \ set (verticesFrom f v') \ {v, u}" apply (simp add: set_eq) by (unfold pre_splitFace_def) auto with fst_u' have "set (v # ws @ [u]) \ set (verticesFrom f v') \ {u} \ set fst_u" by auto moreover from fst_u' have "set fst_v \ set fst_u" by auto ultimately have "(set (v # ws @ [u]) \ set fst_v) \ set (verticesFrom f v') \ {u} \ set fst_u" by auto with fst_u_def fst_v_def have "set (fst (splitAt v (verticesFrom f v')) @ v # ws @ [u]) \ set (verticesFrom f v') \ {u} \ set (fst (splitAt u (verticesFrom f v')))" by auto moreover with False vv' have "v' # tl (fst (splitAt v (verticesFrom f v'))) = fst (splitAt v (verticesFrom f v'))" by auto ultimately have "set ((v' # tl (fst (splitAt v (verticesFrom f v')))) @ v # ws @ [u]) \ set (verticesFrom f v') \ {u} \ set (fst (splitAt u (verticesFrom f v')))" by (simp only:) then have "help": "set (v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws @ [u]) \ set (verticesFrom f v') \ {u} \ set (fst (splitAt u (verticesFrom f v')))" by auto from split_u dist_vf_v' pre_add pre_fdg removeNones_split have "[v \ (v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws @ [u]) @ snd (splitAt u (verticesFrom f v')) . v \ set (removeNones vol)] = removeNones vol \ hd (removeNones vol) \ set (snd (splitAt u (verticesFrom f v')))" apply (rule_tac filter_distinct_at_special) apply assumption+ apply (simp add: pre_subdivFace'_def) apply (rule "help") . with False m1 show ?thesis by auto qed from rule1 have "(hd (removeNones vol)) \ set (snd (splitAt u (verticesFrom f v')))" by auto with m1 m2 dist_f21_v' have rule3: "before (verticesFrom f21 v') u (hd (removeNones vol))" proof - assume hd_ram: "(hd (removeNones vol)) \ set (snd (splitAt u (verticesFrom f v')))" from m1 m2 dist_f21_v' have "distinct (snd (splitAt u (verticesFrom f v')))" apply (cases "v = v'") by auto moreover define z1 where "z1 = fst (splitAt (hd (removeNones vol)) (snd (splitAt u (verticesFrom f v'))))" define z2 where "z2 = snd (splitAt (hd (removeNones vol)) (snd (splitAt u (verticesFrom f v'))))" note z1_def z2_def hd_ram ultimately have "snd (splitAt u (verticesFrom f v')) = z1 @ (hd (removeNones vol)) # z2" by (auto intro: splitAt_ram) with m1 m2 show ?thesis apply (cases "v = v'") apply (auto simp: before_def) apply (intro exI ) apply (subgoal_tac "v' # ws @ u # z1 @ hd (removeNones vol) # z2 = (v' # ws) @ u # z1 @ hd (removeNones vol) # z2") apply assumption apply simp apply (intro exI ) apply (subgoal_tac "v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws @ u # z1 @ hd (removeNones vol) # z2 = (v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws) @ u # z1 @ hd (removeNones vol) # z2") apply assumption by simp qed from rule1 have ne:"(hd (removeNones vol)) \ set (snd (splitAt u (verticesFrom f v')))" by auto with m1 m2 have "last (verticesFrom f21 v') = last (snd (splitAt u (verticesFrom f v')))" apply (cases "snd (splitAt u (verticesFrom f v'))" rule: rev_exhaust) apply simp_all apply (cases "v = v'") by simp_all moreover from ne have "last (fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v'))) = last (snd (splitAt u (verticesFrom f v')))" by auto moreover note split_u ultimately have rule4: "last (verticesFrom f v') = last (verticesFrom f21 v')" by simp have l: "\ a b f v. v \ set (vertices f) \ is_nextElem (vertices f) a b = is_nextElem (verticesFrom f v) a b " apply (rule is_nextElem_congs_eq) by (rule verticesFrom_congs) define f12 where "f12 = fst (split_face f v u ws)" then have f12_fdg: "f12 = fst (splitFace g v u f ws)" by (simp add: splitFace_def split_def) from pre_bet pre_add have bet_eq2[simp]: "between (vertices f) v u = between (verticesFrom f v') v u" apply (drule_tac pre_between_symI) by (auto intro: verticesFrom_between simp: pre_subdivFace'_def) from f12_fdg have f12_split_face: "f12 = fst (split_face f v u ws)" by (simp add: splitFace_def split_def) then have f12: "f12 = Face (rev ws @ v # between (verticesFrom f v') v u @ [u]) Nonfinal" by (simp add: split_face_def) then have "vertices f12 = rev ws @ v # between (verticesFrom f v') v u @ [u]" by simp with pre_add pre_bet' have vert_f12: "vertices f12 = rev ws @ v # snd (splitAt v (fst (splitAt u (verticesFrom f v')))) @ [u]" apply (subgoal_tac "between (verticesFrom f v') v u = fst (splitAt u (snd (splitAt v (verticesFrom f v'))))") apply (simp add: pre_subdivFace'_def) apply (rule between_simp1) apply (simp add: pre_subdivFace'_def) apply (rule pre_between_symI) . with dist_f21_v' have removeNones_vol_not_f12: "\ x. x \ set (removeNones vol) \ x \ set (vertices f12)" apply (frule_tac notinFst_u) apply (drule inSnd_u) apply simp apply (case_tac "v = v'") apply (simp add: m1 m2) apply (rule conjI) apply force apply (rule conjI) apply (rule ccontr) apply simp apply (subgoal_tac "x \ set ws \ set (snd (splitAt u (verticesFrom f v')))") apply simp apply (elim conjE) apply (thin_tac "set ws \ set (snd (splitAt u (verticesFrom f v'))) = {}") apply simp apply force apply (simp add: m1 m2) apply (rule conjI) apply force apply (rule conjI) apply (rule ccontr) apply simp apply (subgoal_tac "x \ set ws \ set (snd (splitAt u (verticesFrom f v')))") apply simp apply (elim conjE) apply (thin_tac "set ws \ set (snd (splitAt u (verticesFrom f v'))) = {}") apply simp by force from pre_fdg f12_split_face have dist_f12: "distinct (vertices f12)" by (auto intro: split_face_distinct1') then have removeNones_vol_edges_not_f12: "\ x y. x \ set (removeNones vol) \ (x,y) \ edges f12" (* ? *) apply (drule_tac removeNones_vol_not_f12) by auto from dist_f12 have removeNones_vol_edges_not_f12': "\ x y. y \ set (removeNones vol) \ (x,y) \ edges f12" apply (drule_tac removeNones_vol_not_f12) by auto from f12_fdg pre_fdg g' fdg have face_set_eq: "\ g' \ {f} = {f12, f21} \ \ g" apply (rule_tac splitFace_faces_1) by (simp_all) have rule5'': "\ a b. (a,b) \ edges g' \ (a,b) \ edges g \ a \ set (removeNones vol) \ b \ set (removeNones vol) \ (a, b) \ edges f21" (* ? *) apply (simp add: edges_graph_def) apply safe apply (case_tac "x = f") apply simp apply (rule rule5') apply safe apply (subgoal_tac "x \ \ g' \ {f}") apply (thin_tac "x \ f") apply (thin_tac "x \ set (faces g')") apply (simp only: add: face_set_eq) apply safe apply (drule removeNones_vol_edges_not_f12) by auto have rule5''': "\ a b. (a,b) \ edges g' \ (a,b) \ edges g \ a \ set (removeNones vol) \ b \ set (removeNones vol) \ (a, b) \ edges f21" (* ? *) apply (simp add: edges_graph_def) apply safe apply (case_tac "x = f") apply simp apply (rule rule5') apply safe apply (subgoal_tac "x \ \ g' \ {f}") apply (thin_tac "x \ f") apply (thin_tac "x \ \ g'") apply (simp only: add: face_set_eq) apply safe apply (drule removeNones_vol_edges_not_f12) by auto from pre_fdg fdg f12_fdg g' have edges_g'1: "ws \ [] \ edges g' = edges g \ Edges ws \ Edges(rev ws) \ {(u, last ws), (hd ws, v), (v, hd ws), (last ws, u)}" apply (rule_tac splitFace_edges_g') apply simp apply (subgoal_tac "(f12, f21, g') = splitFace g v u f ws") apply assumption by auto from pre_fdg fdg f12_fdg g' have edges_g'2: "ws = [] \ edges g' = edges g \ {(v, u), (u, v)}" apply (rule_tac splitFace_edges_g'_vs) apply simp apply (subgoal_tac "(f12, f21, g') = splitFace g v u f []") apply assumption by auto from f12_split_face f21_split_face have split: "(f12,f21) = split_face f v u ws" by simp from pre_add have "\ invalidVertexList g f vol" by (auto simp: pre_subdivFace'_def dest: invalidVertexList_shorten) then have rule5: "\ invalidVertexList g' f21 vol" apply (simp add: invalidVertexList_def) apply (intro allI impI) apply (case_tac "vol!i") apply simp+ apply (case_tac "vol!Suc i") apply simp+ apply (subgoal_tac "\ is_duplicateEdge g f a aa") apply (thin_tac "\i<|vol| - Suc 0. \ (case vol ! i of None \ False | Some a \ case_option False (is_duplicateEdge g f a) (vol ! (i+1)))") apply (simp add: is_duplicateEdge_def) apply (subgoal_tac "a \ set (removeNones vol) \ aa \ set (removeNones vol)") apply (rule conjI) apply (rule impI) apply (case_tac "(a, aa) \ edges f") apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (frule split_face_edges_or [OF split]) apply simp apply (simp add: removeNones_vol_edges_not_f12) apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) apply (case_tac "(aa, a) \ edges f") apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (frule split_face_edges_or [OF split]) apply simp apply (simp add: removeNones_vol_edges_not_f12) apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) apply simp apply (case_tac "ws = []") apply (frule edges_g'2) apply simp apply (subgoal_tac "pre_split_face f v u []") apply (subgoal_tac "(f12, f21) = split_face f v u ws") apply (case_tac "between (vertices f) u v = []") apply (frule split_face_edges_f21_bet_vs) apply simp apply simp apply simp apply (frule split_face_edges_f21_vs) apply simp apply simp apply simp apply (case_tac "a = v \ aa = u") apply simp apply simp apply (rule split) apply (subgoal_tac "pre_split_face f v u ws") apply simp apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) apply (frule edges_g'1) apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (subgoal_tac "(f12, f21) = split_face f v u ws") apply (case_tac "between (vertices f) u v = []") apply (frule split_face_edges_f21_bet) apply simp apply simp apply simp apply simp apply (case_tac "a = u \ aa = last ws") apply simp apply simp apply (case_tac "a = hd ws \ aa = v") apply simp apply simp apply (case_tac "a = v \ aa = hd ws") apply simp apply simp apply (case_tac "a = last ws \ aa = u") apply simp apply simp apply (case_tac "(a, aa) \ Edges ws") apply simp apply simp apply (frule split_face_edges_f21) apply simp apply simp apply simp apply simp apply (force) apply (rule split) apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) apply (rule impI) apply (case_tac "(aa,a) \ edges f") apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (frule split_face_edges_or [OF split]) apply simp apply (simp add: removeNones_vol_edges_not_f12) apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) apply (case_tac "(a,aa) \ edges f") apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (frule split_face_edges_or [OF split]) apply simp apply (simp add: removeNones_vol_edges_not_f12) apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) apply simp apply (case_tac "ws = []") apply (frule edges_g'2) apply simp apply (subgoal_tac "pre_split_face f v u []") apply (subgoal_tac "(f12, f21) = split_face f v u ws") apply (case_tac "between (vertices f) u v = []") apply (frule split_face_edges_f21_bet_vs) apply simp apply simp apply simp apply (frule split_face_edges_f21_vs) apply simp apply simp apply simp apply force apply (rule split) apply (subgoal_tac "pre_split_face f v u ws") apply simp apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) apply (frule edges_g'1) apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (subgoal_tac "(f12, f21) = split_face f v u ws") apply (case_tac "between (vertices f) u v = []") apply (frule split_face_edges_f21_bet) apply simp apply simp apply simp apply (force) apply (frule split_face_edges_f21) apply simp apply simp apply simp apply simp apply (force) apply (rule split) apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) apply (rule conjI) apply (subgoal_tac "Some a \ set vol") apply (induct vol) apply simp apply force apply (subgoal_tac "vol ! i \ set vol") apply simp apply (rule nth_mem) apply arith apply (subgoal_tac "Some aa \ set vol") apply (induct vol) apply simp apply force apply (subgoal_tac "vol ! (Suc i) \ set vol") apply simp apply (rule nth_mem) apply arith by auto from pre_fdg dist_f21 v' have dists: "distinct (vertices f)" "distinct (vertices f12)" "distinct (vertices f21)" "v' \ \ f" apply auto defer apply (drule splitFace_distinct2) apply (simp add: f12_fdg) apply (unfold pre_splitFace_def) by simp with pre_fdg have edges_or: "\ a b. (a,b) \ edges f \ (a,b) \ edges f12 \ (a,b) \ edges f21" apply (rule_tac split_face_edges_or) apply (simp add: f12_split_face f21_split_face) by simp+ from pre_fdg have dist_f: "distinct (vertices f)" apply (unfold pre_splitFace_def) by simp (* lemma *) from g' have edges_g': "edges g' = (UN h:set(replace f [snd (split_face f v u ws)] (faces g)). edges h) \ edges (fst (split_face f v u ws))" by (auto simp add: splitFace_def split_def edges_graph_def) (* lemma *) from pre_fdg edges_g' have edges_g'_or: "\ a b. (a,b) \ edges g' \ (a,b) \ edges g \ (a,b) \ edges f12 \ (a,b) \ edges f21" apply simp apply (case_tac "(a, b) \ edges (fst (split_face f v u ws))") apply (simp add:f12_split_face) apply simp apply (elim bexE) apply (simp add: f12_split_face) apply (case_tac "x \ \ g") apply (induct g) apply (simp add: edges_graph_def) apply (rule disjI1) apply (rule bexI) apply simp apply simp apply (drule replace1) apply simp by (simp add: f21_split_face) have rule6: "0 < |vol| \ \ invalidVertexList g f (Some u # vol) \ (\y. hd vol = Some y) \ \ is_duplicateEdge g' f21 u (the (hd vol))" apply (rule impI) apply (erule exE) apply simp apply (case_tac vol) apply simp+ apply (simp add: invalidVertexList_def) apply (erule allE) apply (erule impE) apply force apply (simp) apply (subgoal_tac "y \ \ f12") defer apply (rule removeNones_vol_not_f12) apply simp apply (simp add: is_duplicateEdge_def) apply (subgoal_tac "y \ set (removeNones vol)") apply (rule conjI) apply (rule impI) apply (case_tac "(u, y) \ edges f") apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (frule split_face_edges_or [OF split]) apply simp apply (simp add: removeNones_vol_edges_not_f12') apply (rule pre_splitFace_pre_split_face) apply simp apply (rule pre_fdg) apply (case_tac "(y, u) \ edges f") apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (frule split_face_edges_or [OF split]) apply simp apply (simp add: removeNones_vol_edges_not_f12) apply (rule pre_splitFace_pre_split_face) apply simp apply (rule pre_fdg) apply simp apply (case_tac "ws = []") apply (frule edges_g'2) apply simp apply (subgoal_tac "pre_split_face f v u []") apply (subgoal_tac "(f12, f21) = split_face f v u ws") apply (case_tac "between (vertices f) u v = []") apply (frule split_face_edges_f21_bet_vs) apply simp apply simp apply simp apply (frule split_face_edges_f21_vs) apply simp apply simp apply simp apply force apply (rule split) apply (subgoal_tac "pre_split_face f v u ws") apply simp apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) apply (frule edges_g'1) apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (subgoal_tac "(f12, f21) = split_face f v u ws") apply (case_tac "between (vertices f) u v = []") apply (frule split_face_edges_f21_bet) apply simp apply simp apply simp apply (force) apply (frule split_face_edges_f21) apply simp apply simp apply simp apply simp apply (force) apply (rule split) apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) apply (rule impI) apply (case_tac "(u, y) \ edges f") apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (frule split_face_edges_or [OF split]) apply simp apply (simp add: removeNones_vol_edges_not_f12') apply (rule pre_splitFace_pre_split_face) apply simp apply (rule pre_fdg) apply (case_tac "(y, u) \ edges f") apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (frule split_face_edges_or [OF split]) apply simp apply (simp add: removeNones_vol_edges_not_f12) apply (rule pre_splitFace_pre_split_face) apply simp apply (rule pre_fdg) apply simp apply (case_tac "ws = []") apply (frule edges_g'2) apply simp apply (subgoal_tac "pre_split_face f v u []") apply (subgoal_tac "(f12, f21) = split_face f v u ws") apply (case_tac "between (vertices f) u v = []") apply (frule split_face_edges_f21_bet_vs) apply simp apply simp apply simp apply (frule split_face_edges_f21_vs) apply simp apply simp apply simp apply force apply (rule split) apply (subgoal_tac "pre_split_face f v u ws") apply simp apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) apply (frule edges_g'1) apply simp apply (subgoal_tac "pre_split_face f v u ws") apply (subgoal_tac "(f12, f21) = split_face f v u ws") apply (case_tac "between (vertices f) u v = []") apply (frule split_face_edges_f21_bet) apply simp apply simp apply simp apply (force) apply (frule split_face_edges_f21) apply simp apply simp apply simp apply simp apply (force) apply (rule split) apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg) by simp have u21: "u \ \ f21" by(simp add:f21) from fdg have "\ final f21" by(simp add:splitFace_def split_face_def split_def) with pre_add rule1 rule2 rule3 rule4 rule5 rule6 dist_f21 False dist u21 show ?thesis by (simp_all add: pre_subdivFace'_def l) qed lemma before_filter: "\ ys. filter P xs = ys \ distinct xs \ before ys u v \ before xs u v" - using subst_all subst_all' [simp del] + supply subst_all [simp del] apply (subgoal_tac "P u") apply (subgoal_tac "P v") apply (subgoal_tac "pre_between xs u v") apply (rule ccontr) apply (simp add: before_xor) apply (subgoal_tac "before ys v u") apply (subgoal_tac "\ before ys v u") apply simp apply (rule before_dist_not1) apply force apply simp apply (simp add: before_def) apply (elim exE) apply simp apply (subgoal_tac "a @ u # b @ v # c = filter P aa @ v # filter P ba @ u # filter P ca") apply (intro exI) apply assumption apply simp apply (subgoal_tac "u \ set ys \ v \ set ys \ u \ v") apply (simp add: pre_between_def) apply force apply (subgoal_tac "distinct ys") apply (simp add: before_def) apply (elim exE) apply simp apply force apply (subgoal_tac "v \ set (filter P xs)") apply force apply (simp add: before_def) apply (elim exE) apply simp apply (subgoal_tac "u \ set (filter P xs)") apply force apply (simp add: before_def) apply (elim exE) by simp lemma pre_subdivFace'_Some2: "pre_subdivFace' g f v' v 0 ((Some u) # vol) \ pre_subdivFace' g f v' u 0 vol" apply (cases "vol = []") apply (simp add: pre_subdivFace'_def) apply (cases "u = v'") apply simp apply(rule verticesFrom_in') apply(rule last_in_set) apply(simp add:verticesFrom_Def) apply clarsimp apply (simp add: pre_subdivFace'_def) apply (elim conjE) apply (thin_tac "v' = v \ u \ last (verticesFrom f v') \ v' \ v") apply auto apply(rule verticesFrom_in'[where v = v']) apply(clarsimp simp:before_def) apply simp apply (rule ovl_shorten) apply simp apply (subgoal_tac "[v \ verticesFrom f v' . v \ set (removeNones ((Some u) # vol))] = removeNones ((Some u) # vol)") apply assumption apply simp apply (rule before_filter) apply assumption apply simp apply (simp add: before_def) apply (intro exI) apply (subgoal_tac "u # removeNones vol = [] @ u # [] @ hd (removeNones vol) # tl (removeNones vol)") apply assumption apply simp apply (subgoal_tac "removeNones vol \ []") apply simp apply (cases vol rule: rev_exhaust) apply simp_all apply (simp add: invalidVertexList_shorten) apply (simp add: is_duplicateEdge_def) apply (case_tac "vol") apply simp apply simp apply (simp add: invalidVertexList_def) apply (elim allE) apply (rotate_tac -1) apply (erule impE) apply (subgoal_tac "0 < Suc |list|") apply assumption apply simp apply simp by (simp add: is_duplicateEdge_def) lemma pre_subdivFace'_preFaceDiv: "pre_subdivFace' g f v' v n ((Some u) # vol) \ f \ \ g \ (f \ v = u \ n \ 0) \ \ f \ \ g \ pre_splitFace g v u f [countVertices g ..< countVertices g + n]" proof - assume pre_add: "pre_subdivFace' g f v' v n ((Some u) # vol)" and f: "f \ \ g" and nextVert: "(f \ v = u \ n \ 0)" and subset: "\ f \ \ g" have "distinct [countVertices g ..< countVertices g + n]" by (induct n) auto moreover have "\ g \ set [countVertices g ..< countVertices g + n] = {}" apply (cases g) by auto with subset have "\ f \ set [countVertices g ..< countVertices g + n] = {}" by auto moreover from pre_add have "\ f = set (verticesFrom f v')" apply (intro congs_pres_nodes verticesFrom_congs) by (simp add: pre_subdivFace'_def) with pre_add have "help": "v \ \ f \ u \ \ f \ v \ u" apply (simp add: pre_subdivFace'_def before_def) apply (elim conjE exE) apply (subgoal_tac "distinct (verticesFrom f v')") apply force apply (rule verticesFrom_distinct) by simp_all moreover from "help" pre_add nextVert have help1: "is_nextElem (vertices f) v u \ 0 < n" apply auto apply (simp add: nextVertex_def) by (simp add: nextElem_is_nextElem pre_subdivFace'_def) moreover have help2: "before (verticesFrom f v') v u \ distinct (verticesFrom f v') \ v \ v' \ \ is_nextElem (verticesFrom f v') u v" apply (simp add: before_def is_nextElem_def verticesFrom_hd is_sublist_def) apply safe apply (frule dist_at) apply simp apply (thin_tac "verticesFrom f v' = a @ v # b @ u # c") apply (subgoal_tac "verticesFrom f v' = (as @ [u]) @ v # bs") apply assumption apply simp apply (subgoal_tac "distinct (a @ v # b @ u # c)") apply force by simp note pre_add f moreover(* have "\ m. {k. k < m} \ {k. m \ k \ k < (m + n)} = {}" by auto moreover*) from pre_add f help2 help1 "help" have "[countVertices g.. (v, u) \ edges f \ (u, v) \ edges f" apply (cases "0 < n") apply (induct g) apply simp+ apply (simp add: pre_subdivFace'_def) apply (rule conjI) apply force apply (simp split: if_split_asm) apply (rule ccontr) apply simp apply (subgoal_tac "v = v'") apply simp apply (elim conjE) apply (simp only:) apply (rule verticesFrom_is_nextElem_last) apply force apply force apply (simp add: verticesFrom_is_nextElem [symmetric]) apply (cases "v = v'") apply simp apply (subgoal_tac "v' \ \ f") apply (thin_tac "u \ \ f") apply (simp add: verticesFrom_is_nextElem) apply (rule ccontr) apply simp apply (subgoal_tac "v' \ \ f") apply (drule verticesFrom_is_nextElem_hd) apply simp+ apply (elim conjE) apply (drule help2) apply simp apply simp apply (subgoal_tac "is_nextElem (vertices f) u v = is_nextElem (verticesFrom f v') u v") apply simp apply (rule verticesFrom_is_nextElem) by simp ultimately show ?thesis apply (simp add: pre_subdivFace'_def) apply (unfold pre_splitFace_def) apply simp apply (cases "0 < n") apply (induct g) apply (simp add: ivl_disj_int) apply (auto simp: invalidVertexList_def is_duplicateEdge_def) done qed lemma pre_subdivFace'_Some1: "pre_subdivFace' g f v' v n ((Some u) # vol) \ f \ \ g \ (f \ v = u \ n \ 0) \ \ f \ \ g \ f21 = fst (snd (splitFace g v u f [countVertices g ..< countVertices g + n])) \ g' = snd (snd (splitFace g v u f [countVertices g ..< countVertices g + n])) \ pre_subdivFace' g' f21 v' u 0 vol" apply (subgoal_tac "pre_splitFace g v u f [countVertices g ..< countVertices g + n]") apply (rule pre_subdivFace'_Some1') apply assumption+ apply (simp) apply (rule pre_subdivFace'_preFaceDiv) by auto end diff --git a/thys/Generalized_Counting_Sort/Conservation.thy b/thys/Generalized_Counting_Sort/Conservation.thy --- a/thys/Generalized_Counting_Sort/Conservation.thy +++ b/thys/Generalized_Counting_Sort/Conservation.thy @@ -1,2036 +1,2036 @@ (* Title: An Efficient Generalization of Counting Sort for Large, possibly Infinite Key Ranges Author: Pasquale Noce Software Engineer at HID Global, Italy pasquale dot noce dot lavoro at gmail dot com pasquale dot noce at hidglobal dot com *) section "Proof of objects' conservation" theory Conservation imports Algorithm "HOL-Library.Multiset" begin text \ \null In this section, it is formally proven that GCsort \emph{conserves objects}, viz. that the objects' list returned by function @{const gcsort} contains as many occurrences of any given object as the input objects' list. Here below, steps 5, 6, and 7 of the proof method are accomplished. Particularly, @{text count_inv} is the predicate that will be shown to be invariant over inductive set @{const gcsort_set}. \null \ fun count_inv :: "('a \ nat) \ nat \ nat list \ 'a list \ bool" where "count_inv f (u, ns, xs) = (\x. count (mset xs) x = f x)" lemma gcsort_count_input: "count_inv (count (mset xs)) (0, [length xs], xs)" by simp lemma gcsort_count_intro: "count_inv f t \ count (mset (gcsort_out t)) x = f x" by (cases t, simp add: gcsort_out_def) text \ \null The main task to be accomplished to prove that GCsort conserves objects is to prove that so does function @{const fill} in case its input offsets' list is computed via the composition of functions @{const offs} and @{const enum}, as happens within function @{const round}. To achieve this result, a multi-step strategy will be adopted. The first step, addressed here below, opens with the definition of predicate @{text offs_pred}, satisfied by an offsets' list $ns$ and an objects' list $xs$ just in case each bucket delimited by $ns$ is sufficiently large to accommodate the corresponding objects in $xs$. Then, lemma @{text offs_pred_cons} shows that this predicate, if satisfied initially, keeps being true if each object in $xs$ is consumed as happens within function @{const fill}, viz. increasing the corresponding offset in $ns$ by one. \null \ definition offs_num :: "nat \ 'a list \ ('a, 'b) index_sign \ ('a \ 'b) \ 'b \ 'b \ nat \ nat" where "offs_num n xs index key mi ma i \ length [x\xs. index key x n mi ma = i]" abbreviation offs_set_next :: "nat list \ 'a list \ ('a, 'b) index_sign \ ('a \ 'b) \ 'b \ 'b \ nat \ nat set" where "offs_set_next ns xs index key mi ma i \ {k. k < length ns \ i < k \ 0 < offs_num (length ns) xs index key mi ma k}" abbreviation offs_set_prev :: "nat list \ 'a list \ ('a, 'b) index_sign \ ('a \ 'b) \ 'b \ 'b \ nat \ nat set" where "offs_set_prev ns xs index key mi ma i \ {k. i < length ns \ k < i \ 0 < offs_num (length ns) xs index key mi ma k}" definition offs_next :: "nat list \ nat \ 'a list \ ('a, 'b) index_sign \ ('a \ 'b) \ 'b \ 'b \ nat \ nat" where "offs_next ns ub xs index key mi ma i \ if offs_set_next ns xs index key mi ma i = {} then ub else ns ! Min (offs_set_next ns xs index key mi ma i)" definition offs_none :: "nat list \ nat \ 'a list \ ('a, 'b) index_sign \ ('a \ 'b) \ 'b \ 'b \ nat \ bool" where "offs_none ns ub xs index key mi ma i \ (\j < length ns. 0 < offs_num (length ns) xs index key mi ma j \ i \ {ns ! j + offs_num (length ns) xs index key mi ma j..< offs_next ns ub xs index key mi ma j}) \ offs_num (length ns) xs index key mi ma 0 = 0 \ i < offs_next ns ub xs index key mi ma 0 \ 0 < offs_num (length ns) xs index key mi ma 0 \ i < ns ! 0" definition offs_pred :: "nat list \ nat \ 'a list \ ('a, 'b) index_sign \ ('a \ 'b) \ 'b \ 'b \ bool" where "offs_pred ns ub xs index key mi ma \ \i < length ns. offs_num (length ns) xs index key mi ma i \ offs_next ns ub xs index key mi ma i - ns ! i" lemma offs_num_cons: "offs_num n (x # xs) index key mi ma i = (if index key x n mi ma = i then Suc else id) (offs_num n xs index key mi ma i)" by (simp add: offs_num_def) lemma offs_next_prev: "(0 < offs_num (length ns) xs index key mi ma i \ offs_set_next ns xs index key mi ma i \ {} \ Min (offs_set_next ns xs index key mi ma i) = j) = (0 < offs_num (length ns) xs index key mi ma j \ offs_set_prev ns xs index key mi ma j \ {} \ Max (offs_set_prev ns xs index key mi ma j) = i)" (is "?P = ?Q") proof (rule iffI, (erule_tac [!] conjE)+) let ?A = "offs_set_next ns xs index key mi ma i" let ?B = "offs_set_prev ns xs index key mi ma j" assume A: "0 < offs_num (length ns) xs index key mi ma i" and B: "?A \ {}" and C: "Min ?A = j" have "Min ?A \ ?A" using B by (rule_tac Min_in, simp) hence D: "j \ ?A" using C by simp hence E: "i \ ?B" using A by simp show ?Q proof (rule conjI, rule_tac [2] conjI) show "0 < offs_num (length ns) xs index key mi ma j" using D by simp next show "?B \ {}" using E by blast next from E show "Max ?B = i" proof (subst Max_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI, (erule_tac conjE)+, rule_tac ccontr, simp) fix k assume F: "k < j" and "j < length ns" hence "k < length ns" by simp moreover assume "\ k \ i" and "0 < offs_num (length ns) xs index key mi ma k" ultimately have "k \ ?A" by simp hence "Min ?A \ k" by (rule_tac Min_le, simp) thus False using C and F by simp qed qed next let ?A = "offs_set_prev ns xs index key mi ma j" let ?B = "offs_set_next ns xs index key mi ma i" assume A: "0 < offs_num (length ns) xs index key mi ma j" and B: "?A \ {}" and C: "Max ?A = i" have "Max ?A \ ?A" using B by (rule_tac Max_in, simp) hence D: "i \ ?A" using C by simp hence E: "j \ ?B" using A by simp show ?P proof (rule conjI, rule_tac [2] conjI) show "0 < offs_num (length ns) xs index key mi ma i" using D by simp next show "?B \ {}" using E by blast next from E show "Min ?B = j" proof (subst Min_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI, (erule_tac conjE)+, rule_tac ccontr, simp) fix k assume "j < length ns" and "\ j \ k" and "0 < offs_num (length ns) xs index key mi ma k" hence "k \ ?A" by simp hence "k \ Max ?A" by (rule_tac Max_ge, simp) moreover assume "i < k" ultimately show False using C by simp qed qed qed lemma offs_next_cons_eq: assumes A: "index key x (length ns) mi ma = i" and B: "i < length ns" and C: "0 < offs_num (length ns) (x # xs) index key mi ma j" shows "offs_set_prev ns (x # xs) index key mi ma i = {} \ Max (offs_set_prev ns (x # xs) index key mi ma i) \ j \ offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j = offs_next ns ub (x # xs) index key mi ma j" (is "?P \ ?Q \ _") proof (simp only: disj_imp, cases "j < i") let ?A = "offs_set_prev ns (x # xs) index key mi ma i" let ?B = "offs_set_next (ns[i := Suc (ns ! i)]) xs index key mi ma j" let ?C = "offs_set_next ns (x # xs) index key mi ma j" case True hence D: "j \ ?A" using B and C by simp hence "j \ Max ?A" by (rule_tac Max_ge, simp) moreover assume "\ ?P \ ?Q" hence ?Q using D by blast ultimately have E: "j < Max ?A" by simp have F: "Max ?A \ ?A" using D by (rule_tac Max_in, simp, blast) have G: "Max ?A \ ?B" proof (simp, rule conjI, rule_tac [2] conjI) show "Max ?A < length ns" using F by auto next show "j < Max ?A" using E . next have "0 < offs_num (length ns) (x # xs) index key mi ma (Max ?A)" using F by blast moreover have "Max ?A < i" using F by blast ultimately show "0 < offs_num (length ns) xs index key mi ma (Max ?A)" using A by (simp add: offs_num_cons) qed hence H: "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j = ns[i := Suc (ns ! i)] ! Min ?B" by (simp only: offs_next_def split: if_split, blast) have "Min ?B \ Max ?A" using G by (rule_tac Min_le, simp) moreover have "Max ?A < i" using F by blast ultimately have I: "Min ?B < i" by simp hence J: "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j = ns ! Min ?B" using H by simp have "Min ?B \ ?B" using G by (rule_tac Min_in, simp, blast) hence K: "Min ?B \ ?C" using A and I by (simp add: offs_num_cons) hence L: "Min ?C \ Min ?B" by (rule_tac Min_le, simp) have "Min ?C \ ?C" using K by (rule_tac Min_in, simp, blast) moreover have "Min ?C < i" using L and I by simp ultimately have "Min ?C \ ?B" using A by (simp add: offs_num_cons) hence "Min ?B \ Min ?C" using G by (rule_tac Min_le, simp) hence "Min ?B = Min ?C" using L by simp moreover have "offs_next ns ub (x # xs) index key mi ma j = ns ! Min ?C" using K by (simp only: offs_next_def split: if_split, blast) ultimately show ?thesis using J by simp next let ?A = "offs_set_next (ns[i := Suc (ns ! i)]) xs index key mi ma j" let ?B = "offs_set_next ns (x # xs) index key mi ma j" case False hence "?A = ?B" using A by (rule_tac set_eqI, simp add: offs_num_cons) thus ?thesis proof (simp only: offs_next_def split: if_split, (rule_tac conjI, blast, rule_tac impI)+) assume "?B \ {}" hence "Min ?B \ ?B" by (rule_tac Min_in, simp) hence "i < Min ?B" using False by simp thus "ns[i := Suc (ns ! i)] ! Min ?B = ns ! Min ?B" by simp qed qed lemma offs_next_cons_neq: assumes A: "index key x (length ns) mi ma = i" and B: "offs_set_prev ns (x # xs) index key mi ma i \ {}" and C: "Max (offs_set_prev ns (x # xs) index key mi ma i) = j" shows "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j = (if 0 < offs_num (length ns) xs index key mi ma i then Suc (ns ! i) else offs_next ns ub (x # xs) index key mi ma i)" proof (simp, rule conjI, rule_tac [!] impI) let ?A = "offs_set_next ns (x # xs) index key mi ma j" assume "0 < offs_num (length ns) xs index key mi ma i" with A have "offs_set_next (ns[i := Suc (ns ! i)]) xs index key mi ma j = ?A" by (rule_tac set_eqI, rule_tac iffI, simp_all add: offs_num_cons split: if_split_asm) moreover have "0 < offs_num (length ns) (x # xs) index key mi ma i" using A by (simp add: offs_num_def) hence "0 < offs_num (length ns) (x # xs) index key mi ma j \ ?A \ {} \ Min ?A = i" using B and C by (subst offs_next_prev, simp) ultimately show "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j = Suc (ns ! i)" using B by (simp only: offs_next_def, simp, subst nth_list_update_eq, blast, simp) next let ?A = "offs_set_prev ns (x # xs) index key mi ma i" assume "offs_num (length ns) xs index key mi ma i = 0" with A have "offs_set_next (ns[i := Suc (ns ! i)]) xs index key mi ma j = offs_set_next ns (x # xs) index key mi ma i" proof (rule_tac set_eqI, rule_tac iffI, simp_all add: offs_num_cons split: if_split_asm, rule_tac conjI, rule_tac notI, simp, rule_tac impI, (erule_tac [!] conjE)+, rule_tac ccontr, simp_all) fix k have "i < length ns" using B by blast moreover assume "i \ k" and "\ i < k" hence "k < i" by simp moreover assume "0 < offs_num (length ns) xs index key mi ma k" ultimately have "k \ ?A" by (simp add: offs_num_cons) hence "k \ Max ?A" by (rule_tac Max_ge, simp) moreover assume "j < k" ultimately show False using C by simp next fix k have "Max ?A \ ?A" using B by (rule_tac Max_in, simp_all) hence "j < i" using C by simp moreover assume "i < k" ultimately show "j < k" by simp qed with B show "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j = offs_next ns ub (x # xs) index key mi ma i" proof (simp only: offs_next_def split: if_split, (rule_tac conjI, rule_tac [!] impI, simp)+, subst nth_list_update, blast, simp (no_asm_simp)) assume "offs_set_next ns (x # xs) index key mi ma i \ {}" hence "Min (offs_set_next ns (x # xs) index key mi ma i) \ offs_set_next ns (x # xs) index key mi ma i" by (rule_tac Min_in, simp) thus "i \ Min (offs_set_next ns (x # xs) index key mi ma i)" by simp qed qed lemma offs_pred_ub_aux [rule_format]: assumes A: "offs_pred ns ub xs index key mi ma" shows "i < length ns \ \j < length ns. i \ j \ 0 < offs_num (length ns) xs index key mi ma j \ ns ! j + offs_num (length ns) xs index key mi ma j \ ub" proof (erule strict_inc_induct, rule_tac [!] allI, (rule_tac [!] impI)+, drule le_less_Suc_eq, simp_all) fix i assume B: "length ns = Suc i" hence "offs_num (Suc i) xs index key mi ma i \ offs_next ns ub xs index key mi ma i - ns ! i" using A by (simp add: offs_pred_def) moreover have "offs_next ns ub xs index key mi ma i = ub" using B by (simp add: offs_next_def) ultimately have "offs_num (Suc i) xs index key mi ma i \ ub - ns ! i" by simp moreover assume "0 < offs_num (Suc i) xs index key mi ma i" ultimately show "ns ! i + offs_num (Suc i) xs index key mi ma i \ ub" by simp next fix i j assume "j < length ns" hence "offs_num (length ns) xs index key mi ma j \ offs_next ns ub xs index key mi ma j - ns ! j" using A by (simp add: offs_pred_def) moreover assume B: "\k < length ns. Suc i \ k \ 0 < offs_num (length ns) xs index key mi ma k \ ns ! k + offs_num (length ns) xs index key mi ma k \ ub" and C: "i \ j" have "offs_next ns ub xs index key mi ma j \ ub" proof (simp only: offs_next_def split: if_split, rule conjI, simp, rule impI) let ?A = "offs_set_next ns xs index key mi ma j" assume "?A \ {}" hence "Min ?A \ ?A" by (rule_tac Min_in, simp) hence "ns ! Min ?A + offs_num (length ns) xs index key mi ma (Min ?A) \ ub" using B and C by simp thus "ns ! Min ?A \ ub" by simp qed ultimately have "offs_num (length ns) xs index key mi ma j \ ub - ns ! j" by simp moreover assume "0 < offs_num (length ns) xs index key mi ma j" ultimately show "ns ! j + offs_num (length ns) xs index key mi ma j \ ub" by simp qed lemma offs_pred_ub: "\offs_pred ns ub xs index key mi ma; i < length ns; 0 < offs_num (length ns) xs index key mi ma i\ \ ns ! i + offs_num (length ns) xs index key mi ma i \ ub" by (drule offs_pred_ub_aux, assumption+, simp) lemma offs_pred_asc_aux [rule_format]: assumes A: "offs_pred ns ub xs index key mi ma" shows "i < length ns \ \j k. k < length ns \ i \ j \ j < k \ 0 < offs_num (length ns) xs index key mi ma j \ 0 < offs_num (length ns) xs index key mi ma k \ ns ! j + offs_num (length ns) xs index key mi ma j \ ns ! k" proof (erule strict_inc_induct, simp, (rule allI)+, (rule impI)+, simp) fix i j k assume B: "k < length ns" and C: "j < k" hence "j < length ns" by simp hence "offs_num (length ns) xs index key mi ma j \ offs_next ns ub xs index key mi ma j - ns ! j" using A by (simp add: offs_pred_def) moreover assume D: "\j k. k < length ns \ Suc i \ j \ j < k \ 0 < offs_num (length ns) xs index key mi ma j \ 0 < offs_num (length ns) xs index key mi ma k \ ns ! j + offs_num (length ns) xs index key mi ma j \ ns ! k" and E: "i \ j" and F: "0 < offs_num (length ns) xs index key mi ma k" have "offs_next ns ub xs index key mi ma j \ ns ! k" proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI, rule ccontr, simp) assume "\n > j. n < length ns \ offs_num (length ns) xs index key mi ma n = 0" hence "offs_num (length ns) xs index key mi ma k = 0" using B and C by simp thus False using F by simp next let ?A = "offs_set_next ns xs index key mi ma j" have G: "k \ ?A" using B and C and F by simp hence "Min ?A \ k" by (rule_tac Min_le, simp) hence "Min ?A < k \ Min ?A = k" by (simp add: le_less) moreover { have "Suc i \ Min ?A \ Min ?A < k \ 0 < offs_num (length ns) xs index key mi ma (Min ?A) \ ns ! Min ?A + offs_num (length ns) xs index key mi ma (Min ?A) \ ns ! k" using B and D and F by simp moreover assume "Min ?A < k" moreover have "Min ?A \ ?A" using G by (rule_tac Min_in, simp, blast) ultimately have "ns ! Min ?A < ns ! k" using E by simp } moreover { assume "Min ?A = k" hence "ns ! Min ?A = ns ! k" by simp } ultimately show "ns ! Min ?A \ ns ! k" by (simp add: le_less, blast) qed ultimately have "offs_num (length ns) xs index key mi ma j \ ns ! k - ns ! j" by simp moreover assume "0 < offs_num (length ns) xs index key mi ma j" ultimately show "ns ! j + offs_num (length ns) xs index key mi ma j \ ns ! k" by simp qed lemma offs_pred_asc: "\offs_pred ns ub xs index key mi ma; i < j; j < length ns; 0 < offs_num (length ns) xs index key mi ma i; 0 < offs_num (length ns) xs index key mi ma j\ \ ns ! i + offs_num (length ns) xs index key mi ma i \ ns ! j" by (drule offs_pred_asc_aux, erule less_trans, assumption+, rule order_refl) lemma offs_pred_next: assumes A: "offs_pred ns ub xs index key mi ma" and B: "i < length ns" and C: "0 < offs_num (length ns) xs index key mi ma i" shows "ns ! i < offs_next ns ub xs index key mi ma i" proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI) have "ns ! i + offs_num (length ns) xs index key mi ma i \ ub" using A and B and C by (rule offs_pred_ub) thus "ns ! i < ub" using C by simp next assume "offs_set_next ns xs index key mi ma i \ {}" hence "Min (offs_set_next ns xs index key mi ma i) \ offs_set_next ns xs index key mi ma i" by (rule_tac Min_in, simp) hence "ns ! i + offs_num (length ns) xs index key mi ma i \ ns ! Min (offs_set_next ns xs index key mi ma i)" using C by (rule_tac offs_pred_asc [OF A], simp_all) thus "ns ! i < ns ! Min (offs_set_next ns xs index key mi ma i)" using C by simp qed lemma offs_pred_next_cons_less: assumes A: "offs_pred ns ub (x # xs) index key mi ma" and B: "index key x (length ns) mi ma = i" and C: "offs_set_prev ns (x # xs) index key mi ma i \ {}" and D: "Max (offs_set_prev ns (x # xs) index key mi ma i) = j" shows "offs_next ns ub (x # xs) index key mi ma j < offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j" (is "?M < ?N") proof - have E: "0 < offs_num (length ns) (x # xs) index key mi ma i" using B by (simp add: offs_num_cons) hence "0 < offs_num (length ns) (x # xs) index key mi ma j \ offs_set_next ns (x # xs) index key mi ma j \ {} \ Min (offs_set_next ns (x # xs) index key mi ma j) = i" using C and D by (subst offs_next_prev, simp) hence F: "?M = ns ! i" by (simp only: offs_next_def, simp) have "?N = (if 0 < offs_num (length ns) xs index key mi ma i then Suc (ns ! i) else offs_next ns ub (x # xs) index key mi ma i)" using B and C and D by (rule offs_next_cons_neq) thus ?thesis proof (split if_split_asm) assume "?N = Suc (ns ! i)" thus ?thesis using F by simp next assume "?N = offs_next ns ub (x # xs) index key mi ma i" moreover have "ns ! i < offs_next ns ub (x # xs) index key mi ma i" using C by (rule_tac offs_pred_next [OF A _ E], blast) ultimately show ?thesis using F by simp qed qed lemma offs_pred_next_cons: assumes A: "offs_pred ns ub (x # xs) index key mi ma" and B: "index key x (length ns) mi ma = i" and C: "i < length ns" and D: "0 < offs_num (length ns) (x # xs) index key mi ma j" shows "offs_next ns ub (x # xs) index key mi ma j \ offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j" (is "?M \ ?N") proof - let ?P = "offs_set_prev ns (x # xs) index key mi ma i \ {} \ Max (offs_set_prev ns (x # xs) index key mi ma i) = j" have "?P \ \ ?P" by blast moreover { assume ?P hence "?M < ?N" using B by (rule_tac offs_pred_next_cons_less [OF A], simp_all) hence ?thesis by simp } moreover { assume "\ ?P" hence "?N = ?M" by (rule_tac offs_next_cons_eq [OF B C D], blast) hence ?thesis by simp } ultimately show ?thesis .. qed lemma offs_pred_cons: assumes A: "offs_pred ns ub (x # xs) index key mi ma" and B: "index key x (length ns) mi ma = i" and C: "i < length ns" shows "offs_pred (ns[i := Suc (ns ! i)]) ub xs index key mi ma" using A proof (simp add: offs_pred_def, rule_tac allI, rule_tac impI, simp) let ?ns' = "ns[i := Suc (ns ! i)]" fix j assume "\j < length ns. offs_num (length ns) (x # xs) index key mi ma j \ offs_next ns ub (x # xs) index key mi ma j - ns ! j" and "j < length ns" hence D: "offs_num (length ns) (x # xs) index key mi ma j \ offs_next ns ub (x # xs) index key mi ma j - ns ! j" by simp from B and C show "offs_num (length ns) xs index key mi ma j \ offs_next ?ns' ub xs index key mi ma j - ?ns' ! j" proof (cases "j = i", case_tac [2] "0 < offs_num (length ns) xs index key mi ma j", simp_all) assume "j = i" hence "offs_num (length ns) xs index key mi ma i \ offs_next ns ub (x # xs) index key mi ma i - Suc (ns ! i)" using B and D by (simp add: offs_num_cons) moreover have "offs_next ns ub (x # xs) index key mi ma i \ offs_next ?ns' ub xs index key mi ma i" using B by (rule_tac offs_pred_next_cons [OF A _ C], simp_all add: offs_num_cons) ultimately show "offs_num (length ns) xs index key mi ma i \ offs_next ?ns' ub xs index key mi ma i - Suc (ns ! i)" by simp next assume "j \ i" hence "offs_num (length ns) xs index key mi ma j \ offs_next ns ub (x # xs) index key mi ma j - ns ! j" using B and D by (simp add: offs_num_cons) moreover assume "0 < offs_num (length ns) xs index key mi ma j" hence "offs_next ns ub (x # xs) index key mi ma j \ offs_next ?ns' ub xs index key mi ma j" using B by (rule_tac offs_pred_next_cons [OF A _ C], simp_all add: offs_num_cons) ultimately show "offs_num (length ns) xs index key mi ma j \ offs_next ?ns' ub xs index key mi ma j - ns ! j" by simp qed qed text \ \null The next step consists of proving, as done in lemma @{text fill_count_item} in what follows, that if certain conditions hold, particularly if offsets' list $ns$ and objects' list $xs$ satisfy predicate @{const offs_pred}, then function @{const fill} conserves objects if called using $xs$ and $ns$ as its input arguments. This lemma is proven by induction on $xs$. Hence, lemma @{thm [source] offs_pred_cons}, proven in the previous step, is used to remove the antecedent containing predicate @{const offs_pred} from the induction hypothesis, which has the form of an implication. \null \ lemma offs_next_zero: assumes A: "i < length ns" and B: "offs_num (length ns) xs index key mi ma i = 0" and C: "offs_set_prev ns xs index key mi ma i = {}" shows "offs_next ns ub xs index key mi ma 0 = offs_next ns ub xs index key mi ma i" proof - have "offs_set_next ns xs index key mi ma 0 = offs_set_next ns xs index key mi ma i" proof (rule set_eqI, rule iffI, simp_all, (erule conjE)+, rule ccontr, simp add: not_less) fix j assume D: "0 < offs_num (length ns) xs index key mi ma j" assume "j \ i" hence "j < i \ j = i" by (simp add: le_less) moreover { assume "j < i" hence "j \ offs_set_prev ns xs index key mi ma i" using A and D by simp hence False using C by simp } moreover { assume "j = i" hence False using B and D by simp } ultimately show False .. qed thus ?thesis by (simp only: offs_next_def) qed lemma offs_next_zero_cons_eq: assumes A: "index key x (length ns) mi ma = i" and B: "offs_num (length ns) (x # xs) index key mi ma 0 = 0" and C: "offs_set_prev ns (x # xs) index key mi ma i \ {}" (is "?A \ _") shows "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma 0 = offs_next ns ub (x # xs) index key mi ma 0" proof - have D: "Min ?A \ ?A" using C by (rule_tac Min_in, simp) moreover have E: "0 < Min ?A" proof (rule ccontr, simp) assume "Min ?A = 0" hence "offs_num (length ns) (x # xs) index key mi ma (Min ?A) = 0" using B by simp moreover have "0 < offs_num (length ns) (x # xs) index key mi ma (Min ?A)" using D by auto ultimately show False by simp qed ultimately have "Min ?A \ offs_set_next ns (x # xs) index key mi ma 0" (is "_ \ ?B") by auto moreover from this have "Min ?B = Min ?A" proof (subst Min_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI, (erule_tac conjE)+, rule_tac ccontr, simp add: not_le) fix j assume F: "j < Min ?A" have "i < length ns" using D by simp moreover have "Min ?A < i" using D by auto hence "j < i" using F by simp moreover assume "0 < offs_num (length ns) (x # xs) index key mi ma j" ultimately have "j \ ?A" by simp hence "Min ?A \ j" by (rule_tac Min_le, simp) thus False using F by simp qed moreover have "Min ?A \ offs_set_next (ns[i := Suc (ns ! i)]) xs index key mi ma 0" (is "_ \ ?C") using D and E by (auto, simp add: offs_num_cons A [symmetric]) moreover from this have "Min ?C = Min ?A" proof (subst Min_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI, (erule_tac conjE)+, rule_tac ccontr, simp add: not_le) fix j assume F: "j < Min ?A" have "i < length ns" using D by simp moreover have "Min ?A < i" using D by auto hence "j < i" using F by simp moreover assume "0 < offs_num (length ns) xs index key mi ma j" hence "0 < offs_num (length ns) (x # xs) index key mi ma j" by (simp add: offs_num_cons) ultimately have "j \ ?A" by simp hence "Min ?A \ j" by (rule_tac Min_le, simp) thus False using F by simp qed moreover have "Min ?A < i" using D by auto ultimately show ?thesis by (simp only: offs_next_def split: if_split, (rule_tac conjI, blast, rule_tac impI)+, simp) qed lemma offs_next_zero_cons_neq: assumes A: "index key x (length ns) mi ma = i" and B: "i < length ns" and C: "0 < i" and D: "offs_set_prev ns (x # xs) index key mi ma i = {}" shows "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma 0 = (if 0 < offs_num (length ns) xs index key mi ma i then Suc (ns ! i) else offs_next ns ub (x # xs) index key mi ma i)" proof (simp, rule conjI, rule_tac [!] impI) let ?ns' = "ns[i := Suc (ns ! i)]" assume "0 < offs_num (length ns) xs index key mi ma i" with A have "offs_set_next ?ns' xs index key mi ma 0 = offs_set_next ns (x # xs) index key mi ma 0" by (rule_tac set_eqI, rule_tac iffI, simp_all add: offs_num_cons split: if_split_asm) moreover have "i \ offs_set_next ns (x # xs) index key mi ma 0" using A and B and C by (simp add: offs_num_cons) moreover from this have "Min (offs_set_next ns (x # xs) index key mi ma 0) = i" proof (subst Min_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI, (erule_tac conjE)+, rule_tac ccontr, simp add: not_le) fix j assume "j < i" and "0 < offs_num (length ns) (x # xs) index key mi ma j" hence "j \ offs_set_prev ns (x # xs) index key mi ma i" using B by simp thus False using D by simp qed ultimately show "offs_next ?ns' ub xs index key mi ma 0 = Suc (ns ! i)" by (simp only: offs_next_def split: if_split, rule_tac conjI, rule_tac [!] impI, simp_all) next let ?ns' = "ns[i := Suc (ns ! i)]" assume "offs_num (length ns) xs index key mi ma i = 0" moreover have "offs_set_prev ?ns' xs index key mi ma i = {}" using D by (simp add: offs_num_cons split: if_split_asm, blast) ultimately have "offs_next ?ns' ub xs index key mi ma 0 = offs_next ?ns' ub xs index key mi ma i" using B by (rule_tac offs_next_zero, simp_all) moreover have "offs_next ?ns' ub xs index key mi ma i = offs_next ns ub (x # xs) index key mi ma i" using A and B and D by (rule_tac offs_next_cons_eq, simp_all add: offs_num_cons) ultimately show "offs_next ?ns' ub xs index key mi ma 0 = offs_next ns ub (x # xs) index key mi ma i" by simp qed lemma offs_pred_zero_cons_less: assumes A: "offs_pred ns ub (x # xs) index key mi ma" and B: "index key x (length ns) mi ma = i" and C: "i < length ns" and D: "0 < i" and E: "offs_set_prev ns (x # xs) index key mi ma i = {}" shows "offs_next ns ub (x # xs) index key mi ma 0 < offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma 0" (is "?M < ?N") proof - have "i \ offs_set_next ns (x # xs) index key mi ma 0" using B and C and D by (simp add: offs_num_cons) moreover from this have "Min (offs_set_next ns (x # xs) index key mi ma 0) = i" proof (subst Min_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI, (erule_tac conjE)+, rule_tac ccontr, simp add: not_le) fix j assume "j < i" and "0 < offs_num (length ns) (x # xs) index key mi ma j" hence "j \ offs_set_prev ns (x # xs) index key mi ma i" using C by simp thus False using E by simp qed ultimately have F: "?M = ns ! i" by (simp only: offs_next_def split: if_split, blast) have "?N = (if 0 < offs_num (length ns) xs index key mi ma i then Suc (ns ! i) else offs_next ns ub (x # xs) index key mi ma i)" using B and C and D and E by (rule offs_next_zero_cons_neq) thus ?thesis proof (split if_split_asm) assume "?N = Suc (ns ! i)" thus ?thesis using F by simp next assume "?N = offs_next ns ub (x # xs) index key mi ma i" moreover have "ns ! i < offs_next ns ub (x # xs) index key mi ma i" using B by (rule_tac offs_pred_next [OF A C], simp add: offs_num_cons) ultimately show ?thesis using F by simp qed qed lemma offs_pred_zero_cons: assumes A: "offs_pred ns ub (x # xs) index key mi ma" and B: "index key x (length ns) mi ma = i" and C: "i < length ns" and D: "offs_num (length ns) (x # xs) index key mi ma 0 = 0" shows "offs_next ns ub (x # xs) index key mi ma 0 \ offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma 0" (is "?M \ ?N") proof (cases "offs_set_prev ns (x # xs) index key mi ma i = {}") case True have "0 < i" using B and D by (rule_tac ccontr, simp add: offs_num_cons) hence "?M < ?N" using True and B by (rule_tac offs_pred_zero_cons_less [OF A _ C]) thus ?thesis by simp next case False hence "?N = ?M" by (rule offs_next_zero_cons_eq [OF B D]) thus ?thesis by simp qed lemma replicate_count: "count (mset (replicate n x)) x = n" by (induction n, simp_all) lemma fill_none [rule_format]: assumes A: "index_less index key" shows "(\x \ set xs. key x \ {mi..ma}) \ ns \ [] \ offs_pred ns ub xs index key mi ma \ offs_none ns ub xs index key mi ma i \ fill xs ns index key ub mi ma ! i = None" proof (induction xs arbitrary: ns, simp add: offs_none_def offs_num_def offs_next_def, (rule impI)+, simp add: Let_def, (erule conjE)+) fix x xs and ns :: "nat list" let ?i' = "index key x (length ns) mi ma" let ?ns' = "ns[?i' := Suc (ns ! ?i')]" assume B: "offs_pred ns ub (x # xs) index key mi ma" and C: "offs_none ns ub (x # xs) index key mi ma i" assume D: "ns \ []" and "mi \ key x" and "key x \ ma" moreover from this have E: "?i' < length ns" using A by (simp add: index_less_def) hence "offs_pred ?ns' ub xs index key mi ma" by (rule_tac offs_pred_cons [OF B], simp) moreover assume "\ns. ns \ [] \ offs_pred ns ub xs index key mi ma \ offs_none ns ub xs index key mi ma i \ fill xs ns index key ub mi ma ! i = None" ultimately have F: "offs_none ?ns' ub xs index key mi ma i \ fill xs ?ns' index key ub mi ma ! i = None" by simp show "(fill xs ?ns' index key ub mi ma)[ns ! ?i' := Some x] ! i = None" - proof (insert C, simp add: offs_none_def, erule disjE, erule_tac [2] disjE, simp_all del: subst_all subst_all' + proof (insert C, simp add: offs_none_def, erule disjE, erule_tac [2] disjE, simp_all del: subst_all add: offs_num_cons split: if_split_asm, erule conjE, rule case_split, drule mp, - assumption, simp_all del: subst_all subst_all', (erule conjE)+, (erule_tac [2] conjE)+, + assumption, simp_all del: subst_all, (erule conjE)+, (erule_tac [2] conjE)+, erule_tac [3] conjE, erule_tac [5] conjE) fix j assume G: "?i' = j" and H: "j < length ns" and I: "Suc (ns ! j + offs_num (length ns) xs index key mi ma j) \ i" and J: "i < offs_next ns ub (x # xs) index key mi ma j" show "fill xs (ns[j := Suc (ns ! j)]) index key ub mi ma ! i = None" proof (cases "0 < offs_num (length ns) xs index key mi ma j", case_tac [2] "offs_set_prev ns (x # xs) index key mi ma j \ {}", simp_all only: not_not not_gr0) have "j < length ns \ 0 < offs_num (length ns) xs index key mi ma j \ ?ns' ! j + offs_num (length ns) xs index key mi ma j \ i \ i < offs_next ?ns' ub xs index key mi ma j \ fill xs ?ns' index key ub mi ma ! i = None" using F by (simp add: offs_none_def, blast) moreover assume "0 < offs_num (length ns) xs index key mi ma j" moreover have "?ns' ! j + offs_num (length ns) xs index key mi ma j \ i" using G and H and I by simp moreover have "0 < offs_num (length ns) (x # xs) index key mi ma j" using G by (simp add: offs_num_cons) hence "offs_next ns ub (x # xs) index key mi ma j \ offs_next ?ns' ub xs index key mi ma j" using G and H by (rule_tac offs_pred_next_cons [OF B], simp_all) hence "i < offs_next ?ns' ub xs index key mi ma j" using J by simp ultimately have "fill xs ?ns' index key ub mi ma ! i = None" using H by blast thus ?thesis using G by simp next let ?j' = "Max (offs_set_prev ns (x # xs) index key mi ma j)" have "?j' < length ns \ 0 < offs_num (length ns) xs index key mi ma ?j' \ ?ns' ! ?j' + offs_num (length ns) xs index key mi ma ?j' \ i \ i < offs_next ?ns' ub xs index key mi ma ?j' \ fill xs ?ns' index key ub mi ma ! i = None" using F by (simp add: offs_none_def, blast) moreover assume K: "offs_set_prev ns (x # xs) index key mi ma j \ {}" hence "?j' \ offs_set_prev ns (x # xs) index key mi ma j" by (rule_tac Max_in, simp) hence L: "?j' < length ns \ ?j' < j \ 0 < offs_num (length ns) xs index key mi ma ?j'" using G by (auto, subst (asm) (2) offs_num_cons, simp) moreover have "ns ! ?j' + offs_num (length ns) (x # xs) index key mi ma ?j' \ ns ! j" using G and H and L by (rule_tac offs_pred_asc [OF B], simp_all add: offs_num_cons) hence "?ns' ! ?j' + offs_num (length ns) xs index key mi ma ?j' \ ns ! j" using G and H and L by (subst nth_list_update, simp_all add: offs_num_cons) hence "?ns' ! ?j' + offs_num (length ns) xs index key mi ma ?j' \ i" using I by simp moreover assume M: "offs_num (length ns) xs index key mi ma j = 0" have "offs_next (ns[j := Suc (ns ! j)]) ub xs index key mi ma ?j' = (if 0 < offs_num (length ns) xs index key mi ma j then Suc (ns ! j) else offs_next ns ub (x # xs) index key mi ma j)" using G and K by (rule_tac offs_next_cons_neq, simp_all) hence "offs_next ?ns' ub xs index key mi ma ?j' = offs_next ns ub (x # xs) index key mi ma j" using G and M by simp hence "i < offs_next ?ns' ub xs index key mi ma ?j'" using J by simp ultimately have "fill xs ?ns' index key ub mi ma ! i = None" by blast thus ?thesis using G by simp next have "offs_num (length ns) xs index key mi ma 0 = 0 \ i < offs_next ?ns' ub xs index key mi ma 0 \ fill xs ?ns' index key ub mi ma ! i = None" using F by (simp add: offs_none_def) moreover assume K: "offs_set_prev ns (x # xs) index key mi ma j = {}" and L: "offs_num (length ns) xs index key mi ma j = 0" have "offs_set_prev ns (x # xs) index key mi ma j = offs_set_prev ?ns' xs index key mi ma j" using G by (rule_tac set_eqI, rule_tac iffI, simp_all add: offs_num_cons split: if_split_asm) hence M: "offs_set_prev ?ns' xs index key mi ma j = {}" using K by simp hence "offs_num (length ns) xs index key mi ma 0 = 0" using H and L by (cases j, simp_all) moreover have N: "offs_next ?ns' ub xs index key mi ma 0 = offs_next ?ns' ub xs index key mi ma j" using H and L and M by (rule_tac offs_next_zero, simp_all) have "offs_next ?ns' ub xs index key mi ma j = offs_next ns ub (x # xs) index key mi ma j" using G and H and K by (subst offs_next_cons_eq, simp_all add: offs_num_cons) hence "i < offs_next ?ns' ub xs index key mi ma 0" using J and N by simp ultimately have "fill xs ?ns' index key ub mi ma ! i = None" by blast thus ?thesis using G by simp qed next fix j assume G: "?i' \ j" and H: "j < length ns" and I: "ns ! j + offs_num (length ns) xs index key mi ma j \ i" and J: "i < offs_next ns ub (x # xs) index key mi ma j" and K: "0 < offs_num (length ns) xs index key mi ma j" from G have "ns ! ?i' \ i" proof (rule_tac notI, cases "?i' < j", simp_all add: not_less le_less) assume "?i' < j" hence "ns ! ?i' + offs_num (length ns) (x # xs) index key mi ma ?i' \ ns ! j" using H and K by (rule_tac offs_pred_asc [OF B], simp_all add: offs_num_cons) moreover assume "ns ! ?i' = i" ultimately have "i < ns ! j" by (simp add: offs_num_cons) thus False using I by simp next assume "j < ?i'" hence L: "?i' \ offs_set_next ns (x # xs) index key mi ma j" (is "_ \ ?A") using E by (simp add: offs_num_cons) hence "Min ?A \ ?i'" by (rule_tac Min_le, simp) hence "Min ?A < ?i' \ Min ?A = ?i'" by (simp add: le_less) hence "ns ! Min ?A \ ns ! ?i'" proof (rule disjE, simp_all) assume "Min ?A < ?i'" moreover have "Min ?A \ ?A" using L by (rule_tac Min_in, simp, blast) hence "0 < offs_num (length ns) (x # xs) index key mi ma (Min ?A)" by simp ultimately have "ns ! Min ?A + offs_num (length ns) (x # xs) index key mi ma (Min ?A) \ ns ! ?i'" using E by (rule_tac offs_pred_asc [OF B], simp_all add: offs_num_cons) thus ?thesis by simp qed hence "offs_next ns ub (x # xs) index key mi ma j \ ns ! ?i'" using L by (simp only: offs_next_def split: if_split, blast) moreover assume "ns ! ?i' = i" ultimately show False using J by simp qed thus "(fill xs ?ns' index key ub mi ma)[ns ! ?i' := Some x] ! i = None" proof simp have "j < length ns \ 0 < offs_num (length ns) xs index key mi ma j \ ?ns' ! j + offs_num (length ns) xs index key mi ma j \ i \ i < offs_next ?ns' ub xs index key mi ma j \ fill xs ?ns' index key ub mi ma ! i = None" using F by (simp add: offs_none_def, blast) moreover have "offs_next ns ub (x # xs) index key mi ma j \ offs_next ?ns' ub xs index key mi ma j" using E and K by (rule_tac offs_pred_next_cons [OF B], simp_all add: offs_num_cons) hence "i < offs_next ?ns' ub xs index key mi ma j" using J by simp ultimately show "fill xs ?ns' index key ub mi ma ! i = None" using G and H and I and K by simp qed next assume G: "0 < ?i'" and H: "offs_num (length ns) xs index key mi ma 0 = 0" and I: "i < offs_next ns ub (x # xs) index key mi ma 0" have "ns ! ?i' \ i" proof have "0 < offs_num (length ns) (x # xs) index key mi ma ?i'" by (simp add: offs_num_cons) hence L: "?i' \ offs_set_next ns (x # xs) index key mi ma 0" (is "_ \ ?A") using E and G by simp hence "Min ?A \ ?i'" by (rule_tac Min_le, simp) hence "Min ?A < ?i' \ Min ?A = ?i'" by (simp add: le_less) hence "ns ! Min ?A \ ns ! ?i'" proof (rule disjE, simp_all) assume "Min ?A < ?i'" moreover have "Min ?A \ ?A" using L by (rule_tac Min_in, simp, blast) hence "0 < offs_num (length ns) (x # xs) index key mi ma (Min ?A)" by simp ultimately have "ns ! Min ?A + offs_num (length ns) (x # xs) index key mi ma (Min ?A) \ ns ! ?i'" using E by (rule_tac offs_pred_asc [OF B], simp_all add: offs_num_cons) thus ?thesis by simp qed hence "offs_next ns ub (x # xs) index key mi ma 0 \ ns ! ?i'" using L by (simp only: offs_next_def split: if_split, blast) moreover assume "ns ! ?i' = i" ultimately show False using I by simp qed thus "(fill xs ?ns' index key ub mi ma)[ns ! ?i' := Some x] ! i = None" proof simp have "offs_num (length ns) xs index key mi ma 0 = 0 \ i < offs_next ?ns' ub xs index key mi ma 0 \ fill xs ?ns' index key ub mi ma ! i = None" using F by (simp add: offs_none_def) moreover have "offs_next ns ub (x # xs) index key mi ma 0 \ offs_next ?ns' ub xs index key mi ma 0" using E and G and H by (rule_tac offs_pred_zero_cons [OF B], simp_all add: offs_num_cons) hence "i < offs_next ?ns' ub xs index key mi ma 0" using I by simp ultimately show "fill xs ?ns' index key ub mi ma ! i = None" using H by simp qed next assume G: "?i' = 0" and H: "i < ns ! 0" show "fill xs (ns[0 := Suc (ns ! 0)]) index key ub mi ma ! i = None" proof (cases "0 < offs_num (length ns) xs index key mi ma 0", simp_all) have "0 < offs_num (length ns) xs index key mi ma 0 \ i < ?ns' ! 0 \ fill xs ?ns' index key ub mi ma ! i = None" using F by (simp add: offs_none_def) moreover assume "0 < offs_num (length ns) xs index key mi ma 0" moreover have "i < ?ns' ! 0" using D and G and H by simp ultimately show ?thesis using G by simp next have "offs_num (length ns) xs index key mi ma 0 = 0 \ i < offs_next ?ns' ub xs index key mi ma 0 \ fill xs ?ns' index key ub mi ma ! i = None" using F by (simp add: offs_none_def) moreover assume "offs_num (length ns) xs index key mi ma 0 = 0" moreover have I: "offs_next ?ns' ub xs index key mi ma 0 = offs_next ns ub (x # xs) index key mi ma 0" using D and G by (rule_tac offs_next_cons_eq, simp_all add: offs_num_cons) have "ns ! 0 < offs_next ns ub (x # xs) index key mi ma 0" using D and G by (rule_tac offs_pred_next [OF B], simp_all add: offs_num_cons) hence "i < offs_next ?ns' ub xs index key mi ma 0" using H and I by simp ultimately show ?thesis using G by simp qed next assume G: "0 < ?i'" and H: "0 < offs_num (length ns) xs index key mi ma 0" and I: "i < ns ! 0" have "ns ! ?i' \ i" proof - have "ns ! 0 + offs_num (length ns) (x # xs) index key mi ma 0 \ ns ! ?i'" using H by (rule_tac offs_pred_asc [OF B G E], simp_all add: offs_num_cons) moreover have "0 < offs_num (length ns) (x # xs) index key mi ma 0" using H by (simp add: offs_num_cons) ultimately show ?thesis using I by simp qed thus "(fill xs ?ns' index key ub mi ma)[ns ! ?i' := Some x] ! i = None" proof simp have "0 < offs_num (length ns) xs index key mi ma 0 \ i < ?ns' ! 0 \ fill xs ?ns' index key ub mi ma ! i = None" using F by (simp add: offs_none_def) moreover have "i < ?ns' ! 0" using G and I by simp ultimately show "fill xs ?ns' index key ub mi ma ! i = None" using H by simp qed qed qed lemma fill_index_none [rule_format]: assumes A: "index_less index key" and B: "key x \ {mi..ma}" and C: "ns \ []" and D: "offs_pred ns ub (x # xs) index key mi ma" shows "\x \ set xs. key x \ {mi..ma} \ fill xs (ns[(index key x (length ns) mi ma) := Suc (ns ! index key x (length ns) mi ma)]) index key ub mi ma ! (ns ! index key x (length ns) mi ma) = None" (is "_ \ fill _ ?ns' _ _ _ _ _ ! (_ ! ?i) = _") using A and B and C and D proof (rule_tac fill_none, simp_all, rule_tac offs_pred_cons, simp_all, simp add: index_less_def, cases "0 < ?i", cases "offs_set_prev ns (x # xs) index key mi ma ?i = {}", case_tac [3] "0 < offs_num (length ns) xs index key mi ma 0") assume E: "0 < ?i" and F: "offs_set_prev ns (x # xs) index key mi ma ?i = {}" have G: "?i < length ns" using A and B and C by (simp add: index_less_def) hence "offs_num (length ns) (x # xs) index key mi ma 0 = 0" using E and F by (rule_tac ccontr, simp) hence "offs_num (length ns) xs index key mi ma 0 = 0" by (simp add: offs_num_cons split: if_split_asm) moreover have "offs_next ?ns' ub xs index key mi ma 0 = (if 0 < offs_num (length ns) xs index key mi ma ?i then Suc (ns ! ?i) else offs_next ns ub (x # xs) index key mi ma ?i)" using E and F and G by (rule_tac offs_next_zero_cons_neq, simp_all) hence "ns ! ?i < offs_next ?ns' ub xs index key mi ma 0" by (simp split: if_split_asm, rule_tac offs_pred_next [OF D G], simp add: offs_num_cons) ultimately show "offs_none ?ns' ub xs index key mi ma (ns ! ?i)" by (simp add: offs_none_def) next assume E: "0 < ?i" and F: "offs_set_prev ns (x # xs) index key mi ma ?i \ {}" (is "?A \ _") have G: "?i < length ns" using A and B and C by (simp add: index_less_def) have H: "Max ?A \ ?A" using F by (rule_tac Max_in, simp) hence I: "Max ?A < ?i" by blast have "Max ?A < length ns" using H by auto moreover have "0 < offs_num (length ns) (x # xs) index key mi ma (Max ?A)" using H by auto hence "0 < offs_num (length ns) xs index key mi ma (Max ?A)" using I by (subst (asm) offs_num_cons, split if_split_asm, simp_all) moreover have "ns ! Max ?A + offs_num (length ns) (x # xs) index key mi ma (Max ?A) \ ns ! ?i" using G and H by (rule_tac offs_pred_asc [OF D], simp_all add: offs_num_cons) hence "?ns' ! Max ?A + offs_num (length ns) xs index key mi ma (Max ?A) \ ns ! ?i" using I by (simp add: offs_num_cons) moreover have "offs_next ?ns' ub xs index key mi ma (Max ?A) = (if 0 < offs_num (length ns) xs index key mi ma ?i then Suc (ns ! ?i) else offs_next ns ub (x # xs) index key mi ma ?i)" using F and I by (rule_tac offs_next_cons_neq, simp_all) hence "ns ! ?i < offs_next ?ns' ub xs index key mi ma (Max ?A)" by (simp split: if_split_asm, rule_tac offs_pred_next [OF D G], simp add: offs_num_cons) ultimately show "offs_none ?ns' ub xs index key mi ma (ns ! ?i)" by (simp add: offs_none_def, blast) next assume "0 < offs_num (length ns) xs index key mi ma 0" and "\ 0 < ?i" moreover have "?i < length ns" using A and B and C by (simp add: index_less_def) ultimately show "offs_none ?ns' ub xs index key mi ma (ns ! ?i)" by (simp add: offs_none_def) next assume E: "\ 0 < ?i" and F: "\ 0 < offs_num (length ns) xs index key mi ma 0" have G: "?i < length ns" using A and B and C by (simp add: index_less_def) have "offs_num (length ns) xs index key mi ma 0 = 0" using F by simp moreover have "offs_next ?ns' ub xs index key mi ma ?i = offs_next ns ub (x # xs) index key mi ma ?i" using E and G by (rule_tac offs_next_cons_eq, simp_all add: offs_num_cons) hence "ns ! ?i < offs_next ?ns' ub xs index key mi ma ?i" by (simp, rule_tac offs_pred_next [OF D G], simp add: offs_num_cons) hence "ns ! ?i < offs_next ?ns' ub xs index key mi ma 0" using E by simp ultimately show "offs_none ?ns' ub xs index key mi ma (ns ! ?i)" by (simp add: offs_none_def) qed lemma fill_count_item [rule_format]: assumes A: "index_less index key" shows "(\x \ set xs. key x \ {mi..ma}) \ ns \ [] \ offs_pred ns ub xs index key mi ma \ length xs \ ub \ count (mset (map the (fill xs ns index key ub mi ma))) x = count (mset xs) x + (if the None = x then ub - length xs else 0)" proof (induction xs arbitrary: ns, simp add: replicate_count, (rule impI)+, simp add: Let_def map_update del: count_add_mset mset_map split del: if_split, (erule conjE)+, subst add_mset_add_single, simp only: count_single count_union) fix y xs and ns :: "nat list" let ?i = "index key y (length ns) mi ma" let ?ns' = "ns[?i := Suc (ns ! ?i)]" assume B: "\x \ set xs. mi \ key x \ key x \ ma" and C: "mi \ key y" and D: "key y \ ma" and E: "ns \ []" and F: "offs_pred ns ub (y # xs) index key mi ma" and G: "Suc (length xs) \ ub" have H: "?i < length ns" using A and C and D and E by (simp add: index_less_def) assume "\ns. ns \ [] \ offs_pred ns ub xs index key mi ma \ count (mset (map the (fill xs ns index key ub mi ma))) x = count (mset xs) x + (if the None = x then ub - length xs else 0)" moreover have "offs_pred ?ns' ub xs index key mi ma" using F and H by (rule_tac offs_pred_cons, simp_all) ultimately have "count (mset (map the (fill xs ?ns' index key ub mi ma))) x = count (mset xs) x + (if the None = x then ub - length xs else 0)" using E by simp moreover have "ns ! ?i + offs_num (length ns) (y # xs) index key mi ma ?i \ ub" using F and H by (rule offs_pred_ub, simp add: offs_num_cons) hence "ns ! ?i < ub" by (simp add: offs_num_cons) ultimately show "count (mset ((map the (fill xs ?ns' index key ub mi ma)) [ns ! ?i := y])) x = count (mset xs) x + (if y = x then 1 else 0) + (if the None = x then ub - length (y # xs) else 0)" proof (subst mset_update, simp add: fill_length, subst add_mset_add_single, simp only: count_diff count_single count_union, subst nth_map, simp add: fill_length, subst add.assoc, subst (3) add.commute, subst add.assoc [symmetric], subst add_right_cancel) have "fill xs ?ns' index key ub mi ma ! (ns ! ?i) = None" using B and C and D and E by (rule_tac fill_index_none [OF A _ _ F], simp_all) thus "count (mset xs) x + (if the None = x then ub - length xs else 0) - (if the (fill xs ?ns' index key ub mi ma ! (ns ! ?i)) = x then 1 else 0) = count (mset xs) x + (if the None = x then ub - length (y # xs) else 0)" using G by simp qed qed text \ \null Finally, lemma @{text offs_enum_pred} here below proves that, if $ns$ is the offsets' list obtained by applying the composition of functions @{const offs} and @{const enum} to objects' list $xs$, then predicate @{const offs_pred} is satisfied by $ns$ and $xs$. This result is in turn used, together with lemma @{thm [source] fill_count_item}, to prove lemma @{text fill_offs_enum_count_item}, which states that function @{const fill} conserves objects if its input offsets' list is computed via the composition of functions @{const offs} and @{const enum}. \null \ lemma enum_offs_num: "i < n \ enum xs index key n mi ma ! i = offs_num n xs index key mi ma i" by (induction xs, simp add: offs_num_def, simp add: Let_def offs_num_cons, subst nth_list_update_eq, simp_all add: enum_length) lemma offs_length: "length (offs ns i) = length ns" by (induction ns arbitrary: i, simp_all) lemma offs_add [rule_format]: "i < length ns \ offs ns k ! i = foldl (+) k (take i ns)" by (induction ns arbitrary: i k, simp, simp add: nth_Cons split: nat.split) lemma offs_mono_aux: "i \ j \ j < length ns \ offs ns k ! i \ offs ns k ! (i + (j - i))" by (simp only: offs_add take_add, simp add: add_le) lemma offs_mono: "i \ j \ j < length ns \ offs ns k ! i \ offs ns k ! j" by (frule offs_mono_aux, simp_all) lemma offs_update: "j < length ns \ offs (ns[i := Suc (ns ! i)]) k ! j = (if j \ i then id else Suc) (offs ns k ! j)" by (simp add: offs_add not_le take_update_swap, rule impI, subst nth_take [symmetric], assumption, subst add_update, simp_all) lemma offs_equal_suc: assumes A: "Suc i < length ns" and B: "ns ! i = 0" shows "offs ns m ! i = offs ns m ! Suc i" proof - have "offs ns m ! i = foldl (+) m (take i ns)" using A by (subst offs_add, simp_all) also have "\ = foldl (+) m (take i ns @ [ns ! i])" using B by simp also have "\ = foldl (+) m (take (Suc i) ns)" using A by (subst take_Suc_conv_app_nth, simp_all) also have "\ = offs ns m ! Suc i" using A by (subst offs_add, simp_all) finally show ?thesis . qed lemma offs_equal [rule_format]: "i < j \ j < length ns \ (\k \ {i.. offs ns m ! i = offs ns m ! j" proof (erule strict_inc_induct, rule_tac [!] impI, simp_all, erule offs_equal_suc, simp) fix i assume A: "i < j" and "j < length ns" hence "Suc i < length ns" by simp moreover assume "\k \ {i.. = offs ns m ! j" finally show "offs ns m ! i = offs ns m ! j" . qed lemma offs_enum_last [rule_format]: assumes A: "index_less index key" and B: "0 < n" and C: "\x \ set xs. key x \ {mi..ma}" shows "offs (enum xs index key n mi ma) k ! (n - Suc 0) + offs_num n xs index key mi ma (n - Suc 0) = length xs + k" proof - let ?ns = "enum xs index key n mi ma" from B have D: "last ?ns = offs_num n xs index key mi ma (n - Suc 0)" by (subst last_conv_nth, subst length_0_conv [symmetric], simp_all add: enum_length, subst enum_offs_num, simp_all) have "offs ?ns k ! (n - Suc 0) = foldl (+) k (take (n - Suc 0) ?ns)" using B by (rule_tac offs_add, simp add: enum_length) also have "\ = foldl (+) k (butlast ?ns)" by (simp add: butlast_conv_take enum_length) finally have "offs ?ns k ! (n - Suc 0) + offs_num n xs index key mi ma (n - Suc 0) = foldl (+) k (butlast ?ns @ [last ?ns])" using D by simp also have "\ = foldl (+) k ?ns" using B by (subst append_butlast_last_id, subst length_0_conv [symmetric], simp_all add: enum_length) also have "\ = foldl (+) 0 ?ns + k" by (rule add_base_zero) also have "\ = length xs + k" using A and B and C by (subst enum_add, simp_all) finally show ?thesis . qed lemma offs_enum_ub [rule_format]: assumes A: "index_less index key" and B: "i < n" and C: "\x \ set xs. key x \ {mi..ma}" shows "offs (enum xs index key n mi ma) k ! i \ length xs + k" proof - let ?ns = "enum xs index key n mi ma" have "offs ?ns k ! i \ offs ?ns k ! (n - Suc 0)" using B by (rule_tac offs_mono, simp_all add: enum_length) also have "\ \ offs ?ns k ! (n - Suc 0) + offs_num n xs index key mi ma (n - Suc 0)" by simp also have "\ = length xs + k" using A and B and C by (rule_tac offs_enum_last, simp_all) finally show ?thesis . qed lemma offs_enum_next_ge [rule_format]: assumes A: "index_less index key" and B: "i < n" shows "\x \ set xs. key x \ {mi..ma} \ offs (enum xs index key n mi ma) k ! i \ offs_next (offs (enum xs index key n mi ma) k) (length xs + k) xs index key mi ma i" (is "_ \ offs ?ns _ ! _ \ _") proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI, rule offs_enum_ub [OF A B], simp) assume "offs_set_next (offs ?ns k) xs index key mi ma i \ {}" (is "?A \ _") hence C: "Min ?A \ ?A" by (rule_tac Min_in, simp) hence "i \ Min ?A" by simp moreover have "Min ?A < length ?ns" using C by (simp add: offs_length) ultimately show "offs ?ns k ! i \ offs ?ns k ! Min ?A" by (rule offs_mono) qed lemma offs_enum_zero_aux [rule_format]: "\index_less index key; 0 < n; \x \ set xs. key x \ {mi..ma}; offs_num n xs index key mi ma (n - Suc 0) = 0\ \ offs (enum xs index key n mi ma) k ! (n - Suc 0) = length xs + k" by (subst offs_enum_last [where key = key and mi = mi and ma = ma, symmetric], simp+) lemma offs_enum_zero [rule_format]: assumes A: "index_less index key" and B: "i < n" and C: "\x \ set xs. key x \ {mi..ma}" and D: "offs_num n xs index key mi ma i = 0" shows "offs (enum xs index key n mi ma) k ! i = offs_next (offs (enum xs index key n mi ma) k) (length xs + k) xs index key mi ma i" proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI, cases "i = n - Suc 0", simp) assume "i = n - Suc 0" thus "offs (enum xs index key n mi ma) k ! (n - Suc 0) = length xs + k" using A and B and C and D by (rule_tac offs_enum_zero_aux, simp_all) next let ?ns = "enum xs index key n mi ma" assume E: "offs_set_next (offs ?ns k) xs index key mi ma i = {}" (is "?A = _") assume "i \ n - Suc 0" hence F: "i < n - Suc 0" using B by simp hence "offs ?ns k ! i = offs ?ns k ! (n - Suc 0)" proof (rule offs_equal, simp_all add: enum_length le_less, erule_tac conjE, erule_tac disjE, rule_tac ccontr, drule_tac [2] sym, simp_all) fix j assume G: "j < n - Suc 0" hence "j < length (offs ?ns k)" by (simp add: offs_length enum_length) moreover assume "i < j" moreover assume "0 < ?ns ! j" hence "0 < offs_num (length (offs ?ns k)) xs index key mi ma j" using G by (subst (asm) enum_offs_num, simp_all add: offs_length enum_length) ultimately have "j \ ?A" by simp thus False using E by simp next show "?ns ! i = 0" using B and D by (subst enum_offs_num, simp_all) qed also from A and B and C have "\ = length xs + k" proof (rule_tac offs_enum_zero_aux, simp_all, rule_tac ccontr, simp) have "n - Suc 0 < length (offs ?ns k)" using B by (simp add: offs_length enum_length) moreover assume "0 < offs_num n xs index key mi ma (n - Suc 0)" hence "0 < offs_num (length (offs ?ns k)) xs index key mi ma (n - Suc 0)" by (simp add: offs_length enum_length) ultimately have "n - Suc 0 \ ?A" using F by simp thus False using E by simp qed finally show "offs (enum xs index key n mi ma) k ! i = length xs + k" . next let ?ns = "enum xs index key n mi ma" assume "offs_set_next (offs ?ns k) xs index key mi ma i \ {}" (is "?A \ _") hence "Min ?A \ ?A" by (rule_tac Min_in, simp) thus "offs ?ns k ! i = offs ?ns k ! Min ?A" proof (rule_tac offs_equal, simp_all add: le_less, simp add: offs_length, (erule_tac conjE)+, erule_tac disjE, rule_tac ccontr, drule_tac [2] sym, simp_all) fix j assume E: "j < Min ?A" and "Min ?A < length (offs ?ns k)" hence F: "j < length (offs ?ns k)" by simp moreover assume "i < j" moreover assume "0 < ?ns ! j" hence "0 < offs_num (length (offs ?ns k)) xs index key mi ma j" using F by (subst (asm) enum_offs_num, simp_all add: offs_length enum_length) ultimately have "j \ ?A" by simp hence "Min ?A \ j" by (rule_tac Min_le, simp) thus False using E by simp next show "?ns ! i = 0" using B and D by (subst enum_offs_num, simp_all) qed qed lemma offs_enum_next_cons [rule_format]: assumes A: "index_less index key" and B: "\x \ set xs. key x \ {mi..ma}" shows "(if i < index key x n mi ma then (\) else (<)) (offs_next (offs (enum xs index key n mi ma) k) (length xs + k) xs index key mi ma i) (offs_next (offs ((enum xs index key n mi ma) [index key x n mi ma := Suc (enum xs index key n mi ma ! index key x n mi ma)]) k) (Suc (length xs + k)) (x # xs) index key mi ma i)" (is "(if i < ?i' then _ else _) (offs_next (offs ?ns _) _ _ _ _ _ _ _) (offs_next (offs ?ns' _) _ _ _ _ _ _ _)") proof (simp_all only: offs_next_def not_less split: if_split, (rule conjI, rule impI)+, simp, simp, (rule_tac [!] impI, (rule_tac [!] conjI)?)+, rule_tac [2-3] FalseE, rule_tac [4] conjI, rule_tac [4-5] impI) assume C: "offs_set_next (offs ?ns k) xs index key mi ma i = {}" (is "?A = _") and D: "offs_set_next (offs ?ns' k) (x # xs) index key mi ma i \ {}" (is "?A' \ _") and E: "i < ?i'" from C have F: "\j \ ?i'. j \ ?A'" by (rule_tac allI, rule_tac impI, rule_tac notI, simp add: enum_length offs_length offs_num_cons split: if_split_asm, (erule_tac conjE)+, simp) from D have "Min ?A' \ ?A'" by (rule_tac Min_in, simp) hence G: "Min ?A' < n" by (simp add: offs_length enum_length) have H: "Min ?A' = ?i'" proof (rule Min_eqI, simp, rule eq_refl, erule contrapos_pp, insert F, simp) have "\j. j \ ?A'" using D by blast then obtain j where "j \ ?A'" .. moreover from this have "j = ?i'" by (erule_tac contrapos_pp, insert F, simp) ultimately show "?i' \ ?A'" by simp qed with G have "offs ?ns' k ! Min ?A' = offs ?ns k ! Min ?A'" by (subst offs_update, simp_all add: enum_length) also from A and B and G and H have "\ = offs_next (offs ?ns k) (length xs + k) xs index key mi ma (Min ?A')" proof (rule_tac offs_enum_zero, simp_all, rule_tac ccontr, simp) assume "?i' < n" and "0 < offs_num n xs index key mi ma ?i'" hence "?i' \ ?A" using E by (simp add: offs_length enum_length) thus False using C by simp qed also have "\ = length xs + k" proof (simp only: offs_next_def split: if_split, rule conjI, simp, rule impI, rule FalseE, simp, erule exE, (erule conjE)+) fix j assume "j < length (offs ?ns k)" moreover assume "Min ?A' < j" hence "i < j" using E and H by simp moreover assume "0 < offs_num (length (offs ?ns k)) xs index key mi ma j" ultimately have "j \ ?A" by simp thus False using C by simp qed finally show "length xs + k \ offs ?ns' k ! Min ?A'" by simp next assume C: "offs_set_next (offs ?ns k) xs index key mi ma i = {}" (is "?A = _") and D: "offs_set_next (offs ?ns' k) (x # xs) index key mi ma i \ {}" (is "?A' \ _") and E: "?i' \ i" have "\j. j \ ?A'" using D by blast then obtain j where F: "j \ ?A'" .. hence "j < length (offs ?ns k)" by (simp add: offs_length) moreover have "i < j" using F by simp moreover from this have "0 < offs_num (length (offs ?ns k)) xs index key mi ma j" using E and F by (simp add: offs_length enum_length offs_num_cons) ultimately have "j \ ?A" by simp thus False using C by simp next assume C: "offs_set_next (offs ?ns k) xs index key mi ma i \ {}" (is "?A \ _") and D: "offs_set_next (offs ?ns' k) (x # xs) index key mi ma i = {}" (is "?A' = _") have "\j. j \ ?A" using C by blast then obtain j where E: "j \ ?A" .. hence "j < length (offs ?ns' k)" by (simp add: offs_length) moreover have "i < j" using E by simp moreover have "0 < offs_num (length (offs ?ns' k)) (x # xs) index key mi ma j" using E by (simp add: offs_length enum_length offs_num_cons) ultimately have "j \ ?A'" by simp thus False using D by simp next assume "offs_set_next (offs ?ns k) xs index key mi ma i \ {}" (is "?A \ _") hence "Min ?A \ ?A" by (rule_tac Min_in, simp) hence C: "Min ?A < n" by (simp add: offs_length enum_length) assume "offs_set_next (offs ?ns' k) (x # xs) index key mi ma i \ {}" (is "?A' \ _") hence D: "Min ?A' \ ?A'" by (rule_tac Min_in, simp) hence E: "Min ?A' < n" by (simp add: offs_length enum_length) have "offs ?ns k ! Min ?A \ offs ?ns k ! Min ?A'" proof (cases "offs_num n xs index key mi ma (Min ?A') = 0") case True have "offs ?ns k ! Min ?A' = offs_next (offs ?ns k) (length xs + k) xs index key mi ma (Min ?A')" using A and B and E and True by (rule_tac offs_enum_zero, simp_all) also from A and B and C have "\ \ offs ?ns k ! Min ?A" proof (simp only: offs_next_def split: if_split, rule_tac conjI, rule_tac [!] impI, rule_tac offs_enum_ub, simp, simp, simp) assume "offs_set_next (offs ?ns k) xs index key mi ma (Min ?A') \ {}" (is "?B \ _") hence "Min ?B \ ?B" by (rule_tac Min_in, simp) hence "Min ?B \ ?A" using D by simp moreover from this have "Min ?A \ Min ?B" by (rule_tac Min_le, simp) ultimately show "offs ?ns k ! Min ?A \ offs ?ns k ! Min ?B" by (rule_tac offs_mono, simp_all add: offs_length) qed finally show ?thesis . next case False hence "Min ?A' \ ?A" using D by (simp add: offs_length enum_length) hence "Min ?A \ Min ?A'" by (rule_tac Min_le, simp) thus ?thesis by (rule offs_mono, simp_all add: enum_length E) qed also have "\ \ offs ?ns' k ! Min ?A'" using E by (subst offs_update, simp_all add: enum_length) finally show "offs ?ns k ! Min ?A \ offs ?ns' k ! Min ?A'" . next let ?A = "offs_set_next (offs ?ns k) xs index key mi ma i" assume "offs_set_next (offs ?ns' k) (x # xs) index key mi ma i \ {}" (is "?A' \ _") hence C: "Min ?A' \ ?A'" by (rule_tac Min_in, simp) hence D: "Min ?A' < n" by (simp add: offs_length enum_length) assume "?i' \ i" hence E: "?i' < Min ?A'" using C by simp hence "0 < offs_num n xs index key mi ma (Min ?A')" using C by (simp add: offs_length enum_length offs_num_cons) hence "Min ?A' \ ?A" using C by (simp add: offs_length enum_length) hence "Min ?A \ Min ?A'" by (rule_tac Min_le, simp) hence "offs ?ns k ! Min ?A \ offs ?ns k ! Min ?A'" by (rule offs_mono, simp_all add: enum_length D) also have "\ < offs ?ns' k ! Min ?A'" using E by (subst offs_update, simp_all add: enum_length D) finally show "offs ?ns k ! Min ?A < offs ?ns' k ! Min ?A'" . qed lemma offs_enum_pred [rule_format]: assumes A: "index_less index key" shows "(\x \ set xs. key x \ {mi..ma}) \ offs_pred (offs (enum xs index key n mi ma) k) (length xs + k) xs index key mi ma" proof (induction xs, simp add: offs_pred_def offs_num_def, simp add: Let_def offs_pred_def offs_length enum_length, rule impI, (erule conjE)+, simp, rule allI, rule impI, erule allE, drule mp, assumption) fix x xs i let ?i' = "index key x n mi ma" let ?ns = "enum xs index key n mi ma" let ?ns' = "?ns[?i' := Suc (?ns ! ?i')]" assume B: "\x \ set xs. mi \ key x \ key x \ ma" and C: "i < n" and D: "offs_num n xs index key mi ma i \ offs_next (offs ?ns k) (length xs + k) xs index key mi ma i - offs ?ns k ! i" have E: "(if i < ?i' then (\) else (<)) (offs_next (offs ?ns k) (length xs + k) xs index key mi ma i) (offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i)" using A and B by (subst offs_enum_next_cons, simp_all) show "offs_num n (x # xs) index key mi ma i \ offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i - offs ?ns' k ! i" proof (subst offs_update, simp add: enum_length C, rule linorder_cases [of i ?i'], simp_all add: offs_num_cons) assume "i < ?i'" hence "offs_next (offs ?ns k) (length xs + k) xs index key mi ma i \ offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i" using E by simp thus "offs_num n xs index key mi ma i \ offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i - offs ?ns k ! i" using D by arith next assume F: "i = ?i'" hence "Suc (offs_num n xs index key mi ma ?i') \ Suc (offs_next (offs ?ns k) (length xs + k) xs index key mi ma ?i' - offs ?ns k ! ?i')" using D by simp also from A and B and C and F have "\ = Suc (offs_next (offs ?ns k) (length xs + k) xs index key mi ma ?i') - offs ?ns k ! ?i'" by (rule_tac Suc_diff_le [symmetric], rule_tac offs_enum_next_ge, simp_all) finally have "Suc (offs_num n xs index key mi ma ?i') \ Suc (offs_next (offs ?ns k) (length xs + k) xs index key mi ma ?i') - offs ?ns k ! ?i'" . moreover have "Suc (offs_next (offs ?ns k) (length xs + k) xs index key mi ma ?i') \ offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma ?i'" using E and F by simp ultimately show "Suc (offs_num n xs index key mi ma ?i') \ offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma ?i' - offs ?ns k ! ?i'" by arith next assume "?i' < i" hence "Suc (offs_next (offs ?ns k) (length xs + k) xs index key mi ma i) \ offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i" using E by simp thus "offs_num n xs index key mi ma i \ offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i - Suc (offs ?ns k ! i)" using D by arith qed qed lemma fill_offs_enum_count_item [rule_format]: "\index_less index key; \x \ set xs. key x \ {mi..ma}; 0 < n\ \ count (mset (map the (fill xs (offs (enum xs index key n mi ma) 0) index key (length xs) mi ma))) x = count (mset xs) x" by (subst fill_count_item, simp_all, simp only: length_greater_0_conv [symmetric] offs_length enum_length, insert offs_enum_pred [of index key xs mi ma n 0], simp) text \ \null Using lemma @{thm [source] fill_offs_enum_count_item}, step 9 of the proof method can now be dealt with. It is accomplished by proving lemma @{text gcsort_count_inv}, which states that the number of the occurrences of whatever object in the objects' list is still the same after any recursive round. \null \ lemma nths_count: "count (mset (nths xs A)) x = count (mset xs) x - card {i. i < length xs \ i \ A \ xs ! i = x}" proof (induction xs arbitrary: A, simp_all add: nths_Cons) fix v xs A let ?B = "{i. i < length xs \ Suc i \ A \ xs ! i = x}" let ?C = "\v. {i. i < Suc (length xs) \ i \ A \ (v # xs) ! i = x}" have A: "\v. ?C v = ?C v \ {0} \ ?C v \ {i. \j. i = Suc j}" by (subst Int_Un_distrib [symmetric], auto, arith) have "\v. card (?C v) = card (?C v \ {0}) + card (?C v \ {i. \j. i = Suc j})" by (subst A, rule card_Un_disjoint, simp_all, blast) moreover have "\v. card ?B = card (?C v \ {i. \j. i = Suc j})" by (rule bij_betw_same_card [of Suc], auto) ultimately show "(0 \ A \ (v = x \ Suc (count (mset xs) x - card ?B) = Suc (count (mset xs) x) - card (?C x)) \ (v \ x \ count (mset xs) x - card ?B = count (mset xs) x - card (?C v))) \ (0 \ A \ (v = x \ count (mset xs) x - card ?B = Suc (count (mset xs) x) - card (?C x)) \ (v \ x \ count (mset xs) x - card ?B = count (mset xs) x - card (?C v)))" proof ((rule_tac [!] conjI, rule_tac [!] impI)+, simp_all) have "card ?B \ count (mset xs) x" by (simp add: count_mset length_filter_conv_card, rule card_mono, simp, blast) thus "Suc (count (mset xs) x - card ?B) = Suc (count (mset xs) x) - card ?B" by (rule Suc_diff_le [symmetric]) qed qed lemma round_count_inv [rule_format]: "index_less index key \ bn_inv p q t \ add_inv n t \ count_inv f t \ count_inv f (round index key p q r t)" proof (induction index key p q r t arbitrary: n f rule: round.induct, (rule_tac [!] impI)+, simp, simp, simp_all only: simp_thms) fix index p q r u ns xs n f and key :: "'a \ 'b" let ?t = "round index key p q r (u, ns, tl xs)" assume "\n f. bn_inv p q (u, ns, tl xs) \ add_inv n (u, ns, tl xs) \ count_inv f (u, ns, tl xs) \ count_inv f ?t" and "bn_inv p q (u, Suc 0 # ns, xs)" and "add_inv n (u, Suc 0 # ns, xs)" and "count_inv f (u, Suc 0 # ns, xs)" thus "count_inv f (round index key p q r (u, Suc 0 # ns, xs))" proof (cases ?t, simp add: add_suc, rule_tac allI, cases xs, simp_all add: disj_imp split: if_split_asm) fix y ys x and xs' :: "'a list" assume "\n f. foldl (+) 0 ns = n \ length ys = n \ (\x. count (mset ys) x = f x) \ (\x. count (mset xs') x = f x)" moreover assume "Suc (foldl (+) 0 ns) = n" and "Suc (length ys) = n" ultimately have "\n f. (\x. count (mset ys) x = f x) \ (\x. count (mset xs') x = f x)" by blast moreover assume A: "\x. (y = x \ Suc (count (mset ys) x) = f x) \ (y \ x \ count (mset ys) x = f x)" have "\x. count (mset ys) x = (f(y := f y - Suc 0)) x" (is "\x. _ = ?f' x") by (simp add: A, insert spec [OF A, where x = y], simp) ultimately have "\x. count (mset xs') x = ?f' x" .. thus "(y = x \ Suc (count (mset xs') x) = f x) \ (y \ x \ count (mset xs') x = f x)" by (simp, insert spec [OF A, where x = y], rule_tac impI, simp) qed next fix index p q r u m ns n f and key :: "'a \ 'b" and xs :: "'a list" let ?ws = "take (Suc (Suc m)) xs" let ?ys = "drop (Suc (Suc m)) xs" let ?r = "\m'. round_suc_suc index key ?ws m m' u" let ?t = "\r' v. round index key p q r' (v, ns, ?ys)" assume A: "index_less index key" assume "\ws a b c d e g h i n f. ws = ?ws \ a = bn_comp m p q r \ (b, c) = bn_comp m p q r \ d = ?r b \ (e, g) = ?r b \ (h, i) = g \ bn_inv p q (e, ns, ?ys) \ add_inv n (e, ns, ?ys) \ count_inv f (e, ns, ?ys) \ count_inv f (?t c e)" and "bn_inv p q (u, Suc (Suc m) # ns, xs)" and "add_inv n (u, Suc (Suc m) # ns, xs)" and "count_inv f (u, Suc (Suc m) # ns, xs)" thus "count_inv f (round index key p q r (u, Suc (Suc m) # ns, xs))" proof (simp split: prod.split, ((rule_tac allI)+, ((rule_tac impI)+)?)+, (erule_tac conjE)+, subst (asm) (2) add_base_zero, simp) fix m' r' v ms' ws' xs' x assume B: "?r m' = (v, ms', ws')" and C: "bn_comp m p q r = (m', r')" and D: "bn_valid m p q" and E: "Suc (Suc (foldl (+) 0 ns + m)) = n" and F: "length xs = n" assume "\ws a b c d e g h i n' f. ws = ?ws \ a = (m', r') \ b = m' \ c = r' \ d = (v, ms', ws') \ e = v \ g = (ms', ws') \ h = ms' \ i = ws' \ foldl (+) 0 ns = n' \ n - Suc (Suc m) = n' \ (\x. count (mset ?ys) x = f x) \ (\x. count (mset xs') x = f x)" hence "foldl (+) 0 ns = n - Suc (Suc m) \ (\x. count (mset xs') x = count (mset ?ys) x)" by simp hence "count (mset xs') x = count (mset ?ys) x" using E by simp moreover assume "\x. count (mset xs) x = f x" ultimately have "f x = count (mset ?ws) x + count (mset xs') x" by (subst (asm) append_take_drop_id [of "Suc (Suc m)", symmetric], subst (asm) mset_append, simp) with B [symmetric] show "count (mset ws') x + count (mset xs') x = f x" proof (simp add: round_suc_suc_def Let_def del: count_add_mset mset_map split: if_split_asm, subst (1 2) add_mset_add_single, simp only: count_single count_union) let ?nmi = "mini ?ws key" let ?nma = "maxi ?ws key" let ?xmi = "?ws ! ?nmi" let ?xma = "?ws ! ?nma" let ?mi = "key ?xmi" let ?ma = "key ?xma" let ?k = "case m of 0 \ m | Suc 0 \ m | Suc (Suc i) \ u + m'" let ?zs = "nths ?ws (- {?nmi, ?nma})" let ?ms = "enum ?zs index key ?k ?mi ?ma" let ?A = "{i. i < Suc (Suc m) \ (i = ?nmi \ i = ?nma) \ ?ws ! i = x}" have G: "length ?ws = Suc (Suc m)" using E and F by simp hence H: "card ?A \ count (mset ?ws) x" by (simp add: count_mset length_filter_conv_card, rule_tac card_mono, simp, blast) show "count (mset (map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma))) x + (if ?xma = x then 1 else 0) + (if ?xmi = x then 1 else 0) = count (mset ?ws) x" proof (cases "m = 0") case True hence I: "length ?zs = 0" using G by (simp add: mini_maxi_nths) have "count (mset ?zs) x = count (mset ?ws) x - card ?A" using G by (subst nths_count, simp) hence J: "count (mset ?ws) x = card ?A" using H and I by simp from I show ?thesis proof (simp, (rule_tac [!] conjI, rule_tac [!] impI)+, simp_all (no_asm_simp) add: True) assume "?xmi = x" and "?xma = x" with G have "?A = {?nmi, ?nma}" by (rule_tac set_eqI, rule_tac iffI, simp_all, erule_tac disjE, insert mini_less [of ?ws key], insert maxi_less [of ?ws key], simp_all) with G have "card ?A = Suc (Suc 0)" by (simp, subst card_insert_disjoint, simp_all, rule_tac mini_maxi_neq, simp) thus "Suc (Suc 0) = count (mset (take (Suc (Suc 0)) xs)) x" using True and J by simp next assume "?xmi \ x" and "?xma = x" with G have "?A = {?nma}" by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+, erule_tac disjE, insert maxi_less [of ?ws key], simp_all) thus "Suc 0 = count (mset (take (Suc (Suc 0)) xs)) x" using True and J by simp next assume "?xmi = x" and "?xma \ x" with G have "?A = {?nmi}" by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+, erule_tac disjE, insert mini_less [of ?ws key], simp_all) thus "Suc 0 = count (mset (take (Suc (Suc 0)) xs)) x" using True and J by simp next assume "?xmi \ x" and "?xma \ x" hence "?A = {}" by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+, erule_tac disjE, simp_all) hence "count (mset ?ws) x = 0" using J by simp thus "x \ set (take (Suc (Suc 0)) xs)" using True by simp qed next case False hence "0 < ?k" by (simp, drule_tac bn_comp_fst_nonzero [OF D], subst (asm) C, simp split: nat.split) hence "count (mset (map the (fill ?zs (offs ?ms 0) index key (length ?zs) ?mi ?ma))) x = count (mset ?zs) x" by (rule_tac fill_offs_enum_count_item [OF A], simp, rule_tac conjI, ((rule_tac mini_lb | rule_tac maxi_ub), erule_tac in_set_nthsD)+) with G show ?thesis proof (simp, (rule_tac [!] conjI, rule_tac [!] impI)+, simp_all add: mini_maxi_nths nths_count) assume "?xmi = x" and "?xma = x" with G have "?A = {?nmi, ?nma}" by (rule_tac set_eqI, rule_tac iffI, simp_all, erule_tac disjE, insert mini_less [of ?ws key], insert maxi_less [of ?ws key], simp_all) with G have "card ?A = Suc (Suc 0)" by (simp, subst card_insert_disjoint, simp_all, rule_tac mini_maxi_neq, simp) thus "Suc (Suc (count (mset ?ws) x - card ?A)) = count (mset ?ws) x" using H by simp next assume "?xmi \ x" and "?xma = x" with G have "?A = {?nma}" by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+, erule_tac disjE, insert maxi_less [of ?ws key], simp_all) thus "Suc (count (mset ?ws) x - card ?A) = count (mset ?ws) x" using H by simp next assume "?xmi = x" and "?xma \ x" with G have "?A = {?nmi}" by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+, erule_tac disjE, insert mini_less [of ?ws key], simp_all) thus "Suc (count (mset ?ws) x - card ?A) = count (mset ?ws) x" using H by simp next assume "?xmi \ x" and "?xma \ x" hence "?A = {}" by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+, erule_tac disjE, simp_all) thus "count (mset ?ws) x - card ?A = count (mset ?ws) x" by (simp (no_asm_simp)) qed qed qed qed qed lemma gcsort_count_inv: assumes A: "index_less index key" and B: "add_inv n t" and C: "n \ p" shows "\t' \ gcsort_set index key p t; count_inv f t\ \ count_inv f t'" by (erule gcsort_set.induct, simp, drule gcsort_add_inv [OF A _ B C], rule round_count_inv [OF A], simp_all del: bn_inv.simps, erule conjE, frule sym, erule subst, rule bn_inv_intro, insert C, simp) text \ \null The only remaining task is to address step 10 of the proof method, which is done by proving theorem @{text gcsort_count}. It holds under the conditions that the objects' number is not larger than the counters' upper bound and function @{text index} satisfies predicate @{const index_less}, and states that for any object, function @{const gcsort} leaves unchanged the number of its occurrences within the input objects' list. \null \ theorem gcsort_count: assumes A: "index_less index key" and B: "length xs \ p" shows "count (mset (gcsort index key p xs)) x = count (mset xs) x" proof - let ?t = "(0, [length xs], xs)" have "count_inv (count (mset xs)) (gcsort_aux index key p ?t)" by (rule gcsort_count_inv [OF A _ B], rule gcsort_add_input, rule gcsort_aux_set, rule gcsort_count_input) hence "count (mset (gcsort_out (gcsort_aux index key p ?t))) x = (count (mset xs)) x" by (rule gcsort_count_intro) moreover have "?t = gcsort_in xs" by (simp add: gcsort_in_def) ultimately show ?thesis by (simp add: gcsort_def) qed end diff --git a/thys/Universal_Turing_Machine/UTM.thy b/thys/Universal_Turing_Machine/UTM.thy --- a/thys/Universal_Turing_Machine/UTM.thy +++ b/thys/Universal_Turing_Machine/UTM.thy @@ -1,4708 +1,4708 @@ (* Title: thys/UTM.thy Author: Jian Xu, Xingyuan Zhang, and Christian Urban Modifications: Sebastiaan Joosten *) chapter \Construction of a Universal Turing Machine\ theory UTM imports Recursive Abacus UF HOL.GCD Turing_Hoare begin section \Wang coding of input arguments\ text \ The direct compilation of the universal function \rec_F\ can not give us UTM, because \rec_F\ is of arity 2, where the first argument represents the Godel coding of the TM being simulated and the second argument represents the right number (in Wang's coding) of the TM tape. (Notice, left number is always \0\ at the very beginning). However, UTM needs to simulate the execution of any TM which may very well take many input arguments. Therefore, a initialization TM needs to run before the TM compiled from \rec_F\, and the sequential composition of these two TMs will give rise to the UTM we are seeking. The purpose of this initialization TM is to transform the multiple input arguments of the TM being simulated into Wang's coding, so that it can be consumed by the TM compiled from \rec_F\ as the second argument. However, this initialization TM (named \t_wcode\) can not be constructed by compiling from any recursive function, because every recursive function takes a fixed number of input arguments, while \t_wcode\ needs to take varying number of arguments and tranform them into Wang's coding. Therefore, this section give a direct construction of \t_wcode\ with just some parts being obtained from recursive functions. \newlength{\basewidth} \settowidth{\basewidth}{xxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxx} \newlength{\baseheight} \settoheight{\baseheight}{$B:R$} \newcommand{\vsep}{5\baseheight} The TM used to generate the Wang's code of input arguments is divided into three TMs executed sequentially, namely $prepare$, $mainwork$ and $adjust$. According to the convention, the start state of ever TM is fixed to state $1$ while the final state is fixed to $0$. The input and output of $prepare$ are illustrated respectively by Figure \ref{prepare_input} and \ref{prepare_output}. \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} [tbox/.style = {draw, thick, inner sep = 5pt}] \node (0) {}; \node (1) [tbox, text height = 3.5pt, right = -0.9pt of 0] {$m$}; \node (2) [tbox, right = -0.9pt of 1] {$0$}; \node (3) [tbox, right = -0.9pt of 2] {$a_1$}; \node (4) [tbox, right = -0.9pt of 3] {$0$}; \node (5) [tbox, right = -0.9pt of 4] {$a_2$}; \node (6) [right = -0.9pt of 5] {\ldots \ldots}; \node (7) [tbox, right = -0.9pt of 6] {$a_n$}; \draw [->, >=latex, thick] (1)+(0, -4\baseheight) -- (1); \end{tikzpicture}} \caption{The input of TM $prepare$} \label{prepare_input} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.5pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$}; \node (7) [right = -0.9pt of 6] {\ldots \ldots}; \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_n$}; \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$0$}; \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$}; \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$1$}; \draw [->, >=latex, thick] (10)+(0, -4\baseheight) -- (10); \end{tikzpicture}} \caption{The output of TM $prepare$} \label{prepare_output} \end{figure} As shown in Figure \ref{prepare_input}, the input of $prepare$ is the same as the the input of UTM, where $m$ is the Godel coding of the TM being interpreted and $a_1$ through $a_n$ are the $n$ input arguments of the TM under interpretation. The purpose of $purpose$ is to transform this initial tape layout to the one shown in Figure \ref{prepare_output}, which is convenient for the generation of Wang's codding of $a_1, \ldots, a_n$. The coding procedure starts from $a_n$ and ends after $a_1$ is encoded. The coding result is stored in an accumulator at the end of the tape (initially represented by the $1$ two blanks right to $a_n$ in Figure \ref{prepare_output}). In Figure \ref{prepare_output}, arguments $a_1, \ldots, a_n$ are separated by two blanks on both ends with the rest so that movement conditions can be implemented conveniently in subsequent TMs, because, by convention, two consecutive blanks are usually used to signal the end or start of a large chunk of data. The diagram of $prepare$ is given in Figure \ref{prepare_diag}. \begin{figure}[h!] \centering \scalebox{0.9}{ \begin{tikzpicture} \node[circle,draw] (1) {$1$}; \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$}; \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$}; \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$}; \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$}; \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$}; \node[circle,draw] (7) at ($(6)+(0.3\basewidth, 0)$) {$7$}; \node[circle,draw] (8) at ($(7)+(0.3\basewidth, 0)$) {$0$}; \draw [->, >=latex] (1) edge [loop above] node[above] {$S_1:L$} (1) ; \draw [->, >=latex] (1) -- node[above] {$S_0:S_1$} (2) ; \draw [->, >=latex] (2) edge [loop above] node[above] {$S_1:R$} (2) ; \draw [->, >=latex] (2) -- node[above] {$S_0:L$} (3) ; \draw [->, >=latex] (3) edge[loop above] node[above] {$S_1:S_0$} (3) ; \draw [->, >=latex] (3) -- node[above] {$S_0:R$} (4) ; \draw [->, >=latex] (4) edge[loop above] node[above] {$S_0:R$} (4) ; \draw [->, >=latex] (4) -- node[above] {$S_0:R$} (5) ; \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:R$} (5) ; \draw [->, >=latex] (5) -- node[above] {$S_0:R$} (6) ; \draw [->, >=latex] (6) edge[bend left = 50] node[below] {$S_1:R$} (5) ; \draw [->, >=latex] (6) -- node[above] {$S_0:R$} (7) ; \draw [->, >=latex] (7) edge[loop above] node[above] {$S_0:S_1$} (7) ; \draw [->, >=latex] (7) -- node[above] {$S_1:L$} (8) ; \end{tikzpicture}} \caption{The diagram of TM $prepare$} \label{prepare_diag} \end{figure} The purpose of TM $mainwork$ is to compute the Wang's encoding of $a_1, \ldots, a_n$. Every bit of $a_1, \ldots, a_n$, including the separating bits, is processed from left to right. In order to detect the termination condition when the left most bit of $a_1$ is reached, TM $mainwork$ needs to look ahead and consider three different situations at the start of every iteration: \begin{enumerate} \item The TM configuration for the first situation is shown in Figure \ref{mainwork_case_one_input}, where the accumulator is stored in $r$, both of the next two bits to be encoded are $1$. The configuration at the end of the iteration is shown in Figure \ref{mainwork_case_one_output}, where the first 1-bit has been encoded and cleared. Notice that the accumulator has been changed to $(r+1) \times 2$ to reflect the encoded bit. \item The TM configuration for the second situation is shown in Figure \ref{mainwork_case_two_input}, where the accumulator is stored in $r$, the next two bits to be encoded are $1$ and $0$. After the first $1$-bit was encoded and cleared, the second $0$-bit is difficult to detect and process. To solve this problem, these two consecutive bits are encoded in one iteration. In this situation, only the first $1$-bit needs to be cleared since the second one is cleared by definition. The configuration at the end of the iteration is shown in Figure \ref{mainwork_case_two_output}. Notice that the accumulator has been changed to $(r+1) \times 4$ to reflect the two encoded bits. \item The third situation corresponds to the case when the last bit of $a_1$ is reached. The TM configurations at the start and end of the iteration are shown in Figure \ref{mainwork_case_three_input} and \ref{mainwork_case_three_output} respectively. For this situation, only the read write head needs to be moved to the left to prepare a initial configuration for TM $adjust$ to start with. \end{enumerate} The diagram of $mainwork$ is given in Figure \ref{mainwork_diag}. The two rectangular nodes labeled with $2 \times x$ and $4 \times x$ are two TMs compiling from recursive functions so that we do not have to design and verify two quite complicated TMs. \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$}; \node (7) [right = -0.9pt of 6] {\ldots \ldots}; \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$}; \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$}; \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$1$}; \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$0$}; \node (12) [right = -0.9pt of 11] {\ldots \ldots}; \node (13) [draw, right = -0.9pt of 12, thick, inner sep = 5pt] {$0$}; \node (14) [draw, text height = 3.9pt, right = -0.9pt of 13, thick, inner sep = 5pt] {$r$}; \draw [->, >=latex, thick] (13)+(0, -4\baseheight) -- (13); \end{tikzpicture}} \caption{The first situation for TM $mainwork$ to consider} \label{mainwork_case_one_input} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$}; \node (7) [right = -0.9pt of 6] {\ldots \ldots}; \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$}; \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$}; \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$}; \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$0$}; \node (12) [right = -0.9pt of 11] {\ldots \ldots}; \node (13) [draw, right = -0.9pt of 12, thick, inner sep = 5pt] {$0$}; \node (14) [draw, text height = 2.7pt, right = -0.9pt of 13, thick, inner sep = 5pt] {$(r+1) \times 2$}; \draw [->, >=latex, thick] (13)+(0, -4\baseheight) -- (13); \end{tikzpicture}} \caption{The output for the first case of TM $mainwork$'s processing} \label{mainwork_case_one_output} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$}; \node (7) [right = -0.9pt of 6] {\ldots \ldots}; \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$}; \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$}; \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$}; \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$1$}; \node (12) [draw, right = -0.9pt of 11, thick, inner sep = 5pt] {$0$}; \node (13) [right = -0.9pt of 12] {\ldots \ldots}; \node (14) [draw, right = -0.9pt of 13, thick, inner sep = 5pt] {$0$}; \node (15) [draw, text height = 3.9pt, right = -0.9pt of 14, thick, inner sep = 5pt] {$r$}; \draw [->, >=latex, thick] (14)+(0, -4\baseheight) -- (14); \end{tikzpicture}} \caption{The second situation for TM $mainwork$ to consider} \label{mainwork_case_two_input} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$}; \node (7) [right = -0.9pt of 6] {\ldots \ldots}; \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$}; \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$}; \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$}; \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$0$}; \node (12) [draw, right = -0.9pt of 11, thick, inner sep = 5pt] {$0$}; \node (13) [right = -0.9pt of 12] {\ldots \ldots}; \node (14) [draw, right = -0.9pt of 13, thick, inner sep = 5pt] {$0$}; \node (15) [draw, text height = 2.7pt, right = -0.9pt of 14, thick, inner sep = 5pt] {$(r+1) \times 4$}; \draw [->, >=latex, thick] (14)+(0, -4\baseheight) -- (14); \end{tikzpicture}} \caption{The output for the second case of TM $mainwork$'s processing} \label{mainwork_case_two_output} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [right = -0.9pt of 5] {\ldots \ldots}; \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {$0$}; \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {$r$}; \draw [->, >=latex, thick] (7)+(0, -4\baseheight) -- (7); \end{tikzpicture}} \caption{The third situation for TM $mainwork$ to consider} \label{mainwork_case_three_input} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [right = -0.9pt of 5] {\ldots \ldots}; \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {$0$}; \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {$r$}; \draw [->, >=latex, thick] (3)+(0, -4\baseheight) -- (3); \end{tikzpicture}} \caption{The output for the third case of TM $mainwork$'s processing} \label{mainwork_case_three_output} \end{figure} \begin{figure}[h!] \centering \scalebox{0.9}{ \begin{tikzpicture} \node[circle,draw] (1) {$1$}; \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$}; \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$}; \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$}; \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$}; \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$}; \node[circle,draw] (7) at ($(2)+(0, -7\baseheight)$) {$7$}; \node[circle,draw] (8) at ($(7)+(0, -7\baseheight)$) {$8$}; \node[circle,draw] (9) at ($(8)+(0.3\basewidth, 0)$) {$9$}; \node[circle,draw] (10) at ($(9)+(0.3\basewidth, 0)$) {$10$}; \node[circle,draw] (11) at ($(10)+(0.3\basewidth, 0)$) {$11$}; \node[circle,draw] (12) at ($(11)+(0.3\basewidth, 0)$) {$12$}; \node[draw] (13) at ($(6)+(0.3\basewidth, 0)$) {$2 \times x$}; \node[circle,draw] (14) at ($(13)+(0.3\basewidth, 0)$) {$j_1$}; \node[draw] (15) at ($(12)+(0.3\basewidth, 0)$) {$4 \times x$}; \node[draw] (16) at ($(15)+(0.3\basewidth, 0)$) {$j_2$}; \node[draw] (17) at ($(7)+(0.3\basewidth, 0)$) {$0$}; \draw [->, >=latex] (1) edge[loop left] node[above] {$S_0:L$} (1) ; \draw [->, >=latex] (1) -- node[above] {$S_1:L$} (2) ; \draw [->, >=latex] (2) -- node[above] {$S_1:R$} (3) ; \draw [->, >=latex] (2) -- node[left] {$S_1:L$} (7) ; \draw [->, >=latex] (3) edge[loop above] node[above] {$S_1:S_0$} (3) ; \draw [->, >=latex] (3) -- node[above] {$S_0:R$} (4) ; \draw [->, >=latex] (4) edge[loop above] node[above] {$S_0:R$} (4) ; \draw [->, >=latex] (4) -- node[above] {$S_1:R$} (5) ; \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:R$} (5) ; \draw [->, >=latex] (5) -- node[above] {$S_0:S_1$} (6) ; \draw [->, >=latex] (6) edge[loop above] node[above] {$S_1:L$} (6) ; \draw [->, >=latex] (6) -- node[above] {$S_0:R$} (13) ; \draw [->, >=latex] (13) -- (14) ; \draw (14) -- ($(14)+(0, 6\baseheight)$) -- ($(1) + (0, 6\baseheight)$) node [above,midway] {$S_1:L$} ; \draw [->, >=latex] ($(1) + (0, 6\baseheight)$) -- (1) ; \draw [->, >=latex] (7) -- node[above] {$S_0:R$} (17) ; \draw [->, >=latex] (7) -- node[left] {$S_1:R$} (8) ; \draw [->, >=latex] (8) -- node[above] {$S_0:R$} (9) ; \draw [->, >=latex] (9) -- node[above] {$S_0:R$} (10) ; \draw [->, >=latex] (10) -- node[above] {$S_1:R$} (11) ; \draw [->, >=latex] (10) edge[loop above] node[above] {$S_0:R$} (10) ; \draw [->, >=latex] (11) edge[loop above] node[above] {$S_1:R$} (11) ; \draw [->, >=latex] (11) -- node[above] {$S_0:S_1$} (12) ; \draw [->, >=latex] (12) -- node[above] {$S_0:R$} (15) ; \draw [->, >=latex] (12) edge[loop above] node[above] {$S_1:L$} (12) ; \draw [->, >=latex] (15) -- (16) ; \draw (16) -- ($(16)+(0, -4\baseheight)$) -- ($(1) + (0, -18\baseheight)$) node [below,midway] {$S_1:L$} ; \draw [->, >=latex] ($(1) + (0, -18\baseheight)$) -- (1) ; \end{tikzpicture}} \caption{The diagram of TM $mainwork$} \label{mainwork_diag} \end{figure} The purpose of TM $adjust$ is to encode the last bit of $a_1$. The initial and final configuration of this TM are shown in Figure \ref{adjust_initial} and \ref{adjust_final} respectively. The diagram of TM $adjust$ is shown in Figure \ref{adjust_diag}. \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [right = -0.9pt of 5] {\ldots \ldots}; \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {$0$}; \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {$r$}; \draw [->, >=latex, thick] (3)+(0, -4\baseheight) -- (3); \end{tikzpicture}} \caption{Initial configuration of TM $adjust$} \label{adjust_initial} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, text height = 2.9pt, right = -0.9pt of 2, thick, inner sep = 5pt] {$r+1$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$0$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [right = -0.9pt of 5] {\ldots \ldots}; \draw [->, >=latex, thick] (1)+(0, -4\baseheight) -- (1); \end{tikzpicture}} \caption{Final configuration of TM $adjust$} \label{adjust_final} \end{figure} \begin{figure}[h!] \centering \scalebox{0.9}{ \begin{tikzpicture} \node[circle,draw] (1) {$1$}; \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$}; \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$}; \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$}; \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$}; \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$}; \node[circle,draw] (7) at ($(6)+(0.3\basewidth, 0)$) {$7$}; \node[circle,draw] (8) at ($(4)+(0, -7\baseheight)$) {$8$}; \node[circle,draw] (9) at ($(8)+(0.3\basewidth, 0)$) {$9$}; \node[circle,draw] (10) at ($(9)+(0.3\basewidth, 0)$) {$10$}; \node[circle,draw] (11) at ($(10)+(0.3\basewidth, 0)$) {$11$}; \node[circle,draw] (12) at ($(11)+(0.3\basewidth, 0)$) {$0$}; \draw [->, >=latex] (1) -- node[above] {$S_1:R$} (2) ; \draw [->, >=latex] (1) edge[loop above] node[above] {$S_0:S_1$} (1) ; \draw [->, >=latex] (2) -- node[above] {$S_1:R$} (3) ; \draw [->, >=latex] (3) edge[loop above] node[above] {$S_0:R$} (3) ; \draw [->, >=latex] (3) -- node[above] {$S_1:R$} (4) ; \draw [->, >=latex] (4) -- node[above] {$S_1:L$} (5) ; \draw [->, >=latex] (4) -- node[right] {$S_0:L$} (8) ; \draw [->, >=latex] (5) -- node[above] {$S_0:L$} (6) ; \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:S_0$} (5) ; \draw [->, >=latex] (6) -- node[above] {$S_1:R$} (7) ; \draw [->, >=latex] (6) edge[loop above] node[above] {$S_0:L$} (6) ; \draw (7) -- ($(7)+(0, 6\baseheight)$) -- ($(2) + (0, 6\baseheight)$) node [above,midway] {$S_0:S_1$} ; \draw [->, >=latex] ($(2) + (0, 6\baseheight)$) -- (2) ; \draw [->, >=latex] (8) edge[loop left] node[left] {$S_1:S_0$} (8) ; \draw [->, >=latex] (8) -- node[above] {$S_0:L$} (9) ; \draw [->, >=latex] (9) edge[loop above] node[above] {$S_0:L$} (9) ; \draw [->, >=latex] (9) -- node[above] {$S_1:L$} (10) ; \draw [->, >=latex] (10) edge[loop above] node[above] {$S_0:L$} (10) ; \draw [->, >=latex] (10) -- node[above] {$S_0:L$} (11) ; \draw [->, >=latex] (11) edge[loop above] node[above] {$S_1:L$} (11) ; \draw [->, >=latex] (11) -- node[above] {$S_0:R$} (12) ; \end{tikzpicture}} \caption{Diagram of TM $adjust$} \label{adjust_diag} \end{figure} \ definition rec_twice :: "recf" where "rec_twice = Cn 1 rec_mult [id 1 0, constn 2]" definition rec_fourtimes :: "recf" where "rec_fourtimes = Cn 1 rec_mult [id 1 0, constn 4]" definition abc_twice :: "abc_prog" where "abc_twice = (let (aprog, ary, fp) = rec_ci rec_twice in aprog [+] dummy_abc ((Suc 0)))" definition abc_fourtimes :: "abc_prog" where "abc_fourtimes = (let (aprog, ary, fp) = rec_ci rec_fourtimes in aprog [+] dummy_abc ((Suc 0)))" definition twice_ly :: "nat list" where "twice_ly = layout_of abc_twice" definition fourtimes_ly :: "nat list" where "fourtimes_ly = layout_of abc_fourtimes" definition t_twice_compile :: "instr list" where "t_twice_compile= (tm_of abc_twice @ (shift (mopup 1) (length (tm_of abc_twice) div 2)))" definition t_twice :: "instr list" where "t_twice = adjust0 t_twice_compile" definition t_fourtimes_compile :: "instr list" where "t_fourtimes_compile= (tm_of abc_fourtimes @ (shift (mopup 1) (length (tm_of abc_fourtimes) div 2)))" definition t_fourtimes :: "instr list" where "t_fourtimes = adjust0 t_fourtimes_compile" definition t_twice_len :: "nat" where "t_twice_len = length t_twice div 2" definition t_wcode_main_first_part:: "instr list" where "t_wcode_main_first_part \ [(L, 1), (L, 2), (L, 7), (R, 3), (R, 4), (W0, 3), (R, 4), (R, 5), (W1, 6), (R, 5), (R, 13), (L, 6), (R, 0), (R, 8), (R, 9), (Nop, 8), (R, 10), (W0, 9), (R, 10), (R, 11), (W1, 12), (R, 11), (R, t_twice_len + 14), (L, 12)]" definition t_wcode_main :: "instr list" where "t_wcode_main = (t_wcode_main_first_part @ shift t_twice 12 @ [(L, 1), (L, 1)] @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])" fun bl_bin :: "cell list \ nat" where "bl_bin [] = 0" | "bl_bin (Bk # xs) = 2 * bl_bin xs" | "bl_bin (Oc # xs) = Suc (2 * bl_bin xs)" declare bl_bin.simps[simp del] type_synonym bin_inv_t = "cell list \ nat \ tape \ bool" fun wcode_before_double :: "bin_inv_t" where "wcode_before_double ires rs (l, r) = (\ ln rn. l = Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\((Suc (Suc rs))) @ Bk\(rn ))" declare wcode_before_double.simps[simp del] fun wcode_after_double :: "bin_inv_t" where "wcode_after_double ires rs (l, r) = (\ ln rn. l = Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(Suc (Suc (Suc 2*rs))) @ Bk\(rn))" declare wcode_after_double.simps[simp del] fun wcode_on_left_moving_1_B :: "bin_inv_t" where "wcode_on_left_moving_1_B ires rs (l, r) = (\ ml mr rn. l = Bk\(ml) @ Oc # Oc # ires \ r = Bk\(mr) @ Oc\(Suc rs) @ Bk\(rn) \ ml + mr > Suc 0 \ mr > 0)" declare wcode_on_left_moving_1_B.simps[simp del] fun wcode_on_left_moving_1_O :: "bin_inv_t" where "wcode_on_left_moving_1_O ires rs (l, r) = (\ ln rn. l = Oc # ires \ r = Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" declare wcode_on_left_moving_1_O.simps[simp del] fun wcode_on_left_moving_1 :: "bin_inv_t" where "wcode_on_left_moving_1 ires rs (l, r) = (wcode_on_left_moving_1_B ires rs (l, r) \ wcode_on_left_moving_1_O ires rs (l, r))" declare wcode_on_left_moving_1.simps[simp del] fun wcode_on_checking_1 :: "bin_inv_t" where "wcode_on_checking_1 ires rs (l, r) = (\ ln rn. l = ires \ r = Oc # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_erase1 :: "bin_inv_t" where "wcode_erase1 ires rs (l, r) = (\ ln rn. l = Oc # ires \ tl r = Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" declare wcode_erase1.simps [simp del] fun wcode_on_right_moving_1 :: "bin_inv_t" where "wcode_on_right_moving_1 ires rs (l, r) = (\ ml mr rn. l = Bk\(ml) @ Oc # ires \ r = Bk\(mr) @ Oc\(Suc rs) @ Bk\(rn) \ ml + mr > Suc 0)" declare wcode_on_right_moving_1.simps [simp del] declare wcode_on_right_moving_1.simps[simp del] fun wcode_goon_right_moving_1 :: "bin_inv_t" where "wcode_goon_right_moving_1 ires rs (l, r) = (\ ml mr ln rn. l = Oc\(ml) @ Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = Suc rs)" declare wcode_goon_right_moving_1.simps[simp del] fun wcode_backto_standard_pos_B :: "bin_inv_t" where "wcode_backto_standard_pos_B ires rs (l, r) = (\ ln rn. l = Bk # Bk\(ln) @ Oc # ires \ r = Bk # Oc\((Suc (Suc rs))) @ Bk\(rn ))" declare wcode_backto_standard_pos_B.simps[simp del] fun wcode_backto_standard_pos_O :: "bin_inv_t" where "wcode_backto_standard_pos_O ires rs (l, r) = (\ ml mr ln rn. l = Oc\(ml) @ Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = Suc (Suc rs) \ mr > 0)" declare wcode_backto_standard_pos_O.simps[simp del] fun wcode_backto_standard_pos :: "bin_inv_t" where "wcode_backto_standard_pos ires rs (l, r) = (wcode_backto_standard_pos_B ires rs (l, r) \ wcode_backto_standard_pos_O ires rs (l, r))" declare wcode_backto_standard_pos.simps[simp del] lemma bin_wc_eq: "bl_bin xs = bl2wc xs" proof(induct xs) show " bl_bin [] = bl2wc []" apply(simp add: bl_bin.simps) done next fix a xs assume "bl_bin xs = bl2wc xs" thus " bl_bin (a # xs) = bl2wc (a # xs)" apply(case_tac a, simp_all add: bl_bin.simps bl2wc.simps) apply(simp_all add: bl2nat.simps bl2nat_double) done qed lemma tape_of_nl_append_one: "lm \ [] \ = @ Bk # Oc\Suc a" apply(induct lm, auto simp: tape_of_nl_cons split:if_splits) done lemma tape_of_nl_rev: "rev () = ()" apply(induct lm, simp, auto) apply(auto simp: tape_of_nl_cons tape_of_nl_append_one split: if_splits) apply(simp add: exp_ind[THEN sym]) done lemma exp_1[simp]: "a\(Suc 0) = [a]" by(simp) lemma tape_of_nl_cons_app1: "() = (Oc\(Suc a) @ Bk # ())" apply(case_tac xs; simp add: tape_of_list_def tape_of_nat_list.simps tape_of_nat_def) done lemma bl_bin_bk_oc[simp]: "bl_bin (xs @ [Bk, Oc]) = bl_bin xs + 2*2^(length xs)" apply(simp add: bin_wc_eq) using bl2nat_cons_oc[of "xs @ [Bk]"] apply(simp add: bl2nat_cons_bk bl2wc.simps) done lemma tape_of_nat[simp]: "() = Oc\(Suc a)" apply(simp add: tape_of_nat_def) done lemma tape_of_nl_cons_app2: "() = ( @ Bk # Oc\(Suc b))" proof(induct "length xs" arbitrary: xs c, simp add: tape_of_list_def) fix x xs c assume ind: "\xs c. x = length xs \ = @ Bk # Oc\(Suc b)" and h: "Suc x = length (xs::nat list)" show " = @ Bk # Oc\(Suc b)" proof(cases xs, simp add: tape_of_list_def) fix a list assume g: "xs = a # list" hence k: " = @ Bk # Oc\(Suc b)" apply(rule_tac ind) using h apply(simp) done from g and k show " = @ Bk # Oc\(Suc b)" apply(simp add: tape_of_list_def) done qed qed lemma length_2_elems[simp]: "length () = Suc (Suc aa) + length ()" apply(simp add: tape_of_list_def) done lemma bl_bin_addition[simp]: "bl_bin (Oc\(Suc aa) @ Bk # tape_of_nat_list (a # lista) @ [Bk, Oc]) = bl_bin (Oc\(Suc aa) @ Bk # tape_of_nat_list (a # lista)) + 2* 2^(length (Oc\(Suc aa) @ Bk # tape_of_nat_list (a # lista)))" using bl_bin_bk_oc[of "Oc\(Suc aa) @ Bk # tape_of_nat_list (a # lista)"] apply(simp) done declare replicate_Suc[simp del] lemma bl_bin_2[simp]: "bl_bin () + (4 * rs + 4) * 2 ^ (length () - Suc 0) = bl_bin (Oc\(Suc aa) @ Bk # ) + rs * (2 * 2 ^ (aa + length ()))" apply(case_tac "list", simp add: add_mult_distrib) apply(simp add: tape_of_nl_cons_app2 add_mult_distrib) apply(simp add: tape_of_list_def) done lemma tape_of_nl_app_Suc: "(()) = () @ [Oc]" proof(induct list) case (Cons a list) then show ?case by(cases list;simp_all add:tape_of_list_def exp_ind) qed (simp add: tape_of_list_def exp_ind) lemma bl_bin_3[simp]: "bl_bin (Oc # Oc\(aa) @ Bk # @ [Oc]) = bl_bin (Oc # Oc\(aa) @ Bk # ) + 2^(length (Oc # Oc\(aa) @ Bk # ))" apply(simp add: bin_wc_eq) apply(simp add: bl2nat_cons_oc bl2wc.simps) using bl2nat_cons_oc[of "Oc # Oc\(aa) @ Bk # "] apply(simp) done lemma bl_bin_4[simp]: "bl_bin (Oc # Oc\(aa) @ Bk # ) + (4 * 2 ^ (aa + length ()) + 4 * (rs * 2 ^ (aa + length ()))) = bl_bin (Oc # Oc\(aa) @ Bk # ) + rs * (2 * 2 ^ (aa + length ()))" apply(simp add: tape_of_nl_app_Suc) done declare tape_of_nat[simp del] fun wcode_double_case_inv :: "nat \ bin_inv_t" where "wcode_double_case_inv st ires rs (l, r) = (if st = Suc 0 then wcode_on_left_moving_1 ires rs (l, r) else if st = Suc (Suc 0) then wcode_on_checking_1 ires rs (l, r) else if st = 3 then wcode_erase1 ires rs (l, r) else if st = 4 then wcode_on_right_moving_1 ires rs (l, r) else if st = 5 then wcode_goon_right_moving_1 ires rs (l, r) else if st = 6 then wcode_backto_standard_pos ires rs (l, r) else if st = 13 then wcode_before_double ires rs (l, r) else False)" declare wcode_double_case_inv.simps[simp del] fun wcode_double_case_state :: "config \ nat" where "wcode_double_case_state (st, l, r) = 13 - st" fun wcode_double_case_step :: "config \ nat" where "wcode_double_case_step (st, l, r) = (if st = Suc 0 then (length l) else if st = Suc (Suc 0) then (length r) else if st = 3 then if hd r = Oc then 1 else 0 else if st = 4 then (length r) else if st = 5 then (length r) else if st = 6 then (length l) else 0)" fun wcode_double_case_measure :: "config \ nat \ nat" where "wcode_double_case_measure (st, l, r) = (wcode_double_case_state (st, l, r), wcode_double_case_step (st, l, r))" definition wcode_double_case_le :: "(config \ config) set" where "wcode_double_case_le \ (inv_image lex_pair wcode_double_case_measure)" lemma wf_lex_pair[intro]: "wf lex_pair" by(auto intro:wf_lex_prod simp:lex_pair_def) lemma wf_wcode_double_case_le[intro]: "wf wcode_double_case_le" by(auto intro:wf_inv_image simp: wcode_double_case_le_def ) lemma fetch_t_wcode_main[simp]: "fetch t_wcode_main (Suc 0) Bk = (L, Suc 0)" "fetch t_wcode_main (Suc 0) Oc = (L, Suc (Suc 0))" "fetch t_wcode_main (Suc (Suc 0)) Oc = (R, 3)" "fetch t_wcode_main (Suc (Suc 0)) Bk = (L, 7)" "fetch t_wcode_main (Suc (Suc (Suc 0))) Bk = (R, 4)" "fetch t_wcode_main (Suc (Suc (Suc 0))) Oc = (W0, 3)" "fetch t_wcode_main 4 Bk = (R, 4)" "fetch t_wcode_main 4 Oc = (R, 5)" "fetch t_wcode_main 5 Oc = (R, 5)" "fetch t_wcode_main 5 Bk = (W1, 6)" "fetch t_wcode_main 6 Bk = (R, 13)" "fetch t_wcode_main 6 Oc = (L, 6)" "fetch t_wcode_main 7 Oc = (R, 8)" "fetch t_wcode_main 7 Bk = (R, 0)" "fetch t_wcode_main 8 Bk = (R, 9)" "fetch t_wcode_main 9 Bk = (R, 10)" "fetch t_wcode_main 9 Oc = (W0, 9)" "fetch t_wcode_main 10 Bk = (R, 10)" "fetch t_wcode_main 10 Oc = (R, 11)" "fetch t_wcode_main 11 Bk = (W1, 12)" "fetch t_wcode_main 11 Oc = (R, 11)" "fetch t_wcode_main 12 Oc = (L, 12)" "fetch t_wcode_main 12 Bk = (R, t_twice_len + 14)" by(auto simp: t_wcode_main_def t_wcode_main_first_part_def fetch.simps numeral) declare wcode_on_checking_1.simps[simp del] lemmas wcode_double_case_inv_simps = wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps wcode_on_left_moving_1_B.simps wcode_on_checking_1.simps wcode_erase1.simps wcode_on_right_moving_1.simps wcode_goon_right_moving_1.simps wcode_backto_standard_pos.simps lemma wcode_on_left_moving_1[simp]: "wcode_on_left_moving_1 ires rs (b, []) = False" "wcode_on_left_moving_1 ires rs (b, r) \ b \ []" by(auto simp: wcode_on_left_moving_1.simps wcode_on_left_moving_1_B.simps wcode_on_left_moving_1_O.simps) lemma wcode_on_left_moving_1E[elim]: "\wcode_on_left_moving_1 ires rs (b, Bk # list); tl b = aa \ hd b # Bk # list = ba\ \ wcode_on_left_moving_1 ires rs (aa, ba)" apply(simp only: wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps wcode_on_left_moving_1_B.simps) apply(erule_tac disjE) apply(erule_tac exE)+ apply(rename_tac ml mr rn) apply(case_tac ml, simp) apply(rule_tac x = "mr - Suc (Suc 0)" in exI, rule_tac x = rn in exI) apply (smt One_nat_def Suc_diff_Suc append_Cons empty_replicate list.sel(3) neq0_conv replicate_Suc replicate_app_Cons_same tl_append2 tl_replicate) apply(rule_tac disjI1) apply (metis add_Suc_shift less_SucI list.exhaust_sel list.inject list.simps(3) replicate_Suc_iff_anywhere) by simp declare replicate_Suc[simp] lemma wcode_on_moving_1_Elim[elim]: "\wcode_on_left_moving_1 ires rs (b, Oc # list); tl b = aa \ hd b # Oc # list = ba\ \ wcode_on_checking_1 ires rs (aa, ba)" apply(simp only: wcode_double_case_inv_simps) apply(erule_tac disjE) apply (metis cell.distinct(1) empty_replicate hd_append2 hd_replicate list.sel(1) not_gr_zero) apply force. lemma wcode_on_checking_1_Elim[elim]: "\wcode_on_checking_1 ires rs (b, Oc # ba);Oc # b = aa \ list = ba\ \ wcode_erase1 ires rs (aa, ba)" apply(simp only: wcode_double_case_inv_simps) apply(erule_tac exE)+ by auto lemma wcode_on_checking_1_simp[simp]: "wcode_on_checking_1 ires rs (b, []) = False" "wcode_on_checking_1 ires rs (b, Bk # list) = False" by(auto simp: wcode_double_case_inv_simps) lemma wcode_erase1_nonempty_snd[simp]: "wcode_erase1 ires rs (b, []) = False" apply(simp add: wcode_double_case_inv_simps) done lemma wcode_on_right_moving_1_nonempty_snd[simp]: "wcode_on_right_moving_1 ires rs (b, []) = False" apply(simp add: wcode_double_case_inv_simps) done lemma wcode_on_right_moving_1_BkE[elim]: "\wcode_on_right_moving_1 ires rs (b, Bk # ba); Bk # b = aa \ list = b\ \ wcode_on_right_moving_1 ires rs (aa, ba)" apply(simp only: wcode_double_case_inv_simps) apply(erule_tac exE)+ apply(rename_tac ml mr rn) apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI, rule_tac x = rn in exI) apply(simp) apply(case_tac mr, simp, simp) done lemma wcode_on_right_moving_1_OcE[elim]: "\wcode_on_right_moving_1 ires rs (b, Oc # ba); Oc # b = aa \ list = ba\ \ wcode_goon_right_moving_1 ires rs (aa, ba)" apply(simp only: wcode_double_case_inv_simps) apply(erule_tac exE)+ apply(rename_tac ml mr rn) apply(rule_tac x = "Suc 0" in exI, rule_tac x = "rs" in exI, rule_tac x = "ml - Suc (Suc 0)" in exI, rule_tac x = rn in exI) apply(case_tac mr, simp_all) apply(case_tac ml, simp, case_tac nat, simp, simp) done lemma wcode_erase1_BkE[elim]: assumes "wcode_erase1 ires rs (b, Bk # ba)" "Bk # b = aa \ list = ba" "c = Bk # ba" shows "wcode_on_right_moving_1 ires rs (aa, ba)" proof - from assms obtain rn ln where "b = Oc # ires" "tl (Bk # ba) = Bk \ ln @ Bk # Bk # Oc \ Suc rs @ Bk \ rn" unfolding wcode_double_case_inv_simps by auto thus ?thesis using assms(2-) unfolding wcode_double_case_inv_simps apply(rule_tac x = "Suc 0" in exI, rule_tac x = "Suc (Suc ln)" in exI, rule_tac x = rn in exI, simp add: exp_ind del: replicate_Suc) done qed lemma wcode_erase1_OcE[elim]: "\wcode_erase1 ires rs (aa, Oc # list); b = aa \ Bk # list = ba\ \ wcode_erase1 ires rs (aa, ba)" unfolding wcode_double_case_inv_simps by auto auto lemma wcode_goon_right_moving_1_emptyE[elim]: assumes "wcode_goon_right_moving_1 ires rs (aa, [])" "b = aa \ [Oc] = ba" shows "wcode_backto_standard_pos ires rs (aa, ba)" proof - from assms obtain ml ln rn mr where "aa = Oc \ ml @ Bk # Bk # Bk \ ln @ Oc # ires" "[] = Oc \ mr @ Bk \ rn" "ml + mr = Suc rs" by(auto simp:wcode_double_case_inv_simps) thus ?thesis using assms(2) apply(simp only: wcode_double_case_inv_simps) apply(rule_tac disjI2) apply(simp only:wcode_backto_standard_pos_O.simps) apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp) done qed lemma wcode_goon_right_moving_1_BkE[elim]: assumes "wcode_goon_right_moving_1 ires rs (aa, Bk # list)" "b = aa \ Oc # list = ba" shows "wcode_backto_standard_pos ires rs (aa, ba)" proof - from assms obtain ln rn where "aa = Oc \ Suc rs @ Bk \ Suc (Suc ln) @ Oc # ires" "Bk # list = Bk \ rn" "b = Oc \ Suc rs @ Bk \ Suc (Suc ln) @ Oc # ires" "ba = Oc # list" by(auto simp:wcode_double_case_inv_simps) thus ?thesis using assms(2) apply(simp only: wcode_double_case_inv_simps wcode_backto_standard_pos_O.simps) apply(rule_tac disjI2) apply(rule exI[of _ "Suc rs"], rule exI[of _ "Suc 0"], rule_tac x = ln in exI, rule_tac x = "rn - Suc 0" in exI, simp) apply(cases rn;auto) done qed lemma wcode_goon_right_moving_1_OcE[elim]: assumes "wcode_goon_right_moving_1 ires rs (b, Oc # ba)" "Oc # b = aa \ list = ba" shows "wcode_goon_right_moving_1 ires rs (aa, ba)" proof - from assms obtain ml mr ln rn where "b = Oc \ ml @ Bk # Bk # Bk \ ln @ Oc # ires \ Oc # ba = Oc \ mr @ Bk \ rn \ ml + mr = Suc rs" unfolding wcode_double_case_inv_simps by auto with assms(2) show ?thesis unfolding wcode_double_case_inv_simps apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) apply(simp) apply(case_tac mr, simp, case_tac rn, simp_all) done qed lemma wcode_backto_standard_pos_BkE[elim]: "\wcode_backto_standard_pos ires rs (b, Bk # ba); Bk # b = aa \ list = ba\ \ wcode_before_double ires rs (aa, ba)" apply(simp only: wcode_double_case_inv_simps wcode_backto_standard_pos_B.simps wcode_backto_standard_pos_O.simps wcode_before_double.simps) apply(erule_tac disjE) apply(erule_tac exE)+ by auto lemma wcode_backto_standard_pos_no_Oc[simp]: "wcode_backto_standard_pos ires rs ([], Oc # list) = False" apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps wcode_backto_standard_pos_O.simps) done lemma wcode_backto_standard_pos_nonempty_snd[simp]: "wcode_backto_standard_pos ires rs (b, []) = False" apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps wcode_backto_standard_pos_O.simps) done lemma wcode_backto_standard_pos_OcE[elim]: "\wcode_backto_standard_pos ires rs (b, Oc # list); tl b = aa; hd b # Oc # list = ba\ \ wcode_backto_standard_pos ires rs (aa, ba)" apply(simp only: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps wcode_backto_standard_pos_O.simps) apply(erule_tac disjE) apply(simp) apply(erule_tac exE)+ apply(simp) apply (rename_tac ml mr ln rn) apply(case_tac ml) apply(rule_tac disjI1, rule_tac conjI) apply(rule_tac x = ln in exI, force, rule_tac x = rn in exI, force, force). declare nth_of.simps[simp del] fetch.simps[simp del] lemma wcode_double_case_first_correctness: "let P = (\ (st, l, r). st = 13) in let Q = (\ (st, l, r). wcode_double_case_inv st ires rs (l, r)) in let f = (\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp) in \ n .P (f n) \ Q (f (n::nat))" proof - let ?P = "(\ (st, l, r). st = 13)" let ?Q = "(\ (st, l, r). wcode_double_case_inv st ires rs (l, r))" let ?f = "(\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp)" have "\ n. ?P (?f n) \ ?Q (?f (n::nat))" proof(rule_tac halt_lemma2) show "wf wcode_double_case_le" by auto next show "\ na. \ ?P (?f na) \ ?Q (?f na) \ ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_double_case_le" proof(rule_tac allI, case_tac "?f na", simp) fix na a b c show "a \ 13 \ wcode_double_case_inv a ires rs (b, c) \ (case step0 (a, b, c) t_wcode_main of (st, x) \ wcode_double_case_inv st ires rs x) \ (step0 (a, b, c) t_wcode_main, a, b, c) \ wcode_double_case_le" apply(rule_tac impI, simp add: wcode_double_case_inv.simps) apply(auto split: if_splits simp: step.simps, case_tac [!] c, simp_all, case_tac [!] "(c::cell list)!0") apply(simp_all add: wcode_double_case_inv.simps wcode_double_case_le_def lex_pair_def) apply(auto split: if_splits) done qed next show "?Q (?f 0)" apply(simp add: steps.simps wcode_double_case_inv.simps wcode_on_left_moving_1.simps wcode_on_left_moving_1_B.simps) apply(rule_tac disjI1) apply(rule_tac x = "Suc m" in exI, simp) apply(rule_tac x = "Suc 0" in exI, simp) done next show "\ ?P (?f 0)" apply(simp add: steps.simps) done qed thus "let P = \(st, l, r). st = 13; Q = \(st, l, r). wcode_double_case_inv st ires rs (l, r); f = steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main in \n. P (f n) \ Q (f n)" apply(simp) done qed lemma tm_append_shift_append_steps: "\steps0 (st, l, r) tp stp = (st', l', r'); 0 < st'; length tp1 mod 2 = 0 \ \ steps0 (st + length tp1 div 2, l, r) (tp1 @ shift tp (length tp1 div 2) @ tp2) stp = (st' + length tp1 div 2, l', r')" proof - assume h: "steps0 (st, l, r) tp stp = (st', l', r')" "0 < st'" "length tp1 mod 2 = 0 " from h have "steps (st + length tp1 div 2, l, r) (tp1 @ shift tp (length tp1 div 2), 0) stp = (st' + length tp1 div 2, l', r')" by(rule_tac tm_append_second_steps_eq, simp_all) then have "steps (st + length tp1 div 2, l, r) ((tp1 @ shift tp (length tp1 div 2)) @ tp2, 0) stp = (st' + length tp1 div 2, l', r')" using h apply(rule_tac tm_append_first_steps_eq, simp_all) done thus "?thesis" by simp qed 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 = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" 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 (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 "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 "tm_of abc_twice = tm_of (a [+] dummy_abc (length [rs]))" using h by(simp add: abc_twice_def) qed thus "?thesis" apply(simp add: tape_of_list_def tape_of_nat_def rec_exec.simps twice_lemma) done qed declare adjust.simps[simp] lemma adjust_fetch0: "\0 < a; a \ length ap div 2; fetch ap a b = (aa, 0)\ \ fetch (adjust0 ap) a b = (aa, Suc (length ap div 2))" apply(case_tac b, auto simp: fetch.simps nth_of.simps nth_map split: if_splits) apply(case_tac [!] a, auto simp: fetch.simps nth_of.simps) done lemma adjust_fetch_norm: "\st > 0; st \ length tp div 2; fetch ap st b = (aa, ns); ns \ 0\ \ fetch (adjust0 ap) st b = (aa, ns)" apply(case_tac b, auto simp: fetch.simps nth_of.simps nth_map split: if_splits) apply(case_tac [!] st, auto simp: fetch.simps nth_of.simps) done declare adjust.simps[simp del] lemma adjust_step_eq: assumes exec: "step0 (st,l,r) ap = (st', l', r')" and wf_tm: "tm_wf (ap, 0)" and notfinal: "st' > 0" shows "step0 (st, l, r) (adjust0 ap) = (st', l', r')" using assms proof - have "st > 0" using assms by(case_tac st, simp_all add: step.simps fetch.simps) moreover hence "st \ (length ap) div 2" using assms apply(case_tac "st \ (length ap) div 2", simp) apply(case_tac st, auto simp: step.simps fetch.simps) apply(case_tac "read r", simp_all add: fetch.simps nth_of.simps adjust.simps tm_wf.simps split: if_splits) apply(auto simp: mod_ex2) done ultimately have "fetch (adjust0 ap) st (read r) = fetch ap st (read r)" using assms apply(case_tac "fetch ap st (read r)") apply(drule_tac adjust_fetch_norm, simp_all) apply(simp add: step.simps) done thus "?thesis" using exec by(simp add: step.simps) qed declare adjust.simps[simp del] lemma adjust_steps_eq: assumes exec: "steps0 (st,l,r) ap stp = (st', l', r')" and wf_tm: "tm_wf (ap, 0)" and notfinal: "st' > 0" shows "steps0 (st, l, r) (adjust0 ap) stp = (st', l', r')" using exec notfinal proof(induct stp arbitrary: st' l' r') case 0 thus "?case" by(simp add: steps.simps) next case (Suc stp st' l' r') have ind: "\st' l' r'. \steps0 (st, l, r) ap stp = (st', l', r'); 0 < st'\ \ steps0 (st, l, r) (adjust0 ap) stp = (st', l', r')" by fact have h: "steps0 (st, l, r) ap (Suc stp) = (st', l', r')" by fact have g: "0 < st'" by fact obtain st'' l'' r'' where a: "steps0 (st, l, r) ap stp = (st'', l'', r'')" by (metis prod_cases3) hence c:"0 < st''" using h g apply(simp add: step_red) apply(case_tac st'', auto) done hence b: "steps0 (st, l, r) (adjust0 ap) stp = (st'', l'', r'')" using a by(rule_tac ind, simp_all) thus "?case" using assms a b h g apply(simp add: step_red) apply(rule_tac adjust_step_eq, simp_all) done qed lemma adjust_halt_eq: assumes exec: "steps0 (1, l, r) ap stp = (0, l', r')" and tm_wf: "tm_wf (ap, 0)" shows "\ stp. steps0 (Suc 0, l, r) (adjust0 ap) stp = (Suc (length ap div 2), l', r')" proof - have "\ stp. \ is_final (steps0 (1, l, r) ap stp) \ (steps0 (1, l, r) ap (Suc stp) = (0, l', r'))" using exec by(erule_tac before_final) then obtain stpa where a: "\ is_final (steps0 (1, l, r) ap stpa) \ (steps0 (1, l, r) ap (Suc stpa) = (0, l', r'))" .. obtain sa la ra where b:"steps0 (1, l, r) ap stpa = (sa, la, ra)" by (metis prod_cases3) hence c: "steps0 (Suc 0, l, r) (adjust0 ap) stpa = (sa, la, ra)" using assms a apply(rule_tac adjust_steps_eq, simp_all) done have d: "sa \ length ap div 2" using steps_in_range[of "(l, r)" ap stpa] a tm_wf b by(simp) obtain ac ns where e: "fetch ap sa (read ra) = (ac, ns)" by (metis prod.exhaust) hence f: "ns = 0" using b a apply(simp add: step_red step.simps) done have k: "fetch (adjust0 ap) sa (read ra) = (ac, Suc (length ap div 2))" using a b c d e f apply(rule_tac adjust_fetch0, simp_all) done from a b e f k and c show "?thesis" apply(rule_tac x = "Suc stpa" in exI) apply(simp add: step_red, auto) apply(simp add: step.simps) done qed declare tm_wf.simps[simp del] lemma tm_wf_t_twice_compile [simp]: "tm_wf (t_twice_compile, 0)" apply(simp only: t_twice_compile_def) apply(rule_tac wf_tm_from_abacus, simp) done lemma t_twice_change_term_state: "\ stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) t_twice stp = (Suc t_twice_len, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" proof - have "\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 = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" by(rule_tac t_twice_correct) then obtain stp ln rn where " 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 = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" by blast hence "\ stp. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (adjust0 t_twice_compile) stp = (Suc (length t_twice_compile div 2), Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" apply(rule_tac stp = stp in adjust_halt_eq) apply(simp add: t_twice_compile_def, auto) done then obtain stpb where "steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (adjust0 t_twice_compile) stpb = (Suc (length t_twice_compile div 2), Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" .. thus "?thesis" apply(simp add: t_twice_def t_twice_len_def) by metis qed lemma length_t_wcode_main_first_part_even[intro]: "length t_wcode_main_first_part mod 2 = 0" apply(auto simp: t_wcode_main_first_part_def) done lemma t_twice_append_pre: "steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) t_twice stp = (Suc t_twice_len, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn)) \ steps0 (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ ([(L, 1), (L, 1)] @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" by(rule_tac tm_append_shift_append_steps, auto) lemma t_twice_append: "\ stp ln rn. steps0 (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ ([(L, 1), (L, 1)] @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" using t_twice_change_term_state[of ires rs n] apply(erule_tac exE) apply(erule_tac exE) apply(erule_tac exE) apply(drule_tac t_twice_append_pre) apply(rename_tac stp ln rn) apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) apply(simp) done lemma mopup_mod2: "length (mopup k) mod 2 = 0" by(auto simp: mopup.simps) lemma fetch_t_wcode_main_Oc[simp]: "fetch t_wcode_main (Suc (t_twice_len + length t_wcode_main_first_part div 2)) Oc = (L, Suc 0)" apply(subgoal_tac "length (t_twice) mod 2 = 0") apply(simp add: t_wcode_main_def nth_append fetch.simps t_wcode_main_first_part_def nth_of.simps t_twice_len_def, auto) apply(simp add: t_twice_def t_twice_compile_def) using mopup_mod2[of 1] apply(simp) done lemma wcode_jump1: "\ stp ln rn. steps0 (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\(m) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(n)) t_wcode_main stp = (Suc 0, Bk\(ln) @ Bk # ires, Bk # Oc\(Suc (2 * rs)) @ Bk\(rn))" apply(rule_tac x = "Suc 0" in exI, rule_tac x = "m" in exI, rule_tac x = n in exI) apply(simp add: steps.simps step.simps exp_ind) apply(case_tac m, simp_all) apply(simp add: exp_ind[THEN sym]) done lemma wcode_main_first_part_len[simp]: "length t_wcode_main_first_part = 24" apply(simp add: t_wcode_main_first_part_def) done lemma wcode_double_case: shows "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (2 * rs + 2)) @ Bk\(rn))" (is "\stp ln rn. ?tm stp ln rn") proof - from wcode_double_case_first_correctness[of ires rs m n] obtain na ln rn where "steps0 (Suc 0, Bk # Bk \ m @ Oc # Oc # ires, Bk # Oc # Oc \ rs @ Bk \ n) t_wcode_main na = (13, Bk # Bk # Bk \ ln @ Oc # ires, Oc # Oc # Oc \ rs @ Bk \ rn)" by(auto simp: wcode_double_case_inv.simps wcode_before_double.simps) hence "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (13, Bk # Bk # Bk\(ln) @ Oc # ires, Oc\(Suc (Suc rs)) @ Bk\(rn))" by(case_tac "steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main na", auto) from this obtain stpa lna rna where stp1: "steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stpa = (13, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (Suc rs)) @ Bk\(rna))" by blast from t_twice_append[of "Bk\(lna) @ Oc # ires" "Suc rs" rna] obtain stp ln rn where "steps0 (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # Bk \ lna @ Oc # ires, Oc \ Suc (Suc rs) @ Bk \ rna) (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)] @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)]) stp = (Suc t_twice_len + length t_wcode_main_first_part div 2, Bk \ ln @ Bk # Bk # Bk \ lna @ Oc # ires, Oc \ Suc (2 * Suc rs) @ Bk \ rn)" by blast hence "\ stp ln rn. steps0 (13, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (Suc rs)) @ Bk\(rna)) t_wcode_main stp = (13 + t_twice_len, Bk # Bk # Bk\(ln) @ Oc # ires, Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rn))" using t_twice_append[of "Bk\(lna) @ Oc # ires" "Suc rs" rna] apply(simp) 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] 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 = (13 + t_twice_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rnb))" by blast from wcode_jump1[of lnb "Oc # ires" "Suc rs" rnb] obtain stp ln rn where "steps0 (Suc t_twice_len + length t_wcode_main_first_part div 2, Bk \ lnb @ Bk # Bk # Oc # ires, Oc \ Suc (2 * Suc rs) @ Bk \ rnb) t_wcode_main stp = (Suc 0, Bk \ ln @ Bk # Oc # ires, Bk # Oc \ Suc (2 * Suc rs) @ Bk \ rn)" by metis hence "steps0 (13 + t_twice_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rnb)) t_wcode_main stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rn))" apply(auto simp add: t_wcode_main_def) apply(subgoal_tac "Bk\(lnb) @ Bk # Bk # Oc # ires = Bk # Bk # Bk\(lnb) @ Oc # ires", simp) apply(simp add: replicate_Suc[THEN sym] exp_ind[THEN sym] del: replicate_Suc) apply(simp) apply(simp add: replicate_Suc[THEN sym] exp_ind del: replicate_Suc) done hence "\stp ln rn. steps0 (13 + t_twice_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rnb)) t_wcode_main stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rn))" by blast from this obtain stpc lnc rnc where stp3: "steps0 (13 + t_twice_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rnb)) t_wcode_main stpc = (Suc 0, Bk # Bk\(lnc) @ Oc # ires, Bk # Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rnc))" by blast from stp1 stp2 stp3 have "?tm (stpa + stpb + stpc) lnc rnc" by simp thus "?thesis" by blast qed (* Begin: fourtime_case*) fun wcode_on_left_moving_2_B :: "bin_inv_t" where "wcode_on_left_moving_2_B ires rs (l, r) = (\ ml mr rn. l = Bk\(ml) @ Oc # Bk # Oc # ires \ r = Bk\(mr) @ Oc\(Suc rs) @ Bk\(rn) \ ml + mr > Suc 0 \ mr > 0)" fun wcode_on_left_moving_2_O :: "bin_inv_t" where "wcode_on_left_moving_2_O ires rs (l, r) = (\ ln rn. l = Bk # Oc # ires \ r = Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_on_left_moving_2 :: "bin_inv_t" where "wcode_on_left_moving_2 ires rs (l, r) = (wcode_on_left_moving_2_B ires rs (l, r) \ wcode_on_left_moving_2_O ires rs (l, r))" fun wcode_on_checking_2 :: "bin_inv_t" where "wcode_on_checking_2 ires rs (l, r) = (\ ln rn. l = Oc#ires \ r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_goon_checking :: "bin_inv_t" where "wcode_goon_checking ires rs (l, r) = (\ ln rn. l = ires \ r = Oc # Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_right_move :: "bin_inv_t" where "wcode_right_move ires rs (l, r) = (\ ln rn. l = Oc # ires \ r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_erase2 :: "bin_inv_t" where "wcode_erase2 ires rs (l, r) = (\ ln rn. l = Bk # Oc # ires \ tl r = Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_on_right_moving_2 :: "bin_inv_t" where "wcode_on_right_moving_2 ires rs (l, r) = (\ ml mr rn. l = Bk\(ml) @ Oc # ires \ r = Bk\(mr) @ Oc\(Suc rs) @ Bk\(rn) \ ml + mr > Suc 0)" fun wcode_goon_right_moving_2 :: "bin_inv_t" where "wcode_goon_right_moving_2 ires rs (l, r) = (\ ml mr ln rn. l = Oc\(ml) @ Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = Suc rs)" fun wcode_backto_standard_pos_2_B :: "bin_inv_t" where "wcode_backto_standard_pos_2_B ires rs (l, r) = (\ ln rn. l = Bk # Bk\(ln) @ Oc # ires \ r = Bk # Oc\(Suc (Suc rs)) @ Bk\(rn))" fun wcode_backto_standard_pos_2_O :: "bin_inv_t" where "wcode_backto_standard_pos_2_O ires rs (l, r) = (\ ml mr ln rn. l = Oc\(ml )@ Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = (Suc (Suc rs)) \ mr > 0)" fun wcode_backto_standard_pos_2 :: "bin_inv_t" where "wcode_backto_standard_pos_2 ires rs (l, r) = (wcode_backto_standard_pos_2_O ires rs (l, r) \ wcode_backto_standard_pos_2_B ires rs (l, r))" fun wcode_before_fourtimes :: "bin_inv_t" where "wcode_before_fourtimes ires rs (l, r) = (\ ln rn. l = Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(Suc (Suc rs)) @ Bk\(rn))" declare wcode_on_left_moving_2_B.simps[simp del] wcode_on_left_moving_2.simps[simp del] wcode_on_left_moving_2_O.simps[simp del] wcode_on_checking_2.simps[simp del] wcode_goon_checking.simps[simp del] wcode_right_move.simps[simp del] wcode_erase2.simps[simp del] wcode_on_right_moving_2.simps[simp del] wcode_goon_right_moving_2.simps[simp del] wcode_backto_standard_pos_2_B.simps[simp del] wcode_backto_standard_pos_2_O.simps[simp del] wcode_backto_standard_pos_2.simps[simp del] lemmas wcode_fourtimes_invs = wcode_on_left_moving_2_B.simps wcode_on_left_moving_2.simps wcode_on_left_moving_2_O.simps wcode_on_checking_2.simps wcode_goon_checking.simps wcode_right_move.simps wcode_erase2.simps wcode_on_right_moving_2.simps wcode_goon_right_moving_2.simps wcode_backto_standard_pos_2_B.simps wcode_backto_standard_pos_2_O.simps wcode_backto_standard_pos_2.simps fun wcode_fourtimes_case_inv :: "nat \ bin_inv_t" where "wcode_fourtimes_case_inv st ires rs (l, r) = (if st = Suc 0 then wcode_on_left_moving_2 ires rs (l, r) else if st = Suc (Suc 0) then wcode_on_checking_2 ires rs (l, r) else if st = 7 then wcode_goon_checking ires rs (l, r) else if st = 8 then wcode_right_move ires rs (l, r) else if st = 9 then wcode_erase2 ires rs (l, r) else if st = 10 then wcode_on_right_moving_2 ires rs (l, r) else if st = 11 then wcode_goon_right_moving_2 ires rs (l, r) else if st = 12 then wcode_backto_standard_pos_2 ires rs (l, r) else if st = t_twice_len + 14 then wcode_before_fourtimes ires rs (l, r) else False)" declare wcode_fourtimes_case_inv.simps[simp del] fun wcode_fourtimes_case_state :: "config \ nat" where "wcode_fourtimes_case_state (st, l, r) = 13 - st" fun wcode_fourtimes_case_step :: "config \ nat" where "wcode_fourtimes_case_step (st, l, r) = (if st = Suc 0 then length l else if st = 9 then (if hd r = Oc then 1 else 0) else if st = 10 then length r else if st = 11 then length r else if st = 12 then length l else 0)" fun wcode_fourtimes_case_measure :: "config \ nat \ nat" where "wcode_fourtimes_case_measure (st, l, r) = (wcode_fourtimes_case_state (st, l, r), wcode_fourtimes_case_step (st, l, r))" definition wcode_fourtimes_case_le :: "(config \ config) set" where "wcode_fourtimes_case_le \ (inv_image lex_pair wcode_fourtimes_case_measure)" lemma wf_wcode_fourtimes_case_le[intro]: "wf wcode_fourtimes_case_le" by(auto simp: wcode_fourtimes_case_le_def) lemma nonempty_snd [simp]: "wcode_on_left_moving_2 ires rs (b, []) = False" "wcode_on_checking_2 ires rs (b, []) = False" "wcode_goon_checking ires rs (b, []) = False" "wcode_right_move ires rs (b, []) = False" "wcode_erase2 ires rs (b, []) = False" "wcode_on_right_moving_2 ires rs (b, []) = False" "wcode_backto_standard_pos_2 ires rs (b, []) = False" "wcode_on_checking_2 ires rs (b, Oc # list) = False" by(auto simp: wcode_fourtimes_invs) lemma wcode_on_left_moving_2[simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \ wcode_on_left_moving_2 ires rs (tl b, hd b # Bk # list)" apply(simp only: wcode_fourtimes_invs) apply(erule_tac disjE) apply(erule_tac exE)+ apply(simp add: gr1_conv_Suc exp_ind replicate_app_Cons_same split:hd_repeat_cases) apply (auto simp add: gr0_conv_Suc[symmetric] replicate_app_Cons_same split:hd_repeat_cases) by force+ lemma wcode_goon_checking_via_2 [simp]: "wcode_on_checking_2 ires rs (b, Bk # list) \ wcode_goon_checking ires rs (tl b, hd b # Bk # list)" unfolding wcode_fourtimes_invs by auto lemma wcode_erase2_via_move [simp]: "wcode_right_move ires rs (b, Bk # list) \ wcode_erase2 ires rs (Bk # b, list)" by (auto simp:wcode_fourtimes_invs ) auto lemma wcode_on_right_moving_2_via_erase2[simp]: "wcode_erase2 ires rs (b, Bk # list) \ wcode_on_right_moving_2 ires rs (Bk # b, list)" apply(auto simp:wcode_fourtimes_invs ) apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: exp_ind) by (metis replicate_Suc_iff_anywhere replicate_app_Cons_same) lemma wcode_on_right_moving_2_move_Bk[simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list) \ wcode_on_right_moving_2 ires rs (Bk # b, list)" apply(auto simp: wcode_fourtimes_invs) apply(rename_tac ml mr rn) apply(rule_tac x = "Suc ml" in exI, simp) apply(rule_tac x = "mr - 1" in exI, case_tac mr,auto) done lemma wcode_backto_standard_pos_2_via_right[simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \ wcode_backto_standard_pos_2 ires rs (b, Oc # list)" apply(simp add: wcode_fourtimes_invs, auto) by (metis add.right_neutral add_Suc_shift append_Cons list.sel(3) replicate.simps(1) replicate_Suc replicate_Suc_iff_anywhere self_append_conv2 tl_replicate zero_less_Suc) lemma wcode_on_checking_2_via_left[simp]: "wcode_on_left_moving_2 ires rs (b, Oc # list) \ wcode_on_checking_2 ires rs (tl b, hd b # Oc # list)" by(auto simp: wcode_fourtimes_invs) lemma wcode_backto_standard_pos_2_empty_via_right[simp]: "wcode_goon_right_moving_2 ires rs (b, []) \ wcode_backto_standard_pos_2 ires rs (b, [Oc])" apply(simp only: wcode_fourtimes_invs) apply(erule_tac exE)+ by(rule_tac disjI1,auto) lemma wcode_goon_checking_cases[simp]: "wcode_goon_checking ires rs (b, Oc # list) \ (b = [] \ wcode_right_move ires rs ([Oc], list)) \ (b \ [] \ wcode_right_move ires rs (Oc # b, list))" apply(simp only: wcode_fourtimes_invs) apply(erule_tac exE)+ apply(auto) done lemma wcode_right_move_no_Oc[simp]: "wcode_right_move ires rs (b, Oc # list) = False" apply(auto simp: wcode_fourtimes_invs) done lemma wcode_erase2_Bk_via_Oc[simp]: "wcode_erase2 ires rs (b, Oc # list) \ wcode_erase2 ires rs (b, Bk # list)" apply(auto simp: wcode_fourtimes_invs) done lemma wcode_goon_right_moving_2_Oc_move[simp]: "wcode_on_right_moving_2 ires rs (b, Oc # list) \ wcode_goon_right_moving_2 ires rs (Oc # b, list)" apply(auto simp: wcode_fourtimes_invs) apply(rule_tac x = "Suc 0" in exI, auto) apply(rule_tac x = "ml - 2" in exI) apply(case_tac ml, simp, case_tac "ml - 1", simp_all) done lemma wcode_backto_standard_pos_2_exists[simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list) \ (\ln. b = Bk # Bk\(ln) @ Oc # ires) \ (\rn. list = Oc\(Suc (Suc rs)) @ Bk\(rn))" by(simp add: wcode_fourtimes_invs) lemma wcode_goon_right_moving_2_move_Oc[simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \ wcode_goon_right_moving_2 ires rs (Oc # b, list)" apply(simp only:wcode_fourtimes_invs, auto) apply(rename_tac ml ln mr rn) apply(case_tac mr;force) done lemma wcode_backto_standard_pos_2_Oc_mv_hd[simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list) \ wcode_backto_standard_pos_2 ires rs (tl b, hd b # Oc # list)" apply(simp only: wcode_fourtimes_invs) apply(erule_tac disjE) apply(erule_tac exE)+ apply(rename_tac ml mr ln rn) by (case_tac ml, force,force,force) lemma nonempty_fst[simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \ b \ []" "wcode_on_checking_2 ires rs (b, Bk # list) \ b \ []" "wcode_goon_checking ires rs (b, Bk # list) = False" "wcode_right_move ires rs (b, Bk # list) \ b\ []" "wcode_erase2 ires rs (b, Bk # list) \ b \ []" "wcode_on_right_moving_2 ires rs (b, Bk # list) \ b \ []" "wcode_goon_right_moving_2 ires rs (b, Bk # list) \ b \ []" "wcode_backto_standard_pos_2 ires rs (b, Bk # list) \ b \ []" "wcode_on_left_moving_2 ires rs (b, Oc # list) \ b \ []" "wcode_goon_right_moving_2 ires rs (b, []) \ b \ []" "wcode_erase2 ires rs (b, Oc # list) \ b \ []" "wcode_on_right_moving_2 ires rs (b, Oc # list) \ b \ []" "wcode_goon_right_moving_2 ires rs (b, Oc # list) \ b \ []" "wcode_backto_standard_pos_2 ires rs (b, Oc # list) \ b \ []" by(auto simp: wcode_fourtimes_invs) lemma wcode_fourtimes_case_first_correctness: shows "let P = (\ (st, l, r). st = t_twice_len + 14) in let Q = (\ (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r)) in let f = (\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp) in \ n .P (f n) \ Q (f (n::nat))" proof - let ?P = "(\ (st, l, r). st = t_twice_len + 14)" let ?Q = "(\ (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r))" let ?f = "(\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp)" have "\ n . ?P (?f n) \ ?Q (?f (n::nat))" proof(rule_tac halt_lemma2) show "wf wcode_fourtimes_case_le" by auto next have "\ ?P (?f na) \ ?Q (?f na) \ ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_fourtimes_case_le" for na apply(cases "?f na", rule_tac impI) apply(simp add: step_red step.simps) apply(case_tac "snd (snd (?f na))", simp, case_tac [2] "hd (snd (snd (?f na)))", simp_all) apply(simp_all add: wcode_fourtimes_case_inv.simps wcode_fourtimes_case_le_def lex_pair_def split: if_splits) by(auto simp: wcode_backto_standard_pos_2.simps wcode_backto_standard_pos_2_O.simps wcode_backto_standard_pos_2_B.simps gr0_conv_Suc) thus "\ na. \ ?P (?f na) \ ?Q (?f na) \ ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_fourtimes_case_le" by auto next show "?Q (?f 0)" apply(simp add: steps.simps wcode_fourtimes_case_inv.simps) apply(simp add: wcode_on_left_moving_2.simps wcode_on_left_moving_2_B.simps wcode_on_left_moving_2_O.simps) apply(rule_tac x = "Suc m" in exI, simp ) apply(rule_tac x ="Suc 0" in exI, auto) done next show "\ ?P (?f 0)" apply(simp add: steps.simps) done qed thus "?thesis" apply(erule_tac exE, simp) done qed definition t_fourtimes_len :: "nat" where "t_fourtimes_len = (length t_fourtimes div 2)" lemma primerec_rec_fourtimes_1[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)) (tm_of abc_fourtimes @ shift (mopup 1) (length (tm_of abc_fourtimes) div 2)) stp = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" 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 (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 "terminate rec_fourtimes [rs]" apply(rule_tac primerec_terminate) by auto next show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc (length [rs]))" using h by(simp add: abc_fourtimes_def) qed thus "?thesis" apply(simp add: tape_of_list_def tape_of_nat_def fourtimes_lemma) done qed lemma wf_fourtimes[intro]: "tm_wf (t_fourtimes_compile, 0)" apply(simp only: t_fourtimes_compile_def) apply(rule_tac wf_tm_from_abacus, simp) done lemma t_fourtimes_change_term_state: "\ stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) t_fourtimes stp = (Suc t_fourtimes_len, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" proof - have "\stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (tm_of abc_fourtimes @ shift (mopup 1) ((length (tm_of abc_fourtimes) div 2))) stp = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" by(rule_tac t_fourtimes_correct) then obtain stp ln rn where "steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (tm_of abc_fourtimes @ shift (mopup 1) ((length (tm_of abc_fourtimes) div 2))) stp = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" by blast hence "\ stp. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (adjust0 t_fourtimes_compile) stp = (Suc (length t_fourtimes_compile div 2), Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" apply(rule_tac stp = stp in adjust_halt_eq) apply(simp add: t_fourtimes_compile_def, auto) done then obtain stpb where "steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (adjust0 t_fourtimes_compile) stpb = (Suc (length t_fourtimes_compile div 2), Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" .. thus "?thesis" apply(simp add: t_fourtimes_def t_fourtimes_len_def) by metis qed lemma length_t_twice_even[intro]: "is_even (length t_twice)" by(auto simp: t_twice_def t_twice_compile_def intro!:mopup_mod2) lemma t_fourtimes_append_pre: "steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) t_fourtimes stp = (Suc t_fourtimes_len, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn)) \ steps0 (Suc 0 + length (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) ((t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) @ shift t_fourtimes (length (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2) @ ([(L, 1), (L, 1)])) stp = ((Suc t_fourtimes_len) + length (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" using length_t_twice_even by(intro tm_append_shift_append_steps, auto) lemma split_26_even[simp]: "(26 + l::nat) div 2 = l div 2 + 13" by(simp) lemma t_twice_len_plust_14[simp]: "t_twice_len + 14 = 14 + length (shift t_twice 12) div 2" apply(simp add: t_twice_def t_twice_len_def) done lemma t_fourtimes_append: "\ stp ln rn. steps0 (Suc 0 + length (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) ((t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)]) stp = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" using t_fourtimes_change_term_state[of ires rs n] apply(erule_tac exE) apply(erule_tac exE) apply(erule_tac exE) apply(drule_tac t_fourtimes_append_pre) apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) apply(simp add: t_twice_len_def) done lemma even_fourtimes_len: "length t_fourtimes mod 2 = 0" apply(auto simp: t_fourtimes_def t_fourtimes_compile_def) by (metis mopup_mod2) lemma t_twice_even[simp]: "2 * (length t_twice div 2) = length t_twice" using length_t_twice_even by arith lemma t_fourtimes_even[simp]: "2 * (length t_fourtimes div 2) = length t_fourtimes" using even_fourtimes_len by arith lemma fetch_t_wcode_14_Oc: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) Oc = (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_t_twice_even t_fourtimes_len_def nth_append) by arith lemma fetch_t_wcode_14_Bk: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) Bk = (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_t_twice_even t_fourtimes_len_def nth_append) by arith lemma fetch_t_wcode_14 [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) b = (L, Suc 0)" apply(case_tac b, simp_all add:fetch_t_wcode_14_Bk fetch_t_wcode_14_Oc) done lemma wcode_jump2: "\ stp ln rn. steps0 (t_twice_len + 14 + t_fourtimes_len , Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (4 * rs + 4)) @ Bk\(rnb)) t_wcode_main stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (4 * rs + 4)) @ Bk\(rn))" apply(rule_tac x = "Suc 0" in exI) apply(simp add: steps.simps) apply(rule_tac x = lnb in exI, rule_tac x = rnb in exI) apply(simp add: step.simps) done lemma wcode_fourtimes_case: shows "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rn))" proof - have "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (t_twice_len + 14, Bk # Bk # Bk\(ln) @ Oc # ires, Oc\(Suc (rs + 1)) @ Bk\(rn))" using wcode_fourtimes_case_first_correctness[of ires rs m n] by (auto simp add: wcode_fourtimes_case_inv.simps) auto from this obtain stpa lna rna where stp1: "steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stpa = (t_twice_len + 14, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (rs + 1)) @ Bk\(rna))" by blast have "\stp ln rn. steps0 (t_twice_len + 14, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (rs + 1)) @ Bk\(rna)) t_wcode_main stp = (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\(ln) @ Oc # ires, Oc\(Suc (4*rs + 4)) @ Bk\(rn))" using t_fourtimes_append[of " Bk\(lna) @ Oc # ires" "rs + 1" rna] apply(erule_tac exE) apply(erule_tac exE) apply(erule_tac exE) apply(simp add: t_wcode_main_def) apply(rename_tac stp ln rn) 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] 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)) t_wcode_main stpb = (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (4*rs + 4)) @ Bk\(rnb))" by blast have "\stp ln rn. steps0 (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (4*rs + 4)) @ Bk\(rnb)) t_wcode_main stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rn))" apply(rule wcode_jump2) done from this obtain stpc lnc rnc where stp3: "steps0 (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (4*rs + 4)) @ Bk\(rnb)) t_wcode_main stpc = (Suc 0, Bk # Bk\(lnc) @ Oc # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rnc))" by blast from stp1 stp2 stp3 show "?thesis" apply(rule_tac x = "stpa + stpb + stpc" in exI, rule_tac x = lnc in exI, rule_tac x = rnc in exI) apply(simp) done qed fun wcode_on_left_moving_3_B :: "bin_inv_t" where "wcode_on_left_moving_3_B ires rs (l, r) = (\ ml mr rn. l = Bk\(ml) @ Oc # Bk # Bk # ires \ r = Bk\(mr) @ Oc\(Suc rs) @ Bk\(rn) \ ml + mr > Suc 0 \ mr > 0 )" fun wcode_on_left_moving_3_O :: "bin_inv_t" where "wcode_on_left_moving_3_O ires rs (l, r) = (\ ln rn. l = Bk # Bk # ires \ r = Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_on_left_moving_3 :: "bin_inv_t" where "wcode_on_left_moving_3 ires rs (l, r) = (wcode_on_left_moving_3_B ires rs (l, r) \ wcode_on_left_moving_3_O ires rs (l, r))" fun wcode_on_checking_3 :: "bin_inv_t" where "wcode_on_checking_3 ires rs (l, r) = (\ ln rn. l = Bk # ires \ r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_goon_checking_3 :: "bin_inv_t" where "wcode_goon_checking_3 ires rs (l, r) = (\ ln rn. l = ires \ r = Bk # Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_stop :: "bin_inv_t" where "wcode_stop ires rs (l, r) = (\ ln rn. l = Bk # ires \ r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_halt_case_inv :: "nat \ bin_inv_t" where "wcode_halt_case_inv st ires rs (l, r) = (if st = 0 then wcode_stop ires rs (l, r) else if st = Suc 0 then wcode_on_left_moving_3 ires rs (l, r) else if st = Suc (Suc 0) then wcode_on_checking_3 ires rs (l, r) else if st = 7 then wcode_goon_checking_3 ires rs (l, r) else False)" fun wcode_halt_case_state :: "config \ nat" where "wcode_halt_case_state (st, l, r) = (if st = 1 then 5 else if st = Suc (Suc 0) then 4 else if st = 7 then 3 else 0)" fun wcode_halt_case_step :: "config \ nat" where "wcode_halt_case_step (st, l, r) = (if st = 1 then length l else 0)" fun wcode_halt_case_measure :: "config \ nat \ nat" where "wcode_halt_case_measure (st, l, r) = (wcode_halt_case_state (st, l, r), wcode_halt_case_step (st, l, r))" definition wcode_halt_case_le :: "(config \ config) set" where "wcode_halt_case_le \ (inv_image lex_pair wcode_halt_case_measure)" lemma wf_wcode_halt_case_le[intro]: "wf wcode_halt_case_le" by(auto simp: wcode_halt_case_le_def) declare wcode_on_left_moving_3_B.simps[simp del] wcode_on_left_moving_3_O.simps[simp del] wcode_on_checking_3.simps[simp del] wcode_goon_checking_3.simps[simp del] wcode_on_left_moving_3.simps[simp del] wcode_stop.simps[simp del] lemmas wcode_halt_invs = wcode_on_left_moving_3_B.simps wcode_on_left_moving_3_O.simps wcode_on_checking_3.simps wcode_goon_checking_3.simps wcode_on_left_moving_3.simps wcode_stop.simps lemma wcode_on_left_moving_3_mv_Bk[simp]: "wcode_on_left_moving_3 ires rs (b, Bk # list) \ wcode_on_left_moving_3 ires rs (tl b, hd b # Bk # list)" apply(simp only: wcode_halt_invs) apply(erule_tac disjE) apply(erule_tac exE)+ apply(rename_tac ml mr rn) apply(case_tac ml, simp) apply(rule_tac x = "mr - 2" in exI, rule_tac x = rn in exI) apply(case_tac mr, force, simp add: exp_ind del: replicate_Suc) apply(case_tac "mr - 1", force, simp add: exp_ind del: replicate_Suc) apply force apply force done lemma wcode_goon_checking_3_cases[simp]: "wcode_goon_checking_3 ires rs (b, Bk # list) \ (b = [] \ wcode_stop ires rs ([Bk], list)) \ (b \ [] \ wcode_stop ires rs (Bk # b, list))" apply(auto simp: wcode_halt_invs) done lemma wcode_on_checking_3_mv_Oc[simp]: "wcode_on_left_moving_3 ires rs (b, Oc # list) \ wcode_on_checking_3 ires rs (tl b, hd b # Oc # list)" by(simp add:wcode_halt_invs) lemma wcode_3_nonempty[simp]: "wcode_on_left_moving_3 ires rs (b, []) = False" "wcode_on_checking_3 ires rs (b, []) = False" "wcode_goon_checking_3 ires rs (b, []) = False" "wcode_on_left_moving_3 ires rs (b, Oc # list) \ b \ []" "wcode_on_checking_3 ires rs (b, Oc # list) = False" "wcode_on_left_moving_3 ires rs (b, Bk # list) \ b \ []" "wcode_on_checking_3 ires rs (b, Bk # list) \ b \ []" "wcode_goon_checking_3 ires rs (b, Oc # list) = False" by(auto simp: wcode_halt_invs) lemma wcode_goon_checking_3_mv_Bk[simp]: "wcode_on_checking_3 ires rs (b, Bk # list) \ wcode_goon_checking_3 ires rs (tl b, hd b # Bk # list)" apply(auto simp: wcode_halt_invs) done lemma t_halt_case_correctness: shows "let P = (\ (st, l, r). st = 0) in let Q = (\ (st, l, r). wcode_halt_case_inv st ires rs (l, r)) in let f = (\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp) in \ n .P (f n) \ Q (f (n::nat))" proof - let ?P = "(\ (st, l, r). st = 0)" let ?Q = "(\ (st, l, r). wcode_halt_case_inv st ires rs (l, r))" let ?f = "(\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp)" have "\ n. ?P (?f n) \ ?Q (?f (n::nat))" proof(rule_tac halt_lemma2) show "wf wcode_halt_case_le" by auto next { fix na obtain a b c where abc:"?f na = (a,b,c)" by(cases "?f na",auto) hence "\ ?P (?f na) \ ?Q (?f na) \ ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_halt_case_le" apply(simp add: step.simps) apply(cases c;cases "hd c") apply(auto simp: wcode_halt_case_le_def lex_pair_def split: if_splits) done } thus "\ na. \ ?P (?f na) \ ?Q (?f na) \ ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_halt_case_le" by blast next show "?Q (?f 0)" apply(simp add: steps.simps wcode_halt_invs) apply(rule_tac x = "Suc m" in exI, simp) apply(rule_tac x = "Suc 0" in exI, auto) done next show "\ ?P (?f 0)" apply(simp add: steps.simps) done qed thus "?thesis" apply(auto) done qed declare wcode_halt_case_inv.simps[simp del] lemma leading_Oc[intro]: "\ xs. ( :: cell list) = Oc # xs" apply(case_tac "rev list", simp) apply(simp add: tape_of_nl_cons) done lemma wcode_halt_case: "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" proof - let ?P = "\(st, l, r). st = 0" let ?Q = "\(st, l, r). wcode_halt_case_inv st ires rs (l, r)" let ?f = "steps0 (Suc 0, Bk # Bk \ m @ Oc # Bk # Bk # ires, Bk # Oc \ Suc rs @ Bk \ n) t_wcode_main" from t_halt_case_correctness[of ires rs m n] obtain n where "?P (?f n) \ ?Q (?f n)" by metis thus ?thesis apply(simp add: wcode_halt_case_inv.simps wcode_stop.simps) apply(case_tac "steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main n") apply(auto simp: wcode_halt_case_inv.simps wcode_stop.simps) by auto qed lemma bl_bin_one[simp]: "bl_bin [Oc] = 1" apply(simp add: bl_bin.simps) done lemma twice_power[intro]: "2 * 2 ^ a = Suc (Suc (2 * bl_bin (Oc \ a)))" apply(induct a, auto simp: bl_bin.simps) done declare replicate_Suc[simp del] lemma t_wcode_main_lemma_pre: "\args \ []; lm = \ \ \ stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin lm + rs * 2^(length lm - 1) ) @ Bk\(rn))" proof(induct "length args" arbitrary: args lm rs m n, simp) fix x args lm rs m n assume ind: "\args lm rs m n. \x = length args; (args::nat list) \ []; lm = \ \ \stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\(rn))" and h: "Suc x = length args" "(args::nat list) \ []" "lm = " from h have "\ (a::nat) xs. args = xs @ [a]" apply(rule_tac x = "last args" in exI) apply(rule_tac x = "butlast args" in exI, auto) done from this obtain a xs where "args = xs @ [a]" by blast from h and this show "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\(rn))" proof(case_tac "xs::nat list", simp) show "\stp ln rn. steps0 (Suc 0, Bk # Bk \ m @ Oc \ Suc a @ Bk # Bk # ires, Bk # Oc \ Suc rs @ Bk \ n) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ (bl_bin (Oc \ Suc a) + rs * 2 ^ a) @ Bk \ rn)" proof(induct "a" arbitrary: m n rs ires, simp) fix m n rs ires show "\stp ln rn. steps0 (Suc 0, Bk # Bk \ m @ Oc # Bk # Bk # ires, Bk # Oc \ Suc rs @ Bk \ n) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ Suc rs @ Bk \ rn)" apply(rule_tac wcode_halt_case) done next fix a m n rs ires assume ind2: "\m n rs ires. \stp ln rn. steps0 (Suc 0, Bk # Bk \ m @ Oc \ Suc a @ Bk # Bk # ires, Bk # Oc \ Suc rs @ Bk \ n) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ (bl_bin (Oc \ Suc a) + rs * 2 ^ a) @ Bk \ rn)" show " \stp ln rn. steps0 (Suc 0, Bk # Bk \ m @ Oc \ Suc (Suc a) @ Bk # Bk # ires, Bk # Oc \ Suc rs @ Bk \ n) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ (bl_bin (Oc \ Suc (Suc a)) + rs * 2 ^ Suc a) @ Bk \ rn)" proof - have "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (Suc 0, Bk # Bk\(ln) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (2 * rs + 2)) @ Bk\(rn))" apply(simp add: tape_of_nat) using wcode_double_case[of m "Oc\(a) @ Bk # Bk # ires" rs n] apply(simp add: replicate_Suc) done from this obtain stpa lna rna where stp1: "steps0 (Suc 0, Bk # Bk\(m) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stpa = (Suc 0, Bk # Bk\(lna) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (2 * rs + 2)) @ Bk\(rna))" by blast moreover have "\stp ln rn. steps0 (Suc 0, Bk # Bk\(lna) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (2 * rs + 2)) @ Bk\(rna)) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin () + (2*rs + 2) * 2 ^ a) @ Bk\(rn))" using ind2[of lna ires "2*rs + 2" rna] by(simp add: tape_of_list_def tape_of_nat_def) from this obtain stpb lnb rnb where stp2: "steps0 (Suc 0, Bk # Bk\(lna) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (2 * rs + 2)) @ Bk\(rna)) t_wcode_main stpb = (0, Bk # ires, Bk # Oc # Bk\(lnb) @ Bk # Bk # Oc\(bl_bin () + (2*rs + 2) * 2 ^ a) @ Bk\(rnb))" by blast from stp1 and stp2 show "?thesis" apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp add: tape_of_nat_def) apply(simp add: bl_bin.simps replicate_Suc) apply(auto) done qed qed next fix aa list assume g: "Suc x = length args" "args \ []" "lm = " "args = xs @ [a::nat]" "xs = (aa::nat) # list" thus "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\(rn))" - proof(induct a arbitrary: m n rs args lm, simp_all add: tape_of_nl_rev del: subst_all subst_all', - simp only: tape_of_nl_cons_app1, simp del: subst_all subst_all') + proof(induct a arbitrary: m n rs args lm, simp_all add: tape_of_nl_rev del: subst_all, + simp only: tape_of_nl_cons_app1, simp del: subst_all) fix m n rs args lm have "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # rev (<(aa::nat) # list>) @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (Suc 0, Bk # Bk\(ln) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rn))" proof(simp add: tape_of_nl_rev) have "\ xs. () = Oc # xs" by auto from this obtain xs where "() = Oc # xs" .. thus "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (Suc 0, Bk # Bk\(ln) @ @ Bk # Bk # ires, Bk # Oc\(5 + 4 * rs) @ Bk\(rn))" apply(simp) using wcode_fourtimes_case[of m "xs @ Bk # Bk # ires" rs n] apply(simp) done qed from this obtain stpa lna rna where stp1: "steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # rev () @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stpa = (Suc 0, Bk # Bk\(lna) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rna))" by blast from g have "\ stp ln rn. steps0 (Suc 0, Bk # Bk\(lna) @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rna)) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin ()+ (4*rs + 4) * 2^(length () - 1) ) @ Bk\(rn))" apply(rule_tac args = "(aa::nat)#list" in ind, simp_all) done from this obtain stpb lnb rnb where stp2: "steps0 (Suc 0, Bk # Bk\(lna) @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rna)) t_wcode_main stpb = (0, Bk # ires, Bk # Oc # Bk\(lnb) @ Bk # Bk # Oc\(bl_bin ()+ (4*rs + 4) * 2^(length () - 1) ) @ Bk\(rnb))" by blast from stp1 and stp2 and h show "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin (Oc\(Suc aa) @ Bk # ) + rs * (2 * 2 ^ (aa + length ()))) @ Bk\(rn))" apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_rev) done next fix ab m n rs args lm assume ind2: "\ m n rs args lm. \lm = ; args = aa # list @ [ab]\ \ \stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin () + rs * 2 ^ (length () - Suc 0)) @ Bk\(rn))" and k: "args = aa # list @ [Suc ab]" "lm = " show "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (0, Bk # ires,Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin () + rs * 2 ^ (length () - Suc 0)) @ Bk\(rn))" proof(simp add: tape_of_nl_cons_app1) have "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc\(Suc (Suc ab)) @ Bk # @ Bk # Bk # ires, Bk # Oc # Oc\(rs) @ Bk\(n)) t_wcode_main stp = (Suc 0, Bk # Bk\(ln) @ Oc\(Suc ab) @ Bk # @ Bk # Bk # ires, Bk # Oc\(Suc (2*rs + 2)) @ Bk\(rn))" using wcode_double_case[of m "Oc\(ab) @ Bk # @ Bk # Bk # ires" rs n] apply(simp add: replicate_Suc) done from this obtain stpa lna rna where stp1: "steps0 (Suc 0, Bk # Bk\(m) @ Oc\(Suc (Suc ab)) @ Bk # @ Bk # Bk # ires, Bk # Oc # Oc\(rs) @ Bk\(n)) t_wcode_main stpa = (Suc 0, Bk # Bk\(lna) @ Oc\(Suc ab) @ Bk # @ Bk # Bk # ires, Bk # Oc\(Suc (2*rs + 2)) @ Bk\(rna))" by blast from k have "\ stp ln rn. steps0 (Suc 0, Bk # Bk\(lna) @ @ Bk # Bk # ires, Bk # Oc\(Suc (2*rs + 2)) @ Bk\(rna)) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin ( ) + (2*rs + 2)* 2^(length () - Suc 0)) @ Bk\(rn))" apply(rule_tac ind2, simp_all) done from this obtain stpb lnb rnb where stp2: "steps0 (Suc 0, Bk # Bk\(lna) @ @ Bk # Bk # ires, Bk # Oc\(Suc (2*rs + 2)) @ Bk\(rna)) t_wcode_main stpb = (0, Bk # ires, Bk # Oc # Bk\(lnb) @ Bk # Bk # Oc\(bl_bin ( ) + (2*rs + 2)* 2^(length () - Suc 0)) @ Bk\(rnb))" by blast from stp1 and stp2 show "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc\(Suc (Suc ab)) @ Bk # @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin (Oc\(Suc aa) @ Bk # ) + rs * (2 * 2 ^ (aa + length ()))) @ Bk\(rn))" apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_cons_app1 replicate_Suc) done qed qed qed qed definition t_wcode_prepare :: "instr list" where "t_wcode_prepare \ [(W1, 2), (L, 1), (L, 3), (R, 2), (R, 4), (W0, 3), (R, 4), (R, 5), (R, 6), (R, 5), (R, 7), (R, 5), (W1, 7), (L, 0)]" fun wprepare_add_one :: "nat \ nat list \ tape \ bool" where "wprepare_add_one m lm (l, r) = (\ rn. l = [] \ (r = @ Bk\(rn) \ r = Bk # @ Bk\(rn)))" fun wprepare_goto_first_end :: "nat \ nat list \ tape \ bool" where "wprepare_goto_first_end m lm (l, r) = (\ ml mr rn. l = Oc\(ml) \ r = Oc\(mr) @ Bk # @ Bk\(rn) \ ml + mr = Suc (Suc m))" fun wprepare_erase :: "nat \ nat list \ tape \ bool" where "wprepare_erase m lm (l, r) = (\ rn. l = Oc\(Suc m) \ tl r = Bk # @ Bk\(rn))" fun wprepare_goto_start_pos_B :: "nat \ nat list \ tape \ bool" where "wprepare_goto_start_pos_B m lm (l, r) = (\ rn. l = Bk # Oc\(Suc m) \ r = Bk # @ Bk\(rn))" fun wprepare_goto_start_pos_O :: "nat \ nat list \ tape \ bool" where "wprepare_goto_start_pos_O m lm (l, r) = (\ rn. l = Bk # Bk # Oc\(Suc m) \ r = @ Bk\(rn))" fun wprepare_goto_start_pos :: "nat \ nat list \ tape \ bool" where "wprepare_goto_start_pos m lm (l, r) = (wprepare_goto_start_pos_B m lm (l, r) \ wprepare_goto_start_pos_O m lm (l, r))" fun wprepare_loop_start_on_rightmost :: "nat \ nat list \ tape \ bool" where "wprepare_loop_start_on_rightmost m lm (l, r) = (\ rn mr. rev l @ r = Oc\(Suc m) @ Bk # Bk # @ Bk\(rn) \ l \ [] \ r = Oc\(mr) @ Bk\(rn))" fun wprepare_loop_start_in_middle :: "nat \ nat list \ tape \ bool" where "wprepare_loop_start_in_middle m lm (l, r) = (\ rn (mr:: nat) (lm1::nat list). rev l @ r = Oc\(Suc m) @ Bk # Bk # @ Bk\(rn) \ l \ [] \ r = Oc\(mr) @ Bk # @ Bk\(rn) \ lm1 \ [])" fun wprepare_loop_start :: "nat \ nat list \ tape \ bool" where "wprepare_loop_start m lm (l, r) = (wprepare_loop_start_on_rightmost m lm (l, r) \ wprepare_loop_start_in_middle m lm (l, r))" fun wprepare_loop_goon_on_rightmost :: "nat \ nat list \ tape \ bool" where "wprepare_loop_goon_on_rightmost m lm (l, r) = (\ rn. l = Bk # @ Bk # Bk # Oc\(Suc m) \ r = Bk\(rn))" fun wprepare_loop_goon_in_middle :: "nat \ nat list \ tape \ bool" where "wprepare_loop_goon_in_middle m lm (l, r) = (\ rn (mr:: nat) (lm1::nat list). rev l @ r = Oc\(Suc m) @ Bk # Bk # @ Bk\(rn) \ l \ [] \ (if lm1 = [] then r = Oc\(mr) @ Bk\(rn) else r = Oc\(mr) @ Bk # @ Bk\(rn)) \ mr > 0)" fun wprepare_loop_goon :: "nat \ nat list \ tape \ bool" where "wprepare_loop_goon m lm (l, r) = (wprepare_loop_goon_in_middle m lm (l, r) \ wprepare_loop_goon_on_rightmost m lm (l, r))" fun wprepare_add_one2 :: "nat \ nat list \ tape \ bool" where "wprepare_add_one2 m lm (l, r) = (\ rn. l = Bk # Bk # @ Bk # Bk # Oc\(Suc m) \ (r = [] \ tl r = Bk\(rn)))" fun wprepare_stop :: "nat \ nat list \ tape \ bool" where "wprepare_stop m lm (l, r) = (\ rn. l = Bk # @ Bk # Bk # Oc\(Suc m) \ r = Bk # Oc # Bk\(rn))" fun wprepare_inv :: "nat \ nat \ nat list \ tape \ bool" where "wprepare_inv st m lm (l, r) = (if st = 0 then wprepare_stop m lm (l, r) else if st = Suc 0 then wprepare_add_one m lm (l, r) else if st = Suc (Suc 0) then wprepare_goto_first_end m lm (l, r) else if st = Suc (Suc (Suc 0)) then wprepare_erase m lm (l, r) else if st = 4 then wprepare_goto_start_pos m lm (l, r) else if st = 5 then wprepare_loop_start m lm (l, r) else if st = 6 then wprepare_loop_goon m lm (l, r) else if st = 7 then wprepare_add_one2 m lm (l, r) else False)" fun wprepare_stage :: "config \ nat" where "wprepare_stage (st, l, r) = (if st \ 1 \ st \ 4 then 3 else if st = 5 \ st = 6 then 2 else 1)" fun wprepare_state :: "config \ nat" where "wprepare_state (st, l, r) = (if st = 1 then 4 else if st = Suc (Suc 0) then 3 else if st = Suc (Suc (Suc 0)) then 2 else if st = 4 then 1 else if st = 7 then 2 else 0)" fun wprepare_step :: "config \ nat" where "wprepare_step (st, l, r) = (if st = 1 then (if hd r = Oc then Suc (length l) else 0) else if st = Suc (Suc 0) then length r else if st = Suc (Suc (Suc 0)) then (if hd r = Oc then 1 else 0) else if st = 4 then length r else if st = 5 then Suc (length r) else if st = 6 then (if r = [] then 0 else Suc (length r)) else if st = 7 then (if (r \ [] \ hd r = Oc) then 0 else 1) else 0)" fun wcode_prepare_measure :: "config \ nat \ nat \ nat" where "wcode_prepare_measure (st, l, r) = (wprepare_stage (st, l, r), wprepare_state (st, l, r), wprepare_step (st, l, r))" definition wcode_prepare_le :: "(config \ config) set" where "wcode_prepare_le \ (inv_image lex_triple wcode_prepare_measure)" lemma wf_wcode_prepare_le[intro]: "wf wcode_prepare_le" by(auto intro:wf_inv_image simp: wcode_prepare_le_def lex_triple_def) declare wprepare_add_one.simps[simp del] wprepare_goto_first_end.simps[simp del] wprepare_erase.simps[simp del] wprepare_goto_start_pos.simps[simp del] wprepare_loop_start.simps[simp del] wprepare_loop_goon.simps[simp del] wprepare_add_one2.simps[simp del] lemmas wprepare_invs = wprepare_add_one.simps wprepare_goto_first_end.simps wprepare_erase.simps wprepare_goto_start_pos.simps wprepare_loop_start.simps wprepare_loop_goon.simps wprepare_add_one2.simps declare wprepare_inv.simps[simp del] lemma fetch_t_wcode_prepare[simp]: "fetch t_wcode_prepare (Suc 0) Bk = (W1, 2)" "fetch t_wcode_prepare (Suc 0) Oc = (L, 1)" "fetch t_wcode_prepare (Suc (Suc 0)) Bk = (L, 3)" "fetch t_wcode_prepare (Suc (Suc 0)) Oc = (R, 2)" "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Bk = (R, 4)" "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Oc = (W0, 3)" "fetch t_wcode_prepare 4 Bk = (R, 4)" "fetch t_wcode_prepare 4 Oc = (R, 5)" "fetch t_wcode_prepare 5 Oc = (R, 5)" "fetch t_wcode_prepare 5 Bk = (R, 6)" "fetch t_wcode_prepare 6 Oc = (R, 5)" "fetch t_wcode_prepare 6 Bk = (R, 7)" "fetch t_wcode_prepare 7 Oc = (L, 0)" "fetch t_wcode_prepare 7 Bk = (W1, 7)" unfolding fetch.simps t_wcode_prepare_def nth_of.simps numeral by auto lemma wprepare_add_one_nonempty_snd[simp]: "lm \ [] \ wprepare_add_one m lm (b, []) = False" apply(simp add: wprepare_invs) done lemma wprepare_goto_first_end_nonempty_snd[simp]: "lm \ [] \ wprepare_goto_first_end m lm (b, []) = False" apply(simp add: wprepare_invs) done lemma wprepare_erase_nonempty_snd[simp]: "lm \ [] \ wprepare_erase m lm (b, []) = False" apply(simp add: wprepare_invs) done lemma wprepare_goto_start_pos_nonempty_snd[simp]: "lm \ [] \ wprepare_goto_start_pos m lm (b, []) = False" apply(simp add: wprepare_invs) done lemma wprepare_loop_start_empty_nonempty_fst[simp]: "\lm \ []; wprepare_loop_start m lm (b, [])\ \ b \ []" apply(simp add: wprepare_invs) done lemma rev_eq: "rev xs = rev ys \ xs = ys" by auto lemma wprepare_loop_goon_Bk_empty_snd[simp]: "\lm \ []; wprepare_loop_start m lm (b, [])\ \ wprepare_loop_goon m lm (Bk # b, [])" apply(simp only: wprepare_invs) apply(erule_tac disjE) apply(rule_tac disjI2) apply(simp add: wprepare_loop_start_on_rightmost.simps wprepare_loop_goon_on_rightmost.simps, auto) apply(rule_tac rev_eq, simp add: tape_of_nl_rev) done lemma wprepare_loop_goon_nonempty_fst[simp]: "\lm \ []; wprepare_loop_goon m lm (b, [])\ \ b \ []" apply(simp only: wprepare_invs, auto) done lemma wprepare_add_one2_Bk_empty[simp]:"\lm \ []; wprepare_loop_goon m lm (b, [])\ \ wprepare_add_one2 m lm (Bk # b, [])" apply(simp only: wprepare_invs, auto split: if_splits) done lemma wprepare_add_one2_nonempty_fst[simp]: "wprepare_add_one2 m lm (b, []) \ b \ []" apply(simp only: wprepare_invs, auto) done lemma wprepare_add_one2_Oc[simp]: "wprepare_add_one2 m lm (b, []) \ wprepare_add_one2 m lm (b, [Oc])" apply(simp only: wprepare_invs, auto) done lemma Bk_not_tape_start[simp]: "(Bk # list = <(m::nat) # lm> @ ys) = False" apply(case_tac lm, auto simp: tape_of_nl_cons replicate_Suc) done lemma wprepare_goto_first_end_cases[simp]: "\lm \ []; wprepare_add_one m lm (b, Bk # list)\ \ (b = [] \ wprepare_goto_first_end m lm ([], Oc # list)) \ (b \ [] \ wprepare_goto_first_end m lm (b, Oc # list))" apply(simp only: wprepare_invs) apply(auto simp: tape_of_nl_cons split: if_splits) apply(cases lm, auto simp add:tape_of_list_def replicate_Suc) done lemma wprepare_goto_first_end_Bk_nonempty_fst[simp]: "wprepare_goto_first_end m lm (b, Bk # list) \ b \ []" apply(simp only: wprepare_invs , auto simp: replicate_Suc) done declare replicate_Suc[simp] lemma wprepare_erase_elem_Bk_rest[simp]: "wprepare_goto_first_end m lm (b, Bk # list) \ wprepare_erase m lm (tl b, hd b # Bk # list)" by(simp add: wprepare_invs) lemma wprepare_erase_Bk_nonempty_fst[simp]: "wprepare_erase m lm (b, Bk # list) \ b \ []" by(simp add: wprepare_invs) lemma wprepare_goto_start_pos_Bk[simp]: "wprepare_erase m lm (b, Bk # list) \ wprepare_goto_start_pos m lm (Bk # b, list)" apply(simp only: wprepare_invs, auto) done lemma wprepare_add_one_Bk_nonempty_snd[simp]: "\wprepare_add_one m lm (b, Bk # list)\ \ list \ []" apply(simp only: wprepare_invs) apply(case_tac lm, simp_all add: tape_of_list_def tape_of_nat_def, auto) done lemma wprepare_goto_first_end_nonempty_snd_tl[simp]: "\lm \ []; wprepare_goto_first_end m lm (b, Bk # list)\ \ list \ []" by(simp only: wprepare_invs, auto) lemma wprepare_erase_Bk_nonempty_list[simp]: "\lm \ []; wprepare_erase m lm (b, Bk # list)\ \ list \ []" apply(simp only: wprepare_invs, auto) done lemma wprepare_goto_start_pos_Bk_nonempty[simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Bk # list)\ \ list \ []" by(cases lm;cases list;simp only: wprepare_invs, auto) lemma wprepare_goto_start_pos_Bk_nonempty_fst[simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Bk # list)\ \ b \ []" apply(simp only: wprepare_invs) apply(auto) done lemma wprepare_loop_goon_Bk_nonempty[simp]: "\lm \ []; wprepare_loop_goon m lm (b, Bk # list)\ \ b \ []" apply(simp only: wprepare_invs, auto) done lemma wprepare_loop_goon_wprepare_add_one2_cases[simp]: "\lm \ []; wprepare_loop_goon m lm (b, Bk # list)\ \ (list = [] \ wprepare_add_one2 m lm (Bk # b, [])) \ (list \ [] \ wprepare_add_one2 m lm (Bk # b, list))" unfolding wprepare_invs apply(cases list;auto split:nat.split if_splits) by (metis list.sel(3) tl_replicate) lemma wprepare_add_one2_nonempty[simp]: "wprepare_add_one2 m lm (b, Bk # list) \ b \ []" apply(simp only: wprepare_invs, simp) done lemma wprepare_add_one2_cases[simp]: "wprepare_add_one2 m lm (b, Bk # list) \ (list = [] \ wprepare_add_one2 m lm (b, [Oc])) \ (list \ [] \ wprepare_add_one2 m lm (b, Oc # list))" apply(simp only: wprepare_invs, auto) done lemma wprepare_goto_first_end_cases_Oc[simp]: "wprepare_goto_first_end m lm (b, Oc # list) \ (b = [] \ wprepare_goto_first_end m lm ([Oc], list)) \ (b \ [] \ wprepare_goto_first_end m lm (Oc # b, list))" apply(simp only: wprepare_invs, auto) apply(rule_tac x = 1 in exI, auto) apply(rename_tac ml mr rn) apply(case_tac mr, simp_all add: ) apply(case_tac ml, simp_all add: ) apply(rule_tac x = "Suc ml" in exI, simp_all add: ) apply(rule_tac x = "mr - 1" in exI, simp) done lemma wprepare_erase_nonempty[simp]: "wprepare_erase m lm (b, Oc # list) \ b \ []" apply(simp only: wprepare_invs, auto simp: ) done lemma wprepare_erase_Bk[simp]: "wprepare_erase m lm (b, Oc # list) \ wprepare_erase m lm (b, Bk # list)" apply(simp only:wprepare_invs, auto simp: ) done lemma wprepare_goto_start_pos_Bk_move[simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Bk # list)\ \ wprepare_goto_start_pos m lm (Bk # b, list)" apply(simp only:wprepare_invs, auto) apply(case_tac [!] lm, simp, simp_all) done lemma wprepare_loop_start_b_nonempty[simp]: "wprepare_loop_start m lm (b, aa) \ b \ []" apply(simp only:wprepare_invs, auto) done lemma exists_exp_of_Bk[elim]: "Bk # list = Oc\(mr) @ Bk\(rn) \ \rn. list = Bk\(rn)" apply(case_tac mr, simp_all) apply(case_tac rn, simp_all) done lemma wprepare_loop_start_in_middle_Bk_False[simp]: "wprepare_loop_start_in_middle m lm (b, [Bk]) = False" by(auto) declare wprepare_loop_start_in_middle.simps[simp del] declare wprepare_loop_start_on_rightmost.simps[simp del] wprepare_loop_goon_in_middle.simps[simp del] wprepare_loop_goon_on_rightmost.simps[simp del] lemma wprepare_loop_goon_in_middle_Bk_False[simp]: "wprepare_loop_goon_in_middle m lm (Bk # b, []) = False" apply(simp add: wprepare_loop_goon_in_middle.simps, auto) done lemma wprepare_loop_goon_Bk[simp]: "\lm \ []; wprepare_loop_start m lm (b, [Bk])\ \ wprepare_loop_goon m lm (Bk # b, [])" unfolding wprepare_invs apply(auto simp add: wprepare_loop_goon_on_rightmost.simps wprepare_loop_start_on_rightmost.simps) apply(rule_tac rev_eq) apply(simp add: tape_of_nl_rev) apply(simp add: exp_ind replicate_Suc[THEN sym] del: replicate_Suc) done lemma wprepare_loop_goon_in_middle_Bk_False2[simp]: "wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista) \ wprepare_loop_goon_in_middle m lm (Bk # b, a # lista) = False" apply(auto simp: wprepare_loop_start_on_rightmost.simps wprepare_loop_goon_in_middle.simps) done lemma wprepare_loop_goon_on_rightbmost_Bk_False[simp]: "\lm \ []; wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)\ \ wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista)" apply(simp only: wprepare_loop_start_on_rightmost.simps wprepare_loop_goon_on_rightmost.simps, auto simp: tape_of_nl_rev) apply(simp add: replicate_Suc[THEN sym] exp_ind tape_of_nl_rev del: replicate_Suc) by (meson Cons_replicate_eq) lemma wprepare_loop_goon_in_middle_Bk_False3[simp]: assumes "lm \ []" "wprepare_loop_start_in_middle m lm (b, Bk # a # lista)" shows "wprepare_loop_goon_in_middle m lm (Bk # b, a # lista)" (is "?t1") and "wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista) = False" (is ?t2) proof - from assms obtain rn mr lm1 where *:"rev b @ Oc \ mr @ Bk # = Oc # Oc \ m @ Bk # Bk # " "b \ []" "Bk # a # lista = Oc \ mr @ Bk # @ Bk \ rn" "lm1 \ []" by(auto simp add: wprepare_loop_start_in_middle.simps) thus ?t1 apply(simp add: wprepare_loop_start_in_middle.simps wprepare_loop_goon_in_middle.simps, auto) apply(rule_tac x = rn in exI, simp) apply(case_tac mr, simp_all add: ) apply(case_tac lm1, simp) apply(rule_tac x = "Suc (hd lm1)" in exI, simp) apply(rule_tac x = "tl lm1" in exI) apply(case_tac "tl lm1", simp_all add: tape_of_list_def tape_of_nat_def) done from * show ?t2 apply(simp add: wprepare_loop_start_in_middle.simps wprepare_loop_goon_on_rightmost.simps del:split_head_repeat, auto simp del:split_head_repeat) apply(case_tac mr) apply(case_tac "lm1::nat list", simp_all, case_tac "tl lm1", simp_all) apply(auto simp add: tape_of_list_def ) apply(case_tac [!] rna, simp_all add: ) apply(case_tac mr, simp_all add: ) apply(case_tac lm1, simp, case_tac list, simp) apply(simp_all add: tape_of_nat_def) by (metis Bk_not_tape_start tape_of_list_def tape_of_nat_list.elims) qed lemma wprepare_loop_goon_Bk2[simp]: "\lm \ []; wprepare_loop_start m lm (b, Bk # a # lista)\ \ wprepare_loop_goon m lm (Bk # b, a # lista)" apply(simp add: wprepare_loop_start.simps wprepare_loop_goon.simps) apply(erule_tac disjE, simp, auto) done lemma start_2_goon: "\lm \ []; wprepare_loop_start m lm (b, Bk # list)\ \ (list = [] \ wprepare_loop_goon m lm (Bk # b, [])) \ (list \ [] \ wprepare_loop_goon m lm (Bk # b, list))" apply(case_tac list, auto) done lemma add_one_2_add_one: "wprepare_add_one m lm (b, Oc # list) \ (hd b = Oc \ (b = [] \ wprepare_add_one m lm ([], Bk # Oc # list)) \ (b \ [] \ wprepare_add_one m lm (tl b, Oc # Oc # list))) \ (hd b \ Oc \ (b = [] \ wprepare_add_one m lm ([], Bk # Oc # list)) \ (b \ [] \ wprepare_add_one m lm (tl b, hd b # Oc # list)))" unfolding wprepare_add_one.simps by auto lemma wprepare_loop_start_on_rightmost_Oc[simp]: "wprepare_loop_start_on_rightmost m lm (b, Oc # list) \ wprepare_loop_start_on_rightmost m lm (Oc # b, list)" apply(simp add: wprepare_loop_start_on_rightmost.simps) by (metis Cons_replicate_eq cell.distinct(1) list.sel(3) self_append_conv2 tl_append2 tl_replicate) lemma wprepare_loop_start_in_middle_Oc[simp]: assumes "wprepare_loop_start_in_middle m lm (b, Oc # list)" shows "wprepare_loop_start_in_middle m lm (Oc # b, list)" proof - from assms obtain mr lm1 rn where "rev b @ Oc \ mr @ Bk # = Oc # Oc \ m @ Bk # Bk # " "Oc # list = Oc \ mr @ Bk # @ Bk \ rn" "lm1 \ []" by(auto simp add: wprepare_loop_start_in_middle.simps) thus ?thesis apply(auto simp add: wprepare_loop_start_in_middle.simps) apply(rule_tac x = rn in exI, auto) apply(case_tac mr, simp, simp add: ) apply(rule_tac x = "mr - 1" in exI, simp) apply(rule_tac x = lm1 in exI, simp) done qed lemma start_2_start: "wprepare_loop_start m lm (b, Oc # list) \ wprepare_loop_start m lm (Oc # b, list)" apply(simp add: wprepare_loop_start.simps) apply(erule_tac disjE, simp_all ) done lemma wprepare_loop_goon_Oc_nonempty[simp]: "wprepare_loop_goon m lm (b, Oc # list) \ b \ []" apply(simp add: wprepare_loop_goon.simps wprepare_loop_goon_in_middle.simps wprepare_loop_goon_on_rightmost.simps) apply(auto) done lemma wprepare_goto_start_pos_Oc_nonempty[simp]: "wprepare_goto_start_pos m lm (b, Oc # list) \ b \ []" apply(simp add: wprepare_goto_start_pos.simps) done lemma wprepare_loop_goon_on_rightmost_Oc_False[simp]: "wprepare_loop_goon_on_rightmost m lm (b, Oc # list) = False" apply(simp add: wprepare_loop_goon_on_rightmost.simps) done lemma wprepare_loop1: "\rev b @ Oc\(mr) = Oc\(Suc m) @ Bk # Bk # ; b \ []; 0 < mr; Oc # list = Oc\(mr) @ Bk\(rn)\ \ wprepare_loop_start_on_rightmost m lm (Oc # b, list)" apply(simp add: wprepare_loop_start_on_rightmost.simps) apply(rule_tac x = rn in exI, simp) apply(case_tac mr, simp, simp) done lemma wprepare_loop2: "\rev b @ Oc\(mr) @ Bk # = Oc\(Suc m) @ Bk # Bk # ; b \ []; Oc # list = Oc\(mr) @ Bk # <(a::nat) # lista> @ Bk\(rn)\ \ wprepare_loop_start_in_middle m lm (Oc # b, list)" apply(simp add: wprepare_loop_start_in_middle.simps) apply(rule_tac x = rn in exI, simp) apply(case_tac mr, simp_all add: ) apply(rename_tac nat) apply(rule_tac x = nat in exI, simp) apply(rule_tac x = "a#lista" in exI, simp) done lemma wprepare_loop_goon_in_middle_cases[simp]: "wprepare_loop_goon_in_middle m lm (b, Oc # list) \ wprepare_loop_start_on_rightmost m lm (Oc # b, list) \ wprepare_loop_start_in_middle m lm (Oc # b, list)" apply(simp add: wprepare_loop_goon_in_middle.simps split: if_splits) apply(rename_tac lm1) apply(case_tac lm1, simp_all add: wprepare_loop1 wprepare_loop2) done lemma wprepare_add_one_b[simp]: "wprepare_add_one m lm (b, Oc # list) \ b = [] \ wprepare_add_one m lm ([], Bk # Oc # list)" "wprepare_loop_goon m lm (b, Oc # list) \ wprepare_loop_start m lm (Oc # b, list)" apply(auto simp add: wprepare_add_one.simps wprepare_loop_goon.simps wprepare_loop_start.simps) done lemma wprepare_loop_start_on_rightmost_Oc2[simp]: "wprepare_goto_start_pos m [a] (b, Oc # list) \ wprepare_loop_start_on_rightmost m [a] (Oc # b, list) " apply(auto simp: wprepare_goto_start_pos.simps wprepare_loop_start_on_rightmost.simps) apply(rename_tac rn) apply(rule_tac x = rn in exI, simp) apply(simp add: replicate_Suc[THEN sym] exp_ind del: replicate_Suc) done lemma wprepare_loop_start_in_middle_2_Oc[simp]: "wprepare_goto_start_pos m (a # aa # listaa) (b, Oc # list) \wprepare_loop_start_in_middle m (a # aa # listaa) (Oc # b, list)" apply(auto simp: wprepare_goto_start_pos.simps wprepare_loop_start_in_middle.simps) apply(rename_tac rn) apply(rule_tac x = rn in exI, simp) apply(simp add: exp_ind[THEN sym]) apply(rule_tac x = a in exI, rule_tac x = "aa#listaa" in exI, simp) apply(simp add: tape_of_nl_cons) done lemma wprepare_loop_start_Oc2[simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Oc # list)\ \ wprepare_loop_start m lm (Oc # b, list)" by(cases lm;cases "tl lm", auto simp add: wprepare_loop_start.simps) lemma wprepare_add_one2_Oc_nonempty[simp]: "wprepare_add_one2 m lm (b, Oc # list) \ b \ []" apply(auto simp: wprepare_add_one2.simps) done lemma add_one_2_stop: "wprepare_add_one2 m lm (b, Oc # list) \ wprepare_stop m lm (tl b, hd b # Oc # list)" apply(simp add: wprepare_add_one2.simps) done declare wprepare_stop.simps[simp del] lemma wprepare_correctness: assumes h: "lm \ []" shows "let P = (\ (st, l, r). st = 0) in let Q = (\ (st, l, r). wprepare_inv st m lm (l, r)) in let f = (\ stp. steps0 (Suc 0, [], ()) t_wcode_prepare stp) in \ n .P (f n) \ Q (f n)" proof - let ?P = "(\ (st, l, r). st = 0)" let ?Q = "(\ (st, l, r). wprepare_inv st m lm (l, r))" let ?f = "(\ stp. steps0 (Suc 0, [], ()) t_wcode_prepare stp)" have "\ n. ?P (?f n) \ ?Q (?f n)" proof(rule_tac halt_lemma2) show "\ n. \ ?P (?f n) \ ?Q (?f n) \ ?Q (?f (Suc n)) \ (?f (Suc n), ?f n) \ wcode_prepare_le" using h apply(rule_tac allI, rule_tac impI) apply(rename_tac n) apply(case_tac "?f n", simp add: step.simps) apply(rename_tac c) apply(case_tac c, simp, case_tac [2] aa) apply(simp_all add: wprepare_inv.simps wcode_prepare_le_def lex_triple_def lex_pair_def split: if_splits) (* slow *) apply(simp_all add: start_2_goon start_2_start add_one_2_add_one add_one_2_stop) apply(auto simp: wprepare_add_one2.simps) done qed (auto simp add: steps.simps wprepare_inv.simps wprepare_invs) thus "?thesis" apply(auto) done qed lemma tm_wf_t_wcode_prepare[intro]: "tm_wf (t_wcode_prepare, 0)" apply(simp add:tm_wf.simps t_wcode_prepare_def) done lemma is_28_even[intro]: "(28 + (length t_twice_compile + length t_fourtimes_compile)) mod 2 = 0" by(auto simp: t_twice_compile_def t_fourtimes_compile_def) lemma b_le_28[elim]: "(a, b) \ set t_wcode_main_first_part \ b \ (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2" apply(auto simp: t_wcode_main_first_part_def t_twice_def) done lemma tm_wf_change_termi: assumes "tm_wf (tp, 0)" shows "list_all (\(acn, st). (st \ Suc (length tp div 2))) (adjust0 tp)" proof - { fix acn st n assume "tp ! n = (acn, st)" "n < length tp" "0 < st" hence "(acn, st)\set tp" by (metis nth_mem) with assms tm_wf.simps have "st \ length tp div 2 + 0" by auto hence "st \ Suc (length tp div 2)" by auto } thus ?thesis by(auto simp: tm_wf.simps List.list_all_length adjust.simps split: if_splits prod.split) qed lemma tm_wf_shift: assumes "list_all (\(acn, st). (st \ y)) tp" shows "list_all (\(acn, st). (st \ y + off)) (shift tp off)" proof - have [dest!]:"\ P Q n. \n. Q n \ P n \ Q n \ P n" by metis from assms show ?thesis by(auto simp: tm_wf.simps List.list_all_length shift.simps) qed declare length_tp'[simp del] lemma length_mopup_1[simp]: "length (mopup (Suc 0)) = 16" apply(auto simp: mopup.simps) done lemma twice_plus_28_elim[elim]: "(a, b) \ set (shift (adjust0 t_twice_compile) 12) \ b \ (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2" apply(simp add: t_twice_compile_def t_fourtimes_compile_def) proof - assume g: "(a, b) \ set (shift (adjust (tm_of abc_twice @ shift (mopup (Suc 0)) (length (tm_of abc_twice) div 2)) (Suc ((length (tm_of abc_twice) + 16) div 2))) 12)" moreover have "length (tm_of abc_twice) mod 2 = 0" by auto moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto ultimately have "list_all (\(acn, st). (st \ (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)) (shift (adjust0 t_twice_compile) 12)" proof(auto simp add: mod_ex1 del: adjust.simps) assume "even (length (tm_of abc_twice))" then obtain q where q:"length (tm_of abc_twice) = 2 * q" by auto assume "even (length (tm_of abc_fourtimes))" then obtain qa where qa:"length (tm_of abc_fourtimes) = 2 * qa" by auto note h = q qa hence "list_all (\(acn, st). st \ (18 + (q + qa)) + 12) (shift (adjust0 t_twice_compile) 12)" proof(rule_tac tm_wf_shift t_twice_compile_def) have "list_all (\(acn, st). st \ Suc (length t_twice_compile div 2)) (adjust0 t_twice_compile)" by(rule_tac tm_wf_change_termi, auto) thus "list_all (\(acn, st). st \ 18 + (q + qa)) (adjust0 t_twice_compile)" using h apply(simp add: t_twice_compile_def, auto simp: List.list_all_length) done qed thus "list_all (\(acn, st). st \ 30 + (length (tm_of abc_twice) div 2 + length (tm_of abc_fourtimes) div 2)) (shift (adjust0 t_twice_compile) 12)" using h by simp qed thus "b \ (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2" using g apply(auto simp:t_twice_compile_def) apply(simp add: Ball_set[THEN sym]) apply(erule_tac x = "(a, b)" in ballE, simp, simp) done qed lemma length_plus_28_elim2[elim]: "(a, b) \ set (shift (adjust0 t_fourtimes_compile) (t_twice_len + 13)) \ b \ (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2" apply(simp add: t_twice_compile_def t_fourtimes_compile_def t_twice_len_def) proof - assume g: "(a, b) \ set (shift (adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) (length (tm_of abc_fourtimes) div 2)) (Suc ((length (tm_of abc_fourtimes) + 16) div 2))) (length t_twice div 2 + 13))" moreover have "length (tm_of abc_twice) mod 2 = 0" by auto moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto ultimately have "list_all (\(acn, st). (st \ (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)) (shift (adjust0 (tm_of abc_fourtimes @ shift (mopup (Suc 0)) (length (tm_of abc_fourtimes) div 2))) (length t_twice div 2 + 13))" proof(auto simp: mod_ex1 t_twice_def t_twice_compile_def) assume "even (length (tm_of abc_twice))" then obtain q where q:"length (tm_of abc_twice) = 2 * q" by auto assume "even (length (tm_of abc_fourtimes))" then obtain qa where qa:"length (tm_of abc_fourtimes) = 2 * qa" by auto note h = q qa hence "list_all (\(acn, st). st \ (9 + qa + (21 + q))) (shift (adjust0 (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)) (21 + q))" proof(rule_tac tm_wf_shift t_twice_compile_def) have "list_all (\(acn, st). st \ Suc (length (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa) div 2)) (adjust0 (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa))" apply(rule_tac tm_wf_change_termi) using wf_fourtimes h apply(simp add: t_fourtimes_compile_def) done thus "list_all (\(acn, st). st \ 9 + qa) (adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa) (Suc (length (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa) div 2)))" using h apply(simp) done qed thus "list_all (\(acn, st). st \ 30 + (length (tm_of abc_twice) div 2 + length (tm_of abc_fourtimes) div 2)) (shift (adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) (length (tm_of abc_fourtimes) div 2)) (9 + length (tm_of abc_fourtimes) div 2)) (21 + length (tm_of abc_twice) div 2))" apply(subgoal_tac "qa + q = q + qa") apply(simp add: h) apply(simp) done qed thus "b \ (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2" using g apply(simp add: Ball_set[THEN sym]) apply(erule_tac x = "(a, b)" in ballE, simp, simp) done qed lemma tm_wf_t_wcode_main[intro]: "tm_wf (t_wcode_main, 0)" by(auto simp: t_wcode_main_def tm_wf.simps t_twice_def t_fourtimes_def del: List.list_all_iff) declare tm_comp.simps[simp del] lemma prepare_mainpart_lemma: "args \ [] \ \ stp ln rn. steps0 (Suc 0, [], ) (t_wcode_prepare |+| t_wcode_main) stp = (0, Bk # Oc\(Suc m), Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin ()) @ Bk\(rn))" proof - let ?P1 = "(\ (l, r). (l::cell list) = [] \ r = )" let ?Q1 = "(\ (l, r). wprepare_stop m args (l, r))" let ?P2 = ?Q1 let ?Q2 = "(\ (l, r). (\ ln rn. l = Bk # Oc\(Suc m) \ r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin ()) @ Bk\(rn)))" let ?P3 = "\ tp. False" assume h: "args \ []" have "{?P1} t_wcode_prepare |+| t_wcode_main {?Q2}" proof(rule_tac Hoare_plus_halt) show "{?P1} t_wcode_prepare {?Q1}" proof(rule_tac Hoare_haltI, auto) show "\n. is_final (steps0 (Suc 0, [], ) t_wcode_prepare n) \ wprepare_stop m args holds_for steps0 (Suc 0, [], ) t_wcode_prepare n" using wprepare_correctness[of args m,OF h] apply(auto simp add: wprepare_inv.simps) by (metis holds_for.simps is_finalI) qed next show "{?P2} t_wcode_main {?Q2}" proof(rule_tac Hoare_haltI, auto) fix l r assume "wprepare_stop m args (l, r)" thus "\n. is_final (steps0 (Suc 0, l, r) t_wcode_main n) \ (\(l, r). l = Bk # Oc # Oc \ m \ (\ln rn. r = Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn)) holds_for steps0 (Suc 0, l, r) t_wcode_main n" proof(auto simp: wprepare_stop.simps) fix rn show " \n. is_final (steps0 (Suc 0, Bk # @ Bk # Bk # Oc # Oc \ m, Bk # Oc # Bk \ rn) t_wcode_main n) \ (\(l, r). l = Bk # Oc # Oc \ m \ (\ln rn. r = Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn)) holds_for steps0 (Suc 0, Bk # @ Bk # Bk # Oc # Oc \ m, Bk # Oc # Bk \ rn) t_wcode_main n" using t_wcode_main_lemma_pre[of "args" "" 0 "Oc\(Suc m)" 0 rn,OF h refl] apply(auto simp: tape_of_nl_rev) apply(rename_tac stp ln rna) apply(rule_tac x = stp in exI, auto) done qed qed next show "tm_wf0 t_wcode_prepare" by auto qed then obtain n where "\ tp. (case tp of (l, r) \ l = [] \ r = ) \ (is_final (steps0 (1, tp) (t_wcode_prepare |+| t_wcode_main) n) \ (\(l, r). \ln rn. l = Bk # Oc \ Suc m \ r = Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn) holds_for steps0 (1, tp) (t_wcode_prepare |+| t_wcode_main) n)" unfolding Hoare_halt_def by auto thus "?thesis" apply(rule_tac x = n in exI) apply(case_tac "(steps0 (Suc 0, [], ) (adjust0 t_wcode_prepare @ shift t_wcode_main (length t_wcode_prepare div 2)) n)") apply(auto simp: tm_comp.simps) done qed definition tinres :: "cell list \ cell list \ bool" where "tinres xs ys = (\n. xs = ys @ Bk \ n \ ys = xs @ Bk \ n)" lemma tinres_fetch_congr[simp]: "tinres r r' \ fetch t ss (read r) = fetch t ss (read r')" apply(simp add: fetch.simps, auto split: if_splits simp: tinres_def) using hd_replicate apply fastforce using hd_replicate apply fastforce done lemma nonempty_hd_tinres[simp]: "\tinres r r'; r \ []; r' \ []\ \ hd r = hd r'" apply(auto simp: tinres_def) done lemma tinres_nonempty[simp]: "\tinres r []; r \ []\ \ hd r = Bk" "\tinres [] r'; r' \ []\ \ hd r' = Bk" "\tinres r []; r \ []\ \ tinres (tl r) []" "tinres r r' \ tinres (b # r) (b # r')" by(auto simp: tinres_def) lemma ex_move_tl[intro]: "\na. tl r = tl (r @ Bk\(n)) @ Bk\(na) \ tl (r @ Bk\(n)) = tl r @ Bk\(na)" apply(case_tac r, simp) by(case_tac n, auto) lemma tinres_tails[simp]: "tinres r r' \ tinres (tl r) (tl r')" apply(auto simp: tinres_def) by(case_tac r', auto) lemma tinres_empty[simp]: "\tinres [] r'\ \ tinres [] (tl r')" "tinres r [] \ tinres (Bk # tl r) [Bk]" "tinres r [] \ tinres (Oc # tl r) [Oc]" by(auto simp: tinres_def) lemma tinres_step2: assumes "tinres r r'" "step0 (ss, l, r) t = (sa, la, ra)" "step0 (ss, l, r') t = (sb, lb, rb)" shows "la = lb \ tinres ra rb \ sa = sb" proof (cases "fetch t ss (read r')") case (Pair a b) have sa:"sa = sb" using assms Pair by(force simp: step.simps) have "la = lb \ tinres ra rb" using assms Pair by(cases a, auto simp: step.simps split: if_splits) thus ?thesis using sa by auto qed lemma tinres_steps2: "\tinres r r'; steps0 (ss, l, r) t stp = (sa, la, ra); steps0 (ss, l, r') t stp = (sb, lb, rb)\ \ la = lb \ tinres ra rb \ sa = sb" proof(induct stp arbitrary: sa la ra sb lb rb) case (Suc stp sa la ra sb lb rb) then show ?case apply(simp) apply(case_tac "(steps0 (ss, l, r) t stp)") apply(case_tac "(steps0 (ss, l, r') t stp)") proof - fix stp a b c aa ba ca assume ind: "\sa la ra sb lb rb. \steps0 (ss, l, r) t stp = (sa, la, ra); steps0 (ss, l, r') t stp = (sb, lb, rb)\ \ la = lb \ tinres ra rb \ sa = sb" and h: " tinres r r'" "step0 (steps0 (ss, l, r) t stp) t = (sa, la, ra)" "step0 (steps0 (ss, l, r') t stp) t = (sb, lb, rb)" "steps0 (ss, l, r) t stp = (a, b, c)" "steps0 (ss, l, r') t stp = (aa, ba, ca)" have "b = ba \ tinres c ca \ a = aa" apply(rule_tac ind, simp_all add: h) done thus "la = lb \ tinres ra rb \ sa = sb" apply(rule_tac l = b and r = c and ss = a and r' = ca and t = t in tinres_step2) using h apply(simp, simp, simp) done qed qed (simp add: steps.simps) definition t_wcode_adjust :: "instr list" where "t_wcode_adjust = [(W1, 1), (R, 2), (Nop, 2), (R, 3), (R, 3), (R, 4), (L, 8), (L, 5), (L, 6), (W0, 5), (L, 6), (R, 7), (W1, 2), (Nop, 7), (L, 9), (W0, 8), (L, 9), (L, 10), (L, 11), (L, 10), (R, 0), (L, 11)]" lemma fetch_t_wcode_adjust[simp]: "fetch t_wcode_adjust (Suc 0) Bk = (W1, 1)" "fetch t_wcode_adjust (Suc 0) Oc = (R, 2)" "fetch t_wcode_adjust (Suc (Suc 0)) Oc = (R, 3)" "fetch t_wcode_adjust (Suc (Suc (Suc 0))) Oc = (R, 4)" "fetch t_wcode_adjust (Suc (Suc (Suc 0))) Bk = (R, 3)" "fetch t_wcode_adjust 4 Bk = (L, 8)" "fetch t_wcode_adjust 4 Oc = (L, 5)" "fetch t_wcode_adjust 5 Oc = (W0, 5)" "fetch t_wcode_adjust 5 Bk = (L, 6)" "fetch t_wcode_adjust 6 Oc = (R, 7)" "fetch t_wcode_adjust 6 Bk = (L, 6)" "fetch t_wcode_adjust 7 Bk = (W1, 2)" "fetch t_wcode_adjust 8 Bk = (L, 9)" "fetch t_wcode_adjust 8 Oc = (W0, 8)" "fetch t_wcode_adjust 9 Oc = (L, 10)" "fetch t_wcode_adjust 9 Bk = (L, 9)" "fetch t_wcode_adjust 10 Bk = (L, 11)" "fetch t_wcode_adjust 10 Oc = (L, 10)" "fetch t_wcode_adjust 11 Oc = (L, 11)" "fetch t_wcode_adjust 11 Bk = (R, 0)" by(auto simp: fetch.simps t_wcode_adjust_def nth_of.simps numeral) fun wadjust_start :: "nat \ nat \ tape \ bool" where "wadjust_start m rs (l, r) = (\ ln rn. l = Bk # Oc\(Suc m) \ tl r = Oc # Bk\(ln) @ Bk # Oc\(Suc rs) @ Bk\(rn))" fun wadjust_loop_start :: "nat \ nat \ tape \ bool" where "wadjust_loop_start m rs (l, r) = (\ ln rn ml mr. l = Oc\(ml) @ Bk # Oc\(Suc m) \ r = Oc # Bk\(ln) @ Bk # Oc\(mr) @ Bk\(rn) \ ml + mr = Suc (Suc rs) \ mr > 0)" fun wadjust_loop_right_move :: "nat \ nat \ tape \ bool" where "wadjust_loop_right_move m rs (l, r) = (\ ml mr nl nr rn. l = Bk\(nl) @ Oc # Oc\(ml) @ Bk # Oc\(Suc m) \ r = Bk\(nr) @ Oc\(mr) @ Bk\(rn) \ ml + mr = Suc (Suc rs) \ mr > 0 \ nl + nr > 0)" fun wadjust_loop_check :: "nat \ nat \ tape \ bool" where "wadjust_loop_check m rs (l, r) = (\ ml mr ln rn. l = Oc # Bk\(ln) @ Bk # Oc # Oc\(ml) @ Bk # Oc\(Suc m) \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = (Suc rs))" fun wadjust_loop_erase :: "nat \ nat \ tape \ bool" where "wadjust_loop_erase m rs (l, r) = (\ ml mr ln rn. l = Bk\(ln) @ Bk # Oc # Oc\(ml) @ Bk # Oc\(Suc m) \ tl r = Oc\(mr) @ Bk\(rn) \ ml + mr = (Suc rs) \ mr > 0)" fun wadjust_loop_on_left_moving_O :: "nat \ nat \ tape \ bool" where "wadjust_loop_on_left_moving_O m rs (l, r) = (\ ml mr ln rn. l = Oc\(ml) @ Bk # Oc\(Suc m )\ r = Oc # Bk\(ln) @ Bk # Bk # Oc\(mr) @ Bk\(rn) \ ml + mr = Suc rs \ mr > 0)" fun wadjust_loop_on_left_moving_B :: "nat \ nat \ tape \ bool" where "wadjust_loop_on_left_moving_B m rs (l, r) = (\ ml mr nl nr rn. l = Bk\(nl) @ Oc # Oc\(ml) @ Bk # Oc\(Suc m) \ r = Bk\(nr) @ Bk # Bk # Oc\(mr) @ Bk\(rn) \ ml + mr = Suc rs \ mr > 0)" fun wadjust_loop_on_left_moving :: "nat \ nat \ tape \ bool" where "wadjust_loop_on_left_moving m rs (l, r) = (wadjust_loop_on_left_moving_O m rs (l, r) \ wadjust_loop_on_left_moving_B m rs (l, r))" fun wadjust_loop_right_move2 :: "nat \ nat \ tape \ bool" where "wadjust_loop_right_move2 m rs (l, r) = (\ ml mr ln rn. l = Oc # Oc\(ml) @ Bk # Oc\(Suc m) \ r = Bk\(ln) @ Bk # Bk # Oc\(mr) @ Bk\(rn) \ ml + mr = Suc rs \ mr > 0)" fun wadjust_erase2 :: "nat \ nat \ tape \ bool" where "wadjust_erase2 m rs (l, r) = (\ ln rn. l = Bk\(ln) @ Bk # Oc # Oc\(Suc rs) @ Bk # Oc\(Suc m) \ tl r = Bk\(rn))" fun wadjust_on_left_moving_O :: "nat \ nat \ tape \ bool" where "wadjust_on_left_moving_O m rs (l, r) = (\ rn. l = Oc\(Suc rs) @ Bk # Oc\(Suc m) \ r = Oc # Bk\(rn))" fun wadjust_on_left_moving_B :: "nat \ nat \ tape \ bool" where "wadjust_on_left_moving_B m rs (l, r) = (\ ln rn. l = Bk\(ln) @ Oc # Oc\(Suc rs) @ Bk # Oc\(Suc m) \ r = Bk\(rn))" fun wadjust_on_left_moving :: "nat \ nat \ tape \ bool" where "wadjust_on_left_moving m rs (l, r) = (wadjust_on_left_moving_O m rs (l, r) \ wadjust_on_left_moving_B m rs (l, r))" fun wadjust_goon_left_moving_B :: "nat \ nat \ tape \ bool" where "wadjust_goon_left_moving_B m rs (l, r) = (\ rn. l = Oc\(Suc m) \ r = Bk # Oc\(Suc (Suc rs)) @ Bk\(rn))" fun wadjust_goon_left_moving_O :: "nat \ nat \ tape \ bool" where "wadjust_goon_left_moving_O m rs (l, r) = (\ ml mr rn. l = Oc\(ml) @ Bk # Oc\(Suc m) \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = Suc (Suc rs) \ mr > 0)" fun wadjust_goon_left_moving :: "nat \ nat \ tape \ bool" where "wadjust_goon_left_moving m rs (l, r) = (wadjust_goon_left_moving_B m rs (l, r) \ wadjust_goon_left_moving_O m rs (l, r))" fun wadjust_backto_standard_pos_B :: "nat \ nat \ tape \ bool" where "wadjust_backto_standard_pos_B m rs (l, r) = (\ rn. l = [] \ r = Bk # Oc\(Suc m )@ Bk # Oc\(Suc (Suc rs)) @ Bk\(rn))" fun wadjust_backto_standard_pos_O :: "nat \ nat \ tape \ bool" where "wadjust_backto_standard_pos_O m rs (l, r) = (\ ml mr rn. l = Oc\(ml) \ r = Oc\(mr) @ Bk # Oc\(Suc (Suc rs)) @ Bk\(rn) \ ml + mr = Suc m \ mr > 0)" fun wadjust_backto_standard_pos :: "nat \ nat \ tape \ bool" where "wadjust_backto_standard_pos m rs (l, r) = (wadjust_backto_standard_pos_B m rs (l, r) \ wadjust_backto_standard_pos_O m rs (l, r))" fun wadjust_stop :: "nat \ nat \ tape \ bool" where "wadjust_stop m rs (l, r) = (\ rn. l = [Bk] \ r = Oc\(Suc m )@ Bk # Oc\(Suc (Suc rs)) @ Bk\(rn))" declare wadjust_start.simps[simp del] wadjust_loop_start.simps[simp del] wadjust_loop_right_move.simps[simp del] wadjust_loop_check.simps[simp del] wadjust_loop_erase.simps[simp del] wadjust_loop_on_left_moving.simps[simp del] wadjust_loop_right_move2.simps[simp del] wadjust_erase2.simps[simp del] wadjust_on_left_moving_O.simps[simp del] wadjust_on_left_moving_B.simps[simp del] wadjust_on_left_moving.simps[simp del] wadjust_goon_left_moving_B.simps[simp del] wadjust_goon_left_moving_O.simps[simp del] wadjust_goon_left_moving.simps[simp del] wadjust_backto_standard_pos.simps[simp del] wadjust_backto_standard_pos_B.simps[simp del] wadjust_backto_standard_pos_O.simps[simp del] wadjust_stop.simps[simp del] fun wadjust_inv :: "nat \ nat \ nat \ tape \ bool" where "wadjust_inv st m rs (l, r) = (if st = Suc 0 then wadjust_start m rs (l, r) else if st = Suc (Suc 0) then wadjust_loop_start m rs (l, r) else if st = Suc (Suc (Suc 0)) then wadjust_loop_right_move m rs (l, r) else if st = 4 then wadjust_loop_check m rs (l, r) else if st = 5 then wadjust_loop_erase m rs (l, r) else if st = 6 then wadjust_loop_on_left_moving m rs (l, r) else if st = 7 then wadjust_loop_right_move2 m rs (l, r) else if st = 8 then wadjust_erase2 m rs (l, r) else if st = 9 then wadjust_on_left_moving m rs (l, r) else if st = 10 then wadjust_goon_left_moving m rs (l, r) else if st = 11 then wadjust_backto_standard_pos m rs (l, r) else if st = 0 then wadjust_stop m rs (l, r) else False )" declare wadjust_inv.simps[simp del] fun wadjust_phase :: "nat \ config \ nat" where "wadjust_phase rs (st, l, r) = (if st = 1 then 3 else if st \ 2 \ st \ 7 then 2 else if st \ 8 \ st \ 11 then 1 else 0)" fun wadjust_stage :: "nat \ config \ nat" where "wadjust_stage rs (st, l, r) = (if st \ 2 \ st \ 7 then rs - length (takeWhile (\ a. a = Oc) (tl (dropWhile (\ a. a = Oc) (rev l @ r)))) else 0)" fun wadjust_state :: "nat \ config \ nat" where "wadjust_state rs (st, l, r) = (if st \ 2 \ st \ 7 then 8 - st else if st \ 8 \ st \ 11 then 12 - st else 0)" fun wadjust_step :: "nat \ config \ nat" where "wadjust_step rs (st, l, r) = (if st = 1 then (if hd r = Bk then 1 else 0) else if st = 3 then length r else if st = 5 then (if hd r = Oc then 1 else 0) else if st = 6 then length l else if st = 8 then (if hd r = Oc then 1 else 0) else if st = 9 then length l else if st = 10 then length l else if st = 11 then (if hd r = Bk then 0 else Suc (length l)) else 0)" fun wadjust_measure :: "(nat \ config) \ nat \ nat \ nat \ nat" where "wadjust_measure (rs, (st, l, r)) = (wadjust_phase rs (st, l, r), wadjust_stage rs (st, l, r), wadjust_state rs (st, l, r), wadjust_step rs (st, l, r))" definition wadjust_le :: "((nat \ config) \ nat \ config) set" where "wadjust_le \ (inv_image lex_square wadjust_measure)" lemma wf_lex_square[intro]: "wf lex_square" by(auto intro:wf_lex_prod simp: Abacus.lex_pair_def lex_square_def Abacus.lex_triple_def) lemma wf_wadjust_le[intro]: "wf wadjust_le" by(auto intro:wf_inv_image simp: wadjust_le_def Abacus.lex_triple_def Abacus.lex_pair_def) lemma wadjust_start_snd_nonempty[simp]: "wadjust_start m rs (c, []) = False" apply(auto simp: wadjust_start.simps) done lemma wadjust_loop_right_move_fst_nonempty[simp]: "wadjust_loop_right_move m rs (c, []) \ c \ []" apply(auto simp: wadjust_loop_right_move.simps) done lemma wadjust_loop_check_fst_nonempty[simp]: "wadjust_loop_check m rs (c, []) \ c \ []" apply(simp only: wadjust_loop_check.simps, auto) done lemma wadjust_loop_start_snd_nonempty[simp]: "wadjust_loop_start m rs (c, []) = False" apply(simp add: wadjust_loop_start.simps) done lemma wadjust_erase2_singleton[simp]: "wadjust_loop_check m rs (c, []) \ wadjust_erase2 m rs (tl c, [hd c])" apply(simp only: wadjust_loop_check.simps wadjust_erase2.simps, auto) done lemma wadjust_loop_on_left_moving_snd_nonempty[simp]: "wadjust_loop_on_left_moving m rs (c, []) = False" "wadjust_loop_right_move2 m rs (c, []) = False" "wadjust_erase2 m rs ([], []) = False" by(auto simp: wadjust_loop_on_left_moving.simps wadjust_loop_right_move2.simps wadjust_erase2.simps) lemma wadjust_on_left_moving_B_Bk1[simp]: "wadjust_on_left_moving_B m rs (Oc # Oc # Oc\(rs) @ Bk # Oc # Oc\(m), [Bk])" apply(simp add: wadjust_on_left_moving_B.simps, auto) done lemma wadjust_on_left_moving_B_Bk2[simp]: "wadjust_on_left_moving_B m rs (Bk\(n) @ Bk # Oc # Oc # Oc\(rs) @ Bk # Oc # Oc\(m), [Bk])" apply(simp add: wadjust_on_left_moving_B.simps , auto) apply(rule_tac x = "Suc n" in exI, simp add: exp_ind del: replicate_Suc) done lemma wadjust_on_left_moving_singleton[simp]: "\wadjust_erase2 m rs (c, []); c \ []\ \ wadjust_on_left_moving m rs (tl c, [hd c])" unfolding wadjust_erase2.simps apply(auto simp add: wadjust_on_left_moving.simps) apply (metis (no_types, lifting) empty_replicate hd_append hd_replicate list.sel(1) list.sel(3) self_append_conv2 tl_append2 tl_replicate wadjust_on_left_moving_B_Bk1 wadjust_on_left_moving_B_Bk2)+ done lemma wadjust_erase2_cases[simp]: "wadjust_erase2 m rs (c, []) \ (c = [] \ wadjust_on_left_moving m rs ([], [Bk])) \ (c \ [] \ wadjust_on_left_moving m rs (tl c, [hd c]))" apply(auto) done lemma wadjust_on_left_moving_nonempty[simp]: "wadjust_on_left_moving m rs ([], []) = False" "wadjust_on_left_moving_O m rs (c, []) = False" apply(auto simp: wadjust_on_left_moving.simps wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps) done lemma wadjust_on_left_moving_B_singleton_Bk[simp]: " \wadjust_on_left_moving_B m rs (c, []); c \ []; hd c = Bk\ \ wadjust_on_left_moving_B m rs (tl c, [Bk])" apply(auto simp add: wadjust_on_left_moving_B.simps hd_append) by (metis cell.distinct(1) empty_replicate list.sel(1) tl_append2 tl_replicate) lemma wadjust_on_left_moving_B_singleton_Oc[simp]: "\wadjust_on_left_moving_B m rs (c, []); c \ []; hd c = Oc\ \ wadjust_on_left_moving_O m rs (tl c, [Oc])" apply(auto simp add: wadjust_on_left_moving_B.simps wadjust_on_left_moving_O.simps hd_append) apply (metis cell.distinct(1) empty_replicate hd_replicate list.sel(3) self_append_conv2)+ done lemma wadjust_on_left_moving_singleton2[simp]: "\wadjust_on_left_moving m rs (c, []); c \ []\ \ wadjust_on_left_moving m rs (tl c, [hd c])" apply(simp add: wadjust_on_left_moving.simps) apply(case_tac "hd c", simp_all) done lemma wadjust_nonempty[simp]: "wadjust_goon_left_moving m rs (c, []) = False" "wadjust_backto_standard_pos m rs (c, []) = False" by(auto simp: wadjust_goon_left_moving.simps wadjust_goon_left_moving_B.simps wadjust_goon_left_moving_O.simps wadjust_backto_standard_pos.simps wadjust_backto_standard_pos_B.simps wadjust_backto_standard_pos_O.simps) lemma wadjust_loop_start_no_Bk[simp]: "wadjust_loop_start m rs (c, Bk # list) = False" apply(auto simp: wadjust_loop_start.simps) done lemma wadjust_loop_check_nonempty[simp]: "wadjust_loop_check m rs (c, b) \ c \ []" apply(simp only: wadjust_loop_check.simps, auto) done lemma wadjust_erase2_via_loop_check_Bk[simp]: "wadjust_loop_check m rs (c, Bk # list) \ wadjust_erase2 m rs (tl c, hd c # Bk # list)" by (auto simp: wadjust_loop_check.simps wadjust_erase2.simps) declare wadjust_loop_on_left_moving_O.simps[simp del] wadjust_loop_on_left_moving_B.simps[simp del] lemma wadjust_loop_on_left_moving_B_via_erase[simp]: "\wadjust_loop_erase m rs (c, Bk # list); hd c = Bk\ \ wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)" unfolding wadjust_loop_erase.simps wadjust_loop_on_left_moving_B.simps apply(erule_tac exE)+ apply(rename_tac ml mr ln rn) apply(rule_tac x = ml in exI, rule_tac x = mr in exI, rule_tac x = ln in exI, rule_tac x = 0 in exI) apply(case_tac ln, auto) apply(simp add: exp_ind [THEN sym]) done lemma wadjust_loop_on_left_moving_O_Bk_via_erase[simp]: "\wadjust_loop_erase m rs (c, Bk # list); c \ []; hd c = Oc\ \ wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)" apply(auto simp: wadjust_loop_erase.simps wadjust_loop_on_left_moving_O.simps) by (metis cell.distinct(1) empty_replicate hd_append hd_replicate list.sel(1)) lemma wadjust_loop_on_left_moving_Bk_via_erase[simp]: "\wadjust_loop_erase m rs (c, Bk # list); c \ []\ \ wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)" apply(case_tac "hd c", simp_all add:wadjust_loop_on_left_moving.simps) done lemma wadjust_loop_on_left_moving_B_Bk_move[simp]: "\wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Bk\ \ wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)" apply(simp only: wadjust_loop_on_left_moving_B.simps) apply(erule_tac exE)+ by (metis (no_types, lifting) cell.distinct(1) list.sel(1) replicate_Suc_iff_anywhere self_append_conv2 tl_append2 tl_replicate) lemma wadjust_loop_on_left_moving_O_Oc_move[simp]: "\wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Oc\ \ wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)" apply(simp only: wadjust_loop_on_left_moving_O.simps wadjust_loop_on_left_moving_B.simps) by (metis cell.distinct(1) empty_replicate hd_append hd_replicate list.sel(3) self_append_conv2) lemma wadjust_loop_erase_nonempty[simp]: "wadjust_loop_erase m rs (c, b) \ c \ []" "wadjust_loop_on_left_moving m rs (c, b) \ c \ []" "wadjust_loop_right_move2 m rs (c, b) \ c \ []" "wadjust_erase2 m rs (c, Bk # list) \ c \ []" "wadjust_on_left_moving m rs (c,b) \ c \ []" "wadjust_on_left_moving_O m rs (c, Bk # list) = False" "wadjust_goon_left_moving m rs (c, b) \ c \ []" "wadjust_loop_on_left_moving_O m rs (c, Bk # list) = False" by(auto simp: wadjust_loop_erase.simps wadjust_loop_on_left_moving.simps wadjust_loop_on_left_moving_O.simps wadjust_loop_on_left_moving_B.simps wadjust_loop_right_move2.simps wadjust_erase2.simps wadjust_on_left_moving.simps wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps wadjust_goon_left_moving.simps wadjust_goon_left_moving_B.simps wadjust_goon_left_moving_O.simps) lemma wadjust_loop_on_left_moving_Bk_move[simp]: "wadjust_loop_on_left_moving m rs (c, Bk # list) \ wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)" apply(simp add: wadjust_loop_on_left_moving.simps) apply(case_tac "hd c", simp_all) done lemma wadjust_loop_start_Oc_via_Bk_move[simp]: "wadjust_loop_right_move2 m rs (c, Bk # list) \ wadjust_loop_start m rs (c, Oc # list)" apply(auto simp: wadjust_loop_right_move2.simps wadjust_loop_start.simps replicate_app_Cons_same) by (metis add_Suc replicate_Suc) lemma wadjust_on_left_moving_Bk_via_erase[simp]: "wadjust_erase2 m rs (c, Bk # list) \ wadjust_on_left_moving m rs (tl c, hd c # Bk # list)" apply(auto simp: wadjust_erase2.simps wadjust_on_left_moving.simps replicate_app_Cons_same wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps) apply (metis exp_ind replicate_append_same)+ done lemma wadjust_on_left_moving_B_Bk_drop_one: "\wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Bk\ \ wadjust_on_left_moving_B m rs (tl c, Bk # Bk # list)" apply(auto simp: wadjust_on_left_moving_B.simps) by (metis cell.distinct(1) hd_append list.sel(1) tl_append2 tl_replicate) lemma wadjust_on_left_moving_B_Bk_drop_Oc: "\wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Oc\ \ wadjust_on_left_moving_O m rs (tl c, Oc # Bk # list)" apply(auto simp: wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps) by (metis cell.distinct(1) empty_replicate hd_append hd_replicate list.sel(3) self_append_conv2) lemma wadjust_on_left_moving_B_drop[simp]: "wadjust_on_left_moving m rs (c, Bk # list) \ wadjust_on_left_moving m rs (tl c, hd c # Bk # list)" by(cases "hd c", auto simp:wadjust_on_left_moving.simps wadjust_on_left_moving_B_Bk_drop_one wadjust_on_left_moving_B_Bk_drop_Oc) lemma wadjust_goon_left_moving_O_no_Bk[simp]: "wadjust_goon_left_moving_O m rs (c, Bk # list) = False" by (auto simp add: wadjust_goon_left_moving_O.simps) lemma wadjust_backto_standard_pos_via_left_Bk[simp]: "wadjust_goon_left_moving m rs (c, Bk # list) \ wadjust_backto_standard_pos m rs (tl c, hd c # Bk # list)" by(case_tac "hd c", simp_all add: wadjust_backto_standard_pos.simps wadjust_goon_left_moving.simps wadjust_goon_left_moving_B.simps wadjust_backto_standard_pos_O.simps) lemma wadjust_loop_right_move_Oc[simp]: "wadjust_loop_start m rs (c, Oc # list) \ wadjust_loop_right_move m rs (Oc # c, list)" apply(auto simp add: wadjust_loop_start.simps wadjust_loop_right_move.simps simp del:split_head_repeat) apply(rename_tac ln rn ml mr) apply(rule_tac x = ml in exI, rule_tac x = mr in exI, rule_tac x = 0 in exI, simp) apply(rule_tac x = "Suc ln" in exI, simp add: exp_ind del: replicate_Suc) done lemma wadjust_loop_check_Oc[simp]: assumes "wadjust_loop_right_move m rs (c, Oc # list)" shows "wadjust_loop_check m rs (Oc # c, list)" proof - from assms obtain ml mr nl nr rn where "c = Bk \ nl @ Oc # Oc \ ml @ Bk # Oc \ m @ [Oc]" "Oc # list = Bk \ nr @ Oc \ mr @ Bk \ rn" "ml + mr = Suc (Suc rs)" "0 < mr" "0 < nl + nr" unfolding wadjust_loop_right_move.simps exp_ind wadjust_loop_check.simps by auto hence "\ln. Oc # c = Oc # Bk \ ln @ Bk # Oc # Oc \ ml @ Bk # Oc \ Suc m" "\rn. list = Oc \ (mr - 1) @ Bk \ rn" "ml + (mr - 1) = Suc rs" by(cases nl;cases nr;cases mr;force simp add: wadjust_loop_right_move.simps exp_ind wadjust_loop_check.simps replicate_append_same)+ thus ?thesis unfolding wadjust_loop_check.simps by auto qed lemma wadjust_loop_erase_move_Oc[simp]: "wadjust_loop_check m rs (c, Oc # list) \ wadjust_loop_erase m rs (tl c, hd c # Oc # list)" apply(simp only: wadjust_loop_check.simps wadjust_loop_erase.simps) apply(erule_tac exE)+ using Cons_replicate_eq by fastforce lemma wadjust_loop_on_move_no_Oc[simp]: "wadjust_loop_on_left_moving_B m rs (c, Oc # list) = False" "wadjust_loop_right_move2 m rs (c, Oc # list) = False" "wadjust_loop_on_left_moving m rs (c, Oc # list) \ wadjust_loop_right_move2 m rs (Oc # c, list)" "wadjust_on_left_moving_B m rs (c, Oc # list) = False" "wadjust_loop_erase m rs (c, Oc # list) \ wadjust_loop_erase m rs (c, Bk # list)" by(auto simp: wadjust_loop_on_left_moving_B.simps wadjust_loop_on_left_moving_O.simps wadjust_loop_right_move2.simps replicate_app_Cons_same wadjust_loop_on_left_moving.simps wadjust_on_left_moving_B.simps wadjust_loop_erase.simps) lemma wadjust_goon_left_moving_B_Bk_Oc: "\wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Bk\ \ wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)" apply(auto simp: wadjust_on_left_moving_O.simps wadjust_goon_left_moving_B.simps ) done lemma wadjust_goon_left_moving_O_Oc_Oc: "\wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Oc\ \ wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)" apply(auto simp: wadjust_on_left_moving_O.simps wadjust_goon_left_moving_O.simps ) apply(auto simp: numeral_2_eq_2) done lemma wadjust_goon_left_moving_Oc[simp]: "wadjust_on_left_moving m rs (c, Oc # list) \ wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)" by(cases "hd c"; force simp: wadjust_on_left_moving.simps wadjust_goon_left_moving.simps wadjust_goon_left_moving_B_Bk_Oc wadjust_goon_left_moving_O_Oc_Oc)+ lemma left_moving_Bk_Oc[simp]: "\wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Bk\ \ wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)" apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps hd_append dest!: gr0_implies_Suc) apply (metis cell.distinct(1) empty_replicate hd_replicate list.sel(3) self_append_conv2) by (metis add_cancel_right_left cell.distinct(1) hd_replicate replicate_Suc_iff_anywhere) lemma left_moving_Oc_Oc[simp]: "\wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Oc\ \ wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)" apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps) apply(rename_tac mlx mrx rnx) apply(rule_tac x = "mlx - 1" in exI, simp) apply(case_tac mlx, simp_all add: ) apply(rule_tac x = "Suc mrx" in exI, auto simp: ) done lemma wadjust_goon_left_moving_B_no_Oc[simp]: "wadjust_goon_left_moving_B m rs (c, Oc # list) = False" apply(auto simp: wadjust_goon_left_moving_B.simps) done lemma wadjust_goon_left_moving_Oc_move[simp]: "wadjust_goon_left_moving m rs (c, Oc # list) \ wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)" by(cases "hd c",auto simp: wadjust_goon_left_moving.simps) lemma wadjust_backto_standard_pos_B_no_Oc[simp]: "wadjust_backto_standard_pos_B m rs (c, Oc # list) = False" apply(simp add: wadjust_backto_standard_pos_B.simps) done lemma wadjust_backto_standard_pos_O_no_Bk[simp]: "wadjust_backto_standard_pos_O m rs (c, Bk # xs) = False" by(simp add: wadjust_backto_standard_pos_O.simps) lemma wadjust_backto_standard_pos_B_Bk_Oc[simp]: "wadjust_backto_standard_pos_O m rs ([], Oc # list) \ wadjust_backto_standard_pos_B m rs ([], Bk # Oc # list)" apply(auto simp: wadjust_backto_standard_pos_O.simps wadjust_backto_standard_pos_B.simps) done lemma wadjust_backto_standard_pos_B_Bk_Oc_via_O[simp]: "\wadjust_backto_standard_pos_O m rs (c, Oc # list); c \ []; hd c = Bk\ \ wadjust_backto_standard_pos_B m rs (tl c, Bk # Oc # list)" apply(simp add:wadjust_backto_standard_pos_O.simps wadjust_backto_standard_pos_B.simps, auto) done lemma wadjust_backto_standard_pos_B_Oc_Oc_via_O[simp]: "\wadjust_backto_standard_pos_O m rs (c, Oc # list); c \ []; hd c = Oc\ \ wadjust_backto_standard_pos_O m rs (tl c, Oc # Oc # list)" apply(simp add: wadjust_backto_standard_pos_O.simps, auto) by force lemma wadjust_backto_standard_pos_cases[simp]: "wadjust_backto_standard_pos m rs (c, Oc # list) \ (c = [] \ wadjust_backto_standard_pos m rs ([], Bk # Oc # list)) \ (c \ [] \ wadjust_backto_standard_pos m rs (tl c, hd c # Oc # list))" apply(auto simp: wadjust_backto_standard_pos.simps) apply(case_tac "hd c", simp_all) done lemma wadjust_loop_right_move_nonempty_snd[simp]: "wadjust_loop_right_move m rs (c, []) = False" proof - {fix nl ml mr rn nr have "(c = Bk \ nl @ Oc # Oc \ ml @ Bk # Oc \ Suc m \ [] = Bk \ nr @ Oc \ mr @ Bk \ rn \ ml + mr = Suc (Suc rs) \ 0 < mr \ 0 < nl + nr) = False" by auto } note t=this thus ?thesis unfolding wadjust_loop_right_move.simps t by blast qed lemma wadjust_loop_erase_nonempty_snd[simp]: "wadjust_loop_erase m rs (c, []) = False" apply(simp only: wadjust_loop_erase.simps, auto) done lemma wadjust_loop_erase_cases2[simp]: "\Suc (Suc rs) = a; wadjust_loop_erase m rs (c, Bk # list)\ \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) = a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list))))" apply(simp only: wadjust_loop_erase.simps) apply(rule_tac disjI2) apply(case_tac c, simp, simp) done lemma dropWhile_exp1: "dropWhile (\a. a = Oc) (Oc\(n) @ xs) = dropWhile (\a. a = Oc) xs" apply(induct n, simp_all add: ) done lemma takeWhile_exp1: "takeWhile (\a. a = Oc) (Oc\(n) @ xs) = Oc\(n) @ takeWhile (\a. a = Oc) xs" apply(induct n, simp_all add: ) done lemma wadjust_correctness_helper_1: assumes "Suc (Suc rs) = a" " wadjust_loop_right_move2 m rs (c, Bk # list)" shows "a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list)))) < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list))))" proof - have "ml + mr = Suc rs \ 0 < mr \ rs - (ml + length (takeWhile (\a. a = Oc) list)) < Suc rs - (ml + length (takeWhile (\a. a = Oc) (Bk \ ln @ Bk # Bk # Oc \ mr @ Bk \ rn))) " for ml mr ln rn by(cases ln, auto) thus ?thesis using assms by (auto simp: wadjust_loop_right_move2.simps dropWhile_exp1 takeWhile_exp1) qed lemma wadjust_correctness_helper_2: "\Suc (Suc rs) = a; wadjust_loop_on_left_moving m rs (c, Bk # list)\ \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) = a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list))))" apply(subgoal_tac "c \ []") apply(case_tac c, simp_all) done lemma wadjust_loop_check_empty_false[simp]: "wadjust_loop_check m rs ([], b) = False" apply(simp add: wadjust_loop_check.simps) done lemma wadjust_loop_check_cases: "\Suc (Suc rs) = a; wadjust_loop_check m rs (c, Oc # list)\ \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Oc # list)))) < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list)))) \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Oc # list)))) = a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list))))" apply(case_tac "c", simp_all) done lemma wadjust_loop_erase_cases_or: "\Suc (Suc rs) = a; wadjust_loop_erase m rs (c, Oc # list)\ \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list)))) \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) = a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list))))" apply(simp add: wadjust_loop_erase.simps) apply(rule_tac disjI2) apply(auto) apply(simp add: dropWhile_exp1 takeWhile_exp1) done lemmas wadjust_correctness_helpers = wadjust_correctness_helper_2 wadjust_correctness_helper_1 wadjust_loop_erase_cases_or wadjust_loop_check_cases declare numeral_2_eq_2[simp del] lemma wadjust_start_Oc[simp]: "wadjust_start m rs (c, Bk # list) \ wadjust_start m rs (c, Oc # list)" apply(auto simp: wadjust_start.simps) done lemma wadjust_stop_Bk[simp]: "wadjust_backto_standard_pos m rs (c, Bk # list) \ wadjust_stop m rs (Bk # c, list)" apply(auto simp: wadjust_backto_standard_pos.simps wadjust_stop.simps wadjust_backto_standard_pos_B.simps) done lemma wadjust_loop_start_Oc[simp]: assumes "wadjust_start m rs (c, Oc # list)" shows "wadjust_loop_start m rs (Oc # c, list)" proof - from assms[unfolded wadjust_start.simps] obtain ln rn where "c = Bk # Oc # Oc \ m" "list = Oc # Bk \ ln @ Bk # Oc # Oc \ rs @ Bk \ rn" by(auto) hence "Oc # c = Oc \ 1 @ Bk # Oc \ Suc m \ list = Oc # Bk \ ln @ Bk # Oc \Suc rs @ Bk \ rn \ 1 + (Suc rs) = Suc (Suc rs) \ 0 < Suc rs" by auto thus ?thesis unfolding wadjust_loop_start.simps by blast qed lemma erase2_Bk_if_Oc[simp]:" wadjust_erase2 m rs (c, Oc # list) \ wadjust_erase2 m rs (c, Bk # list)" apply(auto simp: wadjust_erase2.simps) done lemma wadjust_loop_right_move_Bk[simp]: "wadjust_loop_right_move m rs (c, Bk # list) \ wadjust_loop_right_move m rs (Bk # c, list)" apply(simp only: wadjust_loop_right_move.simps) apply(erule_tac exE)+ apply auto apply (metis cell.distinct(1) empty_replicate hd_append hd_replicate less_SucI list.sel(1) list.sel(3) neq0_conv replicate_Suc_iff_anywhere tl_append2 tl_replicate)+ done lemma wadjust_correctness: shows "let P = (\ (len, st, l, r). st = 0) in let Q = (\ (len, st, l, r). wadjust_inv st m rs (l, r)) in let f = (\ stp. (Suc (Suc rs), steps0 (Suc 0, Bk # Oc\(Suc m), Bk # Oc # Bk\(ln) @ Bk # Oc\(Suc rs) @ Bk\(rn)) t_wcode_adjust stp)) in \ n .P (f n) \ Q (f n)" proof - let ?P = "(\ (len, st, l, r). st = 0)" let ?Q = "\ (len, st, l, r). wadjust_inv st m rs (l, r)" let ?f = "\ stp. (Suc (Suc rs), steps0 (Suc 0, Bk # Oc\(Suc m), Bk # Oc # Bk\(ln) @ Bk # Oc\(Suc rs) @ Bk\(rn)) t_wcode_adjust stp)" have "\ n. ?P (?f n) \ ?Q (?f n)" proof(rule_tac halt_lemma2) show "wf wadjust_le" by auto next { fix n assume a:"\ ?P (?f n) \ ?Q (?f n)" have "?Q (?f (Suc n)) \ (?f (Suc n), ?f n) \ wadjust_le" proof(cases "?f n") case (fields a b c d) then show ?thesis proof(cases d) case Nil then show ?thesis using a fields apply(simp add: step.simps) apply(simp_all only: wadjust_inv.simps split: if_splits) apply(simp_all add: wadjust_inv.simps wadjust_le_def wadjust_correctness_helpers Abacus.lex_triple_def Abacus.lex_pair_def lex_square_def split: if_splits). next case (Cons aa list) then show ?thesis using a fields Nil Cons apply((case_tac aa); simp add: step.simps) apply(simp_all only: wadjust_inv.simps split: if_splits) apply(simp_all) apply(simp_all add: wadjust_inv.simps wadjust_le_def wadjust_correctness_helpers Abacus.lex_triple_def Abacus.lex_pair_def lex_square_def split: if_splits). qed qed } thus "\ n. \ ?P (?f n) \ ?Q (?f n) \ ?Q (?f (Suc n)) \ (?f (Suc n), ?f n) \ wadjust_le" by auto next show "?Q (?f 0)" by(auto simp add: steps.simps wadjust_inv.simps wadjust_start.simps) next show "\ ?P (?f 0)" by (simp add: steps.simps) qed thus"?thesis" by simp qed lemma tm_wf_t_wcode_adjust[intro]: "tm_wf (t_wcode_adjust, 0)" by(auto simp: t_wcode_adjust_def tm_wf.simps) lemma bl_bin_nonzero[simp]: "args \ [] \ bl_bin () > 0" by(cases args) (auto simp: tape_of_nl_cons bl_bin.simps) lemma wcode_lemma_pre': "args \ [] \ \ stp rn. steps0 (Suc 0, [], ) ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp = (0, [Bk], Oc\(Suc m) @ Bk # Oc\(Suc (bl_bin ())) @ Bk\(rn))" proof - let ?P1 = "\ (l, r). l = [] \ r = " let ?Q1 = "\(l, r). l = Bk # Oc\(Suc m) \ (\ln rn. r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin ()) @ Bk\(rn))" let ?P2 = ?Q1 let ?Q2 = "\ (l, r). (wadjust_stop m (bl_bin () - 1) (l, r))" let ?P3 = "\ tp. False" assume h: "args \ []" hence a: "bl_bin () > 0" using h by simp hence "{?P1} (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust {?Q2}" proof(rule_tac Hoare_plus_halt) next show "tm_wf (t_wcode_prepare |+| t_wcode_main, 0)" by(rule_tac tm_comp_wf, auto) next show "{?P1} t_wcode_prepare |+| t_wcode_main {?Q1}" proof(rule_tac Hoare_haltI, auto) show "\n. is_final (steps0 (Suc 0, [], ) (t_wcode_prepare |+| t_wcode_main) n) \ (\(l, r). l = Bk # Oc # Oc \ m \ (\ln rn. r = Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn)) holds_for steps0 (Suc 0, [], ) (t_wcode_prepare |+| t_wcode_main) n" using h prepare_mainpart_lemma[of args m] apply(auto) apply(rename_tac stp ln rn) apply(rule_tac x = stp in exI, simp) apply(rule_tac x = ln in exI, auto) done qed next show "{?P2} t_wcode_adjust {?Q2}" proof(rule_tac Hoare_haltI, auto del: replicate_Suc) fix ln rn obtain n a b where "steps0 (Suc 0, Bk # Oc \ m @ [Oc], Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ (bl_bin () - Suc 0) @ Oc # Bk \ rn) t_wcode_adjust n = (0, a, b)" "wadjust_inv 0 m (bl_bin () - Suc 0) (a, b)" using wadjust_correctness[of m "bl_bin () - 1" "Suc ln" rn,unfolded Let_def] by(simp del: replicate_Suc add: replicate_Suc[THEN sym] exp_ind, auto) thus "\n. is_final (steps0 (Suc 0, Bk # Oc # Oc \ m, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn) t_wcode_adjust n) \ wadjust_stop m (bl_bin () - Suc 0) holds_for steps0 (Suc 0, Bk # Oc # Oc \ m, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn) t_wcode_adjust n" apply(rule_tac x = n in exI) using a apply(case_tac "bl_bin ()", simp, simp del: replicate_Suc add: exp_ind wadjust_inv.simps) by (simp add: replicate_append_same) qed qed thus "?thesis" apply(simp add: Hoare_halt_def, auto) apply(rename_tac n) apply(case_tac "(steps0 (Suc 0, [], <(m::nat) # args>) ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) n)") apply(rule_tac x = n in exI, auto simp: wadjust_stop.simps) using a apply(case_tac "bl_bin ()", simp_all) done qed text \ The initialization TM \t_wcode\. \ definition t_wcode :: "instr list" where "t_wcode = (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust " text \ The correctness of \t_wcode\. \ lemma wcode_lemma_1: "args \ [] \ \ stp ln rn. steps0 (Suc 0, [], ) (t_wcode) stp = (0, [Bk], Oc\(Suc m) @ Bk # Oc\(Suc (bl_bin ())) @ Bk\(rn))" apply(simp add: wcode_lemma_pre' t_wcode_def del: replicate_Suc) done lemma wcode_lemma: "args \ [] \ \ stp ln rn. steps0 (Suc 0, [], ) (t_wcode) stp = (0, [Bk], <[m ,bl_bin ()]> @ Bk\(rn))" using wcode_lemma_1[of args m] apply(simp add: t_wcode_def tape_of_list_def tape_of_nat_def) done section \The universal TM\ text \ This section gives the explicit construction of {\em Universal Turing Machine}, defined as \UTM\ and proves its correctness. It is pretty easy by composing the partial results we have got so far. \ definition UTM :: "instr list" where "UTM = (let (aprog, rs_pos, a_md) = rec_ci rec_F in let abc_F = aprog [+] dummy_abc (Suc (Suc 0)) in (t_wcode |+| (tm_of abc_F @ shift (mopup (Suc (Suc 0))) (length (tm_of abc_F) div 2))))" definition F_aprog :: "abc_prog" where "F_aprog \ (let (aprog, rs_pos, a_md) = rec_ci rec_F in aprog [+] dummy_abc (Suc (Suc 0)))" definition F_tprog :: "instr list" where "F_tprog = tm_of (F_aprog)" definition t_utm :: "instr list" where "t_utm \ F_tprog @ shift (mopup (Suc (Suc 0))) (length F_tprog div 2)" definition UTM_pre :: "instr list" where "UTM_pre = t_wcode |+| t_utm" lemma tinres_step1: assumes "tinres l l'" "step (ss, l, r) (t, 0) = (sa, la, ra)" "step (ss, l', r) (t, 0) = (sb, lb, rb)" shows "tinres la lb \ ra = rb \ sa = sb" proof(cases "r") case Nil then show ?thesis using assms by (cases "(fetch t ss Bk)";cases "fst (fetch t ss Bk)";auto simp:step.simps split:if_splits) next case (Cons a list) then show ?thesis using assms by (cases "(fetch t ss a)";cases "fst (fetch t ss a)";auto simp:step.simps split:if_splits) qed lemma tinres_steps1: "\tinres l l'; steps (ss, l, r) (t, 0) stp = (sa, la, ra); steps (ss, l', r) (t, 0) stp = (sb, lb, rb)\ \ tinres la lb \ ra = rb \ sa = sb" proof (induct stp arbitrary: sa la ra sb lb rb) case (Suc stp) then show ?case apply simp apply(case_tac "(steps (ss, l, r) (t, 0) stp)") apply(case_tac "(steps (ss, l', r) (t, 0) stp)") proof - fix stp sa la ra sb lb rb a b c aa ba ca assume ind: "\sa la ra sb lb rb. \steps (ss, l, r) (t, 0) stp = (sa, (la::cell list), ra); steps (ss, l', r) (t, 0) stp = (sb, lb, rb)\ \ tinres la lb \ ra = rb \ sa = sb" and h: " tinres l l'" "step (steps (ss, l, r) (t, 0) stp) (t, 0) = (sa, la, ra)" "step (steps (ss, l', r) (t, 0) stp) (t, 0) = (sb, lb, rb)" "steps (ss, l, r) (t, 0) stp = (a, b, c)" "steps (ss, l', r) (t, 0) stp = (aa, ba, ca)" have "tinres b ba \ c = ca \ a = aa" using ind h by metis thus "tinres la lb \ ra = rb \ sa = sb" using tinres_step1 h by metis qed qed (simp add: steps.simps) lemma tinres_some_exp[simp]: "tinres (Bk \ m @ [Bk, Bk]) la \ \m. la = Bk \ m" unfolding tinres_def proof - let ?c1 = "\ n. Bk \ m @ [Bk, Bk] = la @ Bk \ n" let ?c2 = "\ n. la = (Bk \ m @ [Bk, Bk]) @ Bk \ n" assume "\n. ?c1 n \ ?c2 n" then obtain n where "?c1 n \ ?c2 n" by auto then consider "?c1 n" | "?c2 n" by blast thus ?thesis proof(cases) case 1 hence "Bk \ Suc (Suc m) = la @ Bk \ n" by (metis exp_ind append_Cons append_eq_append_conv2 self_append_conv2) hence "la = Bk \ (Suc (Suc m) - n)" by (metis replicate_add append_eq_append_conv diff_add_inverse2 length_append length_replicate) then show ?thesis by auto next case 2 hence "la = Bk \ (m + Suc (Suc n))" by (metis append_Cons append_eq_append_conv2 replicate_Suc replicate_add self_append_conv2) then show ?thesis by blast qed qed lemma t_utm_halt_eq: assumes tm_wf: "tm_wf (tp, 0)" and exec: "steps0 (Suc 0, Bk\(l), ) tp stp = (0, Bk\(m), Oc\(rs)@Bk\(n))" and resutl: "0 < rs" shows "\stp m n. steps0 (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\(i)) t_utm stp = (0, Bk\(m), Oc\(rs) @ Bk\(n))" proof - obtain ap arity fp where a: "rec_ci rec_F = (ap, arity, fp)" by (metis prod_cases3) 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 (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 "terminate rec_F [code tp, bl2wc ()]" using assms by(rule_tac terminate_F, simp_all) next 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 (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 (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 moreover obtain sa la ra where "steps0 (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\i) (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp = (sa, la, ra)" apply(case_tac "steps0 (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\i) (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp", auto) done ultimately show "?thesis" using b apply(drule_tac la = "Bk\m @ [Bk, Bk]" in tinres_steps1, auto simp: numeral_2_eq_2) done qed thus "?thesis" apply(auto) apply(rule_tac x = stp in exI, simp add: t_utm_def) using assms apply(case_tac rs, simp_all add: numeral_2_eq_2) done qed lemma tm_wf_t_wcode[intro]: "tm_wf (t_wcode, 0)" apply(simp add: t_wcode_def) apply(rule_tac tm_comp_wf) apply(rule_tac tm_comp_wf, auto) done lemma UTM_halt_lemma_pre: assumes wf_tm: "tm_wf (tp, 0)" and result: "0 < rs" and args: "args \ []" and exec: "steps0 (Suc 0, Bk\(i), ) tp stp = (0, Bk\(m), Oc\(rs)@Bk\(k))" shows "\stp m n. steps0 (Suc 0, [], ) UTM_pre stp = (0, Bk\(m), Oc\(rs) @ Bk\(n))" proof - let ?Q2 = "\ (l, r). (\ ln rn. l = Bk\(ln) \ r = Oc\(rs) @ Bk\(rn))" let ?P1 = "\ (l, r). l = [] \ r = " let ?Q1 = "\ (l, r). (l = [Bk] \ (\ rn. r = Oc\(Suc (code tp)) @ Bk # Oc\(Suc (bl_bin ())) @ Bk\(rn)))" let ?P2 = ?Q1 let ?P3 = "\ (l, r). False" have "{?P1} (t_wcode |+| t_utm) {?Q2}" proof(rule_tac Hoare_plus_halt) show "tm_wf (t_wcode, 0)" by auto next show "{?P1} t_wcode {?Q1}" apply(rule_tac Hoare_haltI, auto) using wcode_lemma_1[of args "code tp"] args apply(auto) by (metis (mono_tags, lifting) holds_for.simps is_finalI old.prod.case) next show "{?P2} t_utm {?Q2}" proof(rule_tac Hoare_haltI, auto) fix rn show "\n. is_final (steps0 (Suc 0, [Bk], Oc # Oc \ code tp @ Bk # Oc # Oc \ bl_bin () @ Bk \ rn) t_utm n) \ (\(l, r). (\ln. l = Bk \ ln) \ (\rn. r = Oc \ rs @ Bk \ rn)) holds_for steps0 (Suc 0, [Bk], Oc # Oc \ code tp @ Bk # Oc # Oc \ bl_bin () @ Bk \ rn) t_utm n" using t_utm_halt_eq[of tp i "args" stp m rs k rn] assms apply(auto simp: bin_wc_eq tape_of_list_def tape_of_nat_def) apply(rename_tac stpa) apply(rule_tac x = stpa in exI, simp) done qed qed thus "?thesis" apply(auto simp: Hoare_halt_def UTM_pre_def) apply(case_tac "steps0 (Suc 0, [], ) (t_wcode |+| t_utm) n",simp) by auto qed text \ The correctness of \UTM\, the halt case. \ lemma UTM_halt_lemma': assumes tm_wf: "tm_wf (tp, 0)" and result: "0 < rs" and args: "args \ []" and exec: "steps0 (Suc 0, Bk\(i), ) tp stp = (0, Bk\(m), Oc\(rs)@Bk\(k))" shows "\stp m n. steps0 (Suc 0, [], ) UTM stp = (0, Bk\(m), Oc\(rs) @ Bk\(n))" using UTM_halt_lemma_pre[of tp rs args i stp m k] assms apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def) apply(case_tac "rec_ci rec_F", simp) done definition TSTD:: "config \ bool" where "TSTD c = (let (st, l, r) = c in st = 0 \ (\ m. l = Bk\(m)) \ (\ rs n. r = Oc\(Suc rs) @ Bk\(n)))" lemma nstd_case1: "0 < a \ NSTD (trpl_code (a, b, c))" by(simp add: NSTD.simps trpl_code.simps) lemma nonzero_bl2wc[simp]: "\m. b \ Bk\(m) \ 0 < bl2wc b" proof - have "\m. b \ Bk \ m \ bl2wc b = 0 \ False" proof(induct b) case (Cons a b) then show ?case apply(simp add: bl2wc.simps, case_tac a, simp_all add: bl2nat.simps bl2nat_double) apply(case_tac "\ m. b = Bk\(m)", erule exE) apply(metis append_Nil2 replicate_Suc_iff_anywhere) by simp qed auto thus "\m. b \ Bk\(m) \ 0 < bl2wc b" by auto qed lemma nstd_case2: "\m. b \ Bk\(m) \ NSTD (trpl_code (a, b, c))" apply(simp add: NSTD.simps trpl_code.simps) done lemma even_not_odd[elim]: "Suc (2 * x) = 2 * y \ RR" proof(induct x arbitrary: y) case (Suc x) thus ?case by(cases y;auto) qed auto declare replicate_Suc[simp del] lemma bl2nat_zero_eq[simp]: "(bl2nat c 0 = 0) = (\n. c = Bk\(n))" proof(induct c) case (Cons a c) then show ?case by (cases a;auto simp: bl2nat.simps bl2nat_double Cons_replicate_eq) qed (auto simp: bl2nat.simps) lemma bl2wc_exp_ex: "\Suc (bl2wc c) = 2 ^ m\ \ \ rs n. c = Oc\(rs) @ Bk\(n)" proof(induct c arbitrary: m) case (Cons a c m) { fix n have "Bk # Bk \ n = Oc \ 0 @ Bk \ Suc n" by (auto simp:replicate_Suc) hence "\rs na. Bk # Bk \ n = Oc \ rs @ Bk \ na" by blast } with Cons show ?case apply(cases a, auto) apply(case_tac m, simp_all add: bl2wc.simps, auto) apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double Cons) apply(case_tac m, simp, simp add: bin_wc_eq bl2wc.simps twice_power ) by (metis Cons.hyps Suc_pred bl2wc.simps neq0_conv power_not_zero replicate_Suc_iff_anywhere zero_neq_numeral) qed (simp add: bl2wc.simps bl2nat.simps) lemma lg_bin: assumes "\rs n. c \ Oc\(Suc rs) @ Bk\(n)" "bl2wc c = 2 ^ lg (Suc (bl2wc c)) 2 - Suc 0" shows "bl2wc c = 0" proof - from assms obtain rs nat n where *:"2 ^ rs - Suc 0 = nat" "c = Oc \ rs @ Bk \ n" using bl2wc_exp_ex[of c "lg (Suc (bl2wc c)) 2"] by(case_tac "(2::nat) ^ lg (Suc (bl2wc c)) 2", simp, simp, erule_tac exE, erule_tac exE, simp) have r:"bl2wc (Oc \ rs) = nat" by (metis "*"(1) bl2nat_exp_zero bl2wc.elims) hence "Suc (bl2wc c) = 2^rs" using * by(case_tac "(2::nat)^rs", auto) thus ?thesis using * assms(1) apply(drule_tac bl2wc_exp_ex, simp, erule_tac exE, erule_tac exE) by(case_tac rs, simp, simp) qed lemma nstd_case3: "\rs n. c \ Oc\(Suc rs) @ Bk\(n) \ NSTD (trpl_code (a, b, c))" apply(simp add: NSTD.simps trpl_code.simps) apply(auto) apply(drule_tac lg_bin, simp_all) done lemma NSTD_1: "\ TSTD (a, b, c) \ rec_exec rec_NSTD [trpl_code (a, b, c)] = Suc 0" using NSTD_lemma1[of "trpl_code (a, b, c)"] NSTD_lemma2[of "trpl_code (a, b, c)"] apply(simp add: TSTD_def) apply(erule_tac disjE, erule_tac nstd_case1) apply(erule_tac disjE, erule_tac nstd_case2) apply(erule_tac nstd_case3) done lemma nonstop_t_uhalt_eq: "\tm_wf (tp, 0); steps0 (Suc 0, Bk\(l), ) tp stp = (a, b, c); \ TSTD (a, b, c)\ \ rec_exec rec_nonstop [code tp, bl2wc (), stp] = Suc 0" apply(simp add: rec_nonstop_def rec_exec.simps) apply(subgoal_tac "rec_exec rec_conf [code tp, bl2wc (), stp] = trpl_code (a, b, c)", simp) apply(erule_tac NSTD_1) using rec_t_eq_steps[of tp l lm stp] apply(simp) done lemma nonstop_true: "\tm_wf (tp, 0); \ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))\ \ \y. rec_exec rec_nonstop ([code tp, bl2wc (), y]) = (Suc 0)" proof fix y assume a:"tm_wf0 tp" "\stp. \ TSTD (steps0 (Suc 0, Bk \ l, ) tp stp)" hence "\ TSTD (steps0 (Suc 0, Bk \ l, ) tp y)" by auto thus "rec_exec rec_nonstop [code tp, bl2wc (), y] = Suc 0" by (cases "steps0 (Suc 0, Bk\(l), ) tp y") (auto intro: nonstop_t_uhalt_eq[OF a(1)]) qed 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: 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)\ \ {\ 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 simp: abc_Hoare_unhalt_def, case_tac "abc_steps_l (0, [code tp, bl2wc ()]) ap n", simp) fix n a b assume h: "\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 a: "ap \ []" using h rec_ci_not_null[of "rec_F" pos md] by auto ultimately show "a < length 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 n ss nl] h by(simp) qed qed lemma uabc_uhalt: "\tm_wf (tp, 0); \ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))\ \ {\ nl. nl = [code tp, bl2wc ()]} F_aprog \ " proof - obtain a b c where abc:"rec_ci rec_F = (a,b,c)" by (cases "rec_ci rec_F") force assume a:"tm_wf (tp, 0)" "\ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))" from uabc_uhalt'[OF a abc] abc_Hoare_plus_unhalt1 show "{\ nl. nl = [code tp, bl2wc ()]} F_aprog \" by(simp add: F_aprog_def abc) qed 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, 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 ()]) (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) []" by(auto simp: crsp.simps start_of.simps) next fix stp a b show "abc_steps_l (0, [code tp, bl2wc ()]) F_aprog stp = (a, b) \ a < length F_aprog" using assms 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 lemma inres_tape: "\steps0 (st, l, r) tp stp = (a, b, c); steps0 (st, l', r') tp stp = (a', b', c'); tinres l l'; tinres r r'\ \ a = a' \ tinres b b' \ tinres c c'" proof(case_tac "steps0 (st, l', r) tp stp") fix aa ba ca assume h: "steps0 (st, l, r) tp stp = (a, b, c)" "steps0 (st, l', r') tp stp = (a', b', c')" "tinres l l'" "tinres r r'" "steps0 (st, l', r) tp stp = (aa, ba, ca)" have "tinres b ba \ c = ca \ a = aa" using h apply(rule_tac tinres_steps1, auto) done moreover have "b' = ba \ tinres c' ca \ a' = aa" using h apply(rule_tac tinres_steps2, auto intro: tinres_commute) done ultimately show "?thesis" apply(auto intro: tinres_commute) done qed lemma tape_normalize: assumes "\ stp. \ is_final(steps0 (Suc 0, [Bk,Bk], <[code tp, bl2wc ()]>) t_utm stp)" shows "\ stp. \ is_final (steps0 (Suc 0, Bk\(m), <[code tp, bl2wc ()]> @ Bk\(n)) t_utm stp)" (is "\ stp. ?P stp") proof fix stp from assms[rule_format,of stp] show "?P stp" apply(case_tac "steps0 (Suc 0, Bk\(m), <[code tp, bl2wc ()]> @ Bk\(n)) t_utm stp", simp) apply(case_tac "steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) t_utm stp", simp) apply(drule_tac inres_tape, auto) apply(auto simp: tinres_def) apply(case_tac "m > Suc (Suc 0)") apply(rule_tac x = "m - Suc (Suc 0)" in exI) apply(case_tac m, simp_all) apply(metis Suc_lessD Suc_pred replicate_Suc) 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 qed lemma tutm_uhalt: "\tm_wf (tp,0); \ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))\ \ \ stp. \ is_final (steps0 (Suc 0, Bk\(m), <[code tp, bl2wc ()]> @ Bk\(n)) t_utm stp)" apply(rule_tac tape_normalize) apply(rule_tac tutm_uhalt'[simplified], simp_all) done lemma UTM_uhalt_lemma_pre: assumes tm_wf: "tm_wf (tp, 0)" and exec: "\ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))" and args: "args \ []" shows "\ stp. \ is_final (steps0 (Suc 0, [], ) UTM_pre stp)" proof - let ?P1 = "\ (l, r). l = [] \ r = " let ?Q1 = "\ (l, r). (l = [Bk] \ (\ rn. r = Oc\(Suc (code tp)) @ Bk # Oc\(Suc (bl_bin ())) @ Bk\(rn)))" let ?P2 = ?Q1 have "{?P1} (t_wcode |+| t_utm) \" proof(rule_tac Hoare_plus_unhalt) show "tm_wf (t_wcode, 0)" by auto next show "{?P1} t_wcode {?Q1}" apply(rule_tac Hoare_haltI, auto) using wcode_lemma_1[of args "code tp"] args apply(auto) by (metis (mono_tags, lifting) holds_for.simps is_finalI old.prod.case) next show "{?P2} t_utm \" proof(rule_tac Hoare_unhaltI, auto) fix n rn assume h: "is_final (steps0 (Suc 0, [Bk], Oc \ Suc (code tp) @ Bk # Oc \ Suc (bl_bin ()) @ Bk \ rn) t_utm n)" have "\ stp. \ is_final (steps0 (Suc 0, Bk\(Suc 0), <[code tp, bl2wc ()]> @ Bk\(rn)) t_utm stp)" using assms apply(rule_tac tutm_uhalt, simp_all) done thus "False" using h apply(erule_tac x = n in allE) apply(simp add: tape_of_list_def bin_wc_eq tape_of_nat_def) done qed qed thus "?thesis" apply(simp add: Hoare_unhalt_def UTM_pre_def) done qed text \ The correctness of \UTM\, the unhalt case. \ lemma UTM_uhalt_lemma': assumes tm_wf: "tm_wf (tp, 0)" and unhalt: "\ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))" and args: "args \ []" shows " \ stp. \ is_final (steps0 (Suc 0, [], ) UTM stp)" using UTM_uhalt_lemma_pre[of tp l args] assms apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def) apply(case_tac "rec_ci rec_F", simp) done lemma UTM_halt_lemma: assumes tm_wf: "tm_wf (p, 0)" and resut: "rs > 0" and args: "(args::nat list) \ []" and exec: "{(\tp. tp = (Bk\i, ))} p {(\tp. tp = (Bk\m, Oc\rs @ Bk\k))}" shows "{(\tp. tp = ([], ))} UTM {(\tp. (\ m n. tp = (Bk\m, Oc\rs @ Bk\n)))}" proof - let ?steps0 = "steps0 (Suc 0, [], )" let ?stepsBk = "steps0 (Suc 0, Bk\i, ) p" from wcode_lemma_1[OF args,of "code p"] obtain stp ln rn where wcl1:"?steps0 t_wcode stp = (0, [Bk], Oc \ Suc (code p) @ Bk # Oc \ Suc (bl_bin ()) @ Bk \ rn)" by fast from exec Hoare_halt_def obtain n where n:"{\tp. tp = (Bk \ i, )} p {\tp. tp = (Bk \ m, Oc \ rs @ Bk \ k)}" "is_final (?stepsBk n)" "(\tp. tp = (Bk \ m, Oc \ rs @ Bk \ k)) holds_for steps0 (Suc 0, Bk \ i, ) p n" by auto obtain a where a:"a = fst (rec_ci rec_F)" by blast have "{(\ (l, r). l = [] \ r = )} (t_wcode |+| t_utm) {(\ (l, r). (\ m. l = Bk\m) \ (\ n. r = Oc\rs @ Bk\n))}" proof(rule_tac Hoare_plus_halt) show "{\(l, r). l = [] \ r = } t_wcode {\ (l, r). (l = [Bk] \ (\ rn. r = Oc\(Suc (code p)) @ Bk # Oc\(Suc (bl_bin ())) @ Bk\(rn)))}" using wcl1 by (auto intro!:Hoare_haltI exI[of _ stp]) next have "\ stp. (?stepsBk stp = (0, Bk\m, Oc\rs @ Bk\k))" using n by (case_tac "?stepsBk n", auto) then obtain stp where k: "steps0 (Suc 0, Bk\i, ) p stp = (0, Bk\m, Oc\rs @ Bk\k)" .. thus "{\(l, r). l = [Bk] \ (\rn. r = Oc \ Suc (code p) @ Bk # Oc \ Suc (bl_bin ()) @ Bk \ rn)} t_utm {\(l, r). (\m. l = Bk \ m) \ (\n. r = Oc \ rs @ Bk \ n)}" proof(rule_tac Hoare_haltI, auto) fix rn from t_utm_halt_eq[OF assms(1) k assms(2),of rn] assms k have "\ ma n stp. steps0 (Suc 0, [Bk], <[code p, bl2wc ()]> @ Bk \ rn) t_utm stp = (0, Bk \ ma, Oc \ rs @ Bk \ n)" by (auto simp add: bin_wc_eq) then obtain stpx m' n' where t:"steps0 (Suc 0, [Bk], <[code p, bl2wc ()]> @ Bk \ rn) t_utm stpx = (0, Bk \ m', Oc \ rs @ Bk \ n')" by auto show "\n. is_final (steps0 (Suc 0, [Bk], Oc \ Suc (code p) @ Bk # Oc \ Suc (bl_bin ()) @ Bk \ rn) t_utm n) \ (\(l, r). (\m. l = Bk \ m) \ (\n. r = Oc \ rs @ Bk \ n)) holds_for steps0 (Suc 0, [Bk], Oc \ Suc (code p) @ Bk # Oc \ Suc (bl_bin ()) @ Bk \ rn) t_utm n" using t by(auto simp: bin_wc_eq tape_of_list_def tape_of_nat_def intro:exI[of _ stpx]) qed next show "tm_wf0 t_wcode" by auto qed then obtain n where "is_final (?steps0 (t_wcode |+| t_utm) n)" "(\(l, r). (\m. l = Bk \ m) \ (\n. r = Oc \ rs @ Bk \ n)) holds_for ?steps0 (t_wcode |+| t_utm) n" by(auto simp add: Hoare_halt_def a) thus "?thesis" apply(case_tac "rec_ci rec_F") apply(auto simp add: UTM_def Hoare_halt_def) apply(case_tac "(?steps0 (t_wcode |+| t_utm) n)") apply(rule_tac x="n" in exI) apply(auto simp add:a t_utm_def F_aprog_def F_tprog_def) done qed lemma UTM_halt_lemma2: assumes tm_wf: "tm_wf (p, 0)" and args: "(args::nat list) \ []" and exec: "{(\tp. tp = ([], ))} p {(\tp. tp = (Bk\m, <(n::nat)> @ Bk\k))}" shows "{(\tp. tp = ([], ))} UTM {(\tp. (\ m k. tp = (Bk\m, @ Bk\k)))}" using UTM_halt_lemma[OF assms(1) _ assms(2), where i="0"] using assms(3) by(simp add: tape_of_nat_def) lemma UTM_unhalt_lemma: assumes tm_wf: "tm_wf (p, 0)" and unhalt: "{(\tp. tp = (Bk\i, ))} p \" and args: "args \ []" shows "{(\tp. tp = ([], ))} UTM \" proof - have "(\ TSTD (steps0 (Suc 0, Bk\(i), ) p stp))" for stp (* in unhalt, we substitute inner 'forall' n\stp *) using unhalt[unfolded Hoare_unhalt_def,rule_format,OF refl,of stp] by(cases "steps0 (Suc 0, Bk \ i, ) p stp",auto simp: Hoare_unhalt_def TSTD_def) then have "\ stp. \ is_final (steps0 (Suc 0, [], ) UTM stp)" using assms by(intro UTM_uhalt_lemma', auto) thus "?thesis" by(simp add: Hoare_unhalt_def) qed lemma UTM_unhalt_lemma2: assumes tm_wf: "tm_wf (p, 0)" and unhalt: "{(\tp. tp = ([], ))} p \" and args: "args \ []" shows "{(\tp. tp = ([], ))} UTM \" using UTM_unhalt_lemma[OF assms(1), where i="0"] using assms(2-3) by(simp add: tape_of_nat_def) end \ No newline at end of file