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,4960 +1,4958 @@ (* 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 (metis dist_at2 dist_f12 last_rev last_snoc list.inject vertices_face.simps) 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" supply [[simproc del: defined_all]] 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" supply [[simproc del: defined_all]] 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" supply [[simproc del: defined_all]] 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" supply [[simproc del: defined_all]] 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" supply [[simproc del: defined_all]] 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" supply [[simproc del: defined_all]] 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" supply [[simproc del: defined_all]] 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" supply [[simproc del: defined_all]] 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 + by (auto simp: less_Suc_eq_0_disj) 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" 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" by (meson pre_subdivFace'_Some1' pre_subdivFace'_preFaceDiv) end diff --git a/thys/Universal_Turing_Machine/UF.thy b/thys/Universal_Turing_Machine/UF.thy --- a/thys/Universal_Turing_Machine/UF.thy +++ b/thys/Universal_Turing_Machine/UF.thy @@ -1,4206 +1,4202 @@ (* Title: thys/UF.thy Author: Jian Xu, Xingyuan Zhang, and Christian Urban Modifications: Sebastiaan Joosten *) chapter \Construction of a Universal Function\ theory UF imports Rec_Def HOL.GCD Abacus begin text \ This theory file constructs the Universal Function \rec_F\, which is the UTM defined in terms of recursive functions. This \rec_F\ is essentially an interpreter of Turing Machines. Once the correctness of \rec_F\ is established, UTM can easil be obtained by compling \rec_F\ into the corresponding Turing Machine. \ section \Universal Function\ subsection \The construction of component functions\ text \ The recursive function used to do arithmetic addition. \ definition rec_add :: "recf" where "rec_add \ Pr 1 (id 1 0) (Cn 3 s [id 3 2])" text \ The recursive function used to do arithmetic multiplication. \ definition rec_mult :: "recf" where "rec_mult = Pr 1 z (Cn 3 rec_add [id 3 0, id 3 2])" text \ The recursive function used to do arithmetic precede. \ definition rec_pred :: "recf" where "rec_pred = Cn 1 (Pr 1 z (id 3 1)) [id 1 0, id 1 0]" text \ The recursive function used to do arithmetic subtraction. \ definition rec_minus :: "recf" where "rec_minus = Pr 1 (id 1 0) (Cn 3 rec_pred [id 3 2])" text \ \constn n\ is the recursive function which computes nature number \n\. \ fun constn :: "nat \ recf" where "constn 0 = z" | "constn (Suc n) = Cn 1 s [constn n]" text \ Sign function, which returns 1 when the input argument is greater than \0\. \ definition rec_sg :: "recf" where "rec_sg = Cn 1 rec_minus [constn 1, Cn 1 rec_minus [constn 1, id 1 0]]" text \ \rec_less\ compares its two arguments, returns \1\ if the first is less than the second; otherwise returns \0\. \ definition rec_less :: "recf" where "rec_less = Cn 2 rec_sg [Cn 2 rec_minus [id 2 1, id 2 0]]" text \ \rec_not\ inverse its argument: returns \1\ when the argument is \0\; returns \0\ otherwise. \ definition rec_not :: "recf" where "rec_not = Cn 1 rec_minus [constn 1, id 1 0]" text \ \rec_eq\ compares its two arguments: returns \1\ if they are equal; return \0\ otherwise. \ definition rec_eq :: "recf" where "rec_eq = Cn 2 rec_minus [Cn 2 (constn 1) [id 2 0], Cn 2 rec_add [Cn 2 rec_minus [id 2 0, id 2 1], Cn 2 rec_minus [id 2 1, id 2 0]]]" text \ \rec_conj\ computes the conjunction of its two arguments, returns \1\ if both of them are non-zero; returns \0\ otherwise. \ definition rec_conj :: "recf" where "rec_conj = Cn 2 rec_sg [Cn 2 rec_mult [id 2 0, id 2 1]] " text \ \rec_disj\ computes the disjunction of its two arguments, returns \0\ if both of them are zero; returns \0\ otherwise. \ definition rec_disj :: "recf" where "rec_disj = Cn 2 rec_sg [Cn 2 rec_add [id 2 0, id 2 1]]" text \ Computes the arity of recursive function. \ fun arity :: "recf \ nat" where "arity z = 1" | "arity s = 1" | "arity (id m n) = m" | "arity (Cn n f gs) = n" | "arity (Pr n f g) = Suc n" | "arity (Mn n f) = n" text \ \get_fstn_args n (Suc k)\ returns \[id n 0, id n 1, id n 2, \, id n k]\, the effect of which is to take out the first \Suc k\ arguments out of the \n\ input arguments. \ fun get_fstn_args :: "nat \ nat \ recf list" where "get_fstn_args n 0 = []" | "get_fstn_args n (Suc y) = get_fstn_args n y @ [id n y]" text \ \rec_sigma f\ returns the recursive functions which sums up the results of \f\: \[ (rec\_sigma f)(x, y) = f(x, 0) + f(x, 1) + \cdots + f(x, y) \] \ fun rec_sigma :: "recf \ recf" where "rec_sigma rf = (let vl = arity rf in Pr (vl - 1) (Cn (vl - 1) rf (get_fstn_args (vl - 1) (vl - 1) @ [Cn (vl - 1) (constn 0) [id (vl - 1) 0]])) (Cn (Suc vl) rec_add [id (Suc vl) vl, Cn (Suc vl) rf (get_fstn_args (Suc vl) (vl - 1) @ [Cn (Suc vl) s [id (Suc vl) (vl - 1)]])]))" text \ \rec_exec\ is the interpreter function for reursive functions. The function is defined such that it always returns meaningful results for primitive recursive functions. \ declare rec_exec.simps[simp del] constn.simps[simp del] text \ Correctness of \rec_add\. \ lemma add_lemma: "\ x y. rec_exec rec_add [x, y] = x + y" by(induct_tac y, auto simp: rec_add_def rec_exec.simps) text \ Correctness of \rec_mult\. \ lemma mult_lemma: "\ x y. rec_exec rec_mult [x, y] = x * y" by(induct_tac y, auto simp: rec_mult_def rec_exec.simps add_lemma) text \ Correctness of \rec_pred\. \ lemma pred_lemma: "\ x. rec_exec rec_pred [x] = x - 1" by(induct_tac x, auto simp: rec_pred_def rec_exec.simps) text \ Correctness of \rec_minus\. \ lemma minus_lemma: "\ x y. rec_exec rec_minus [x, y] = x - y" by(induct_tac y, auto simp: rec_exec.simps rec_minus_def pred_lemma) text \ Correctness of \rec_sg\. \ lemma sg_lemma: "\ x. rec_exec rec_sg [x] = (if x = 0 then 0 else 1)" by(auto simp: rec_sg_def minus_lemma rec_exec.simps constn.simps) text \ Correctness of \constn\. \ lemma constn_lemma: "rec_exec (constn n) [x] = n" by(induct n, auto simp: rec_exec.simps constn.simps) text \ Correctness of \rec_less\. \ lemma less_lemma: "\ x y. rec_exec rec_less [x, y] = (if x < y then 1 else 0)" by(induct_tac y, auto simp: rec_exec.simps rec_less_def minus_lemma sg_lemma) text \ Correctness of \rec_not\. \ lemma not_lemma: "\ x. rec_exec rec_not [x] = (if x = 0 then 1 else 0)" by(induct_tac x, auto simp: rec_exec.simps rec_not_def constn_lemma minus_lemma) text \ Correctness of \rec_eq\. \ lemma eq_lemma: "\ x y. rec_exec rec_eq [x, y] = (if x = y then 1 else 0)" by(induct_tac y, auto simp: rec_exec.simps rec_eq_def constn_lemma add_lemma minus_lemma) text \ Correctness of \rec_conj\. \ lemma conj_lemma: "\ x y. rec_exec rec_conj [x, y] = (if x = 0 \ y = 0 then 0 else 1)" by(induct_tac y, auto simp: rec_exec.simps sg_lemma rec_conj_def mult_lemma) text \ Correctness of \rec_disj\. \ lemma disj_lemma: "\ x y. rec_exec rec_disj [x, y] = (if x = 0 \ y = 0 then 0 else 1)" by(induct_tac y, auto simp: rec_disj_def sg_lemma add_lemma rec_exec.simps) text \ \primrec recf n\ is true iff \recf\ is a primitive recursive function with arity \n\. \ inductive primerec :: "recf \ nat \ bool" where prime_z[intro]: "primerec z (Suc 0)" | prime_s[intro]: "primerec s (Suc 0)" | prime_id[intro!]: "\n < m\ \ primerec (id m n) m" | prime_cn[intro!]: "\primerec f k; length gs = k; \ i < length gs. primerec (gs ! i) m; m = n\ \ primerec (Cn n f gs) m" | prime_pr[intro!]: "\primerec f n; primerec g (Suc (Suc n)); m = Suc n\ \ primerec (Pr n f g) m" inductive_cases prime_cn_reverse'[elim]: "primerec (Cn n f gs) n" inductive_cases prime_mn_reverse: "primerec (Mn n f) m" inductive_cases prime_z_reverse[elim]: "primerec z n" inductive_cases prime_s_reverse[elim]: "primerec s n" inductive_cases prime_id_reverse[elim]: "primerec (id m n) k" inductive_cases prime_cn_reverse[elim]: "primerec (Cn n f gs) m" inductive_cases prime_pr_reverse[elim]: "primerec (Pr n f g) m" declare mult_lemma[simp] add_lemma[simp] pred_lemma[simp] minus_lemma[simp] sg_lemma[simp] constn_lemma[simp] less_lemma[simp] not_lemma[simp] eq_lemma[simp] conj_lemma[simp] disj_lemma[simp] text \ \Sigma\ is the logical specification of the recursive function \rec_sigma\. \ function Sigma :: "(nat list \ nat) \ nat list \ nat" where "Sigma g xs = (if last xs = 0 then g xs else (Sigma g (butlast xs @ [last xs - 1]) + g xs)) " by pat_completeness auto termination proof show "wf (measure (\ (f, xs). last xs))" by auto next fix g xs assume "last (xs::nat list) \ 0" thus "((g, butlast xs @ [last xs - 1]), g, xs) \ measure (\(f, xs). last xs)" by auto qed declare rec_exec.simps[simp del] get_fstn_args.simps[simp del] arity.simps[simp del] Sigma.simps[simp del] rec_sigma.simps[simp del] lemma rec_pr_Suc_simp_rewrite: "rec_exec (Pr n f g) (xs @ [Suc x]) = rec_exec g (xs @ [x] @ [rec_exec (Pr n f g) (xs @ [x])])" by(simp add: rec_exec.simps) lemma Sigma_0_simp_rewrite: "Sigma f (xs @ [0]) = f (xs @ [0])" by(simp add: Sigma.simps) lemma Sigma_Suc_simp_rewrite: "Sigma f (xs @ [Suc x]) = Sigma f (xs @ [x]) + f (xs @ [Suc x])" by(simp add: Sigma.simps) lemma append_access_1[simp]: "(xs @ ys) ! (Suc (length xs)) = ys ! 1" by(simp add: nth_append) lemma get_fstn_args_take: "\length xs = m; n \ m\ \ map (\ f. rec_exec f xs) (get_fstn_args m n)= take n xs" proof(induct n) case 0 thus "?case" by(simp add: get_fstn_args.simps) next case (Suc n) thus "?case" by(simp add: get_fstn_args.simps rec_exec.simps take_Suc_conv_app_nth) qed lemma arity_primerec[simp]: "primerec f n \ arity f = n" apply(cases f) apply(auto simp: arity.simps ) apply(erule_tac prime_mn_reverse) done lemma rec_sigma_Suc_simp_rewrite: "primerec f (Suc (length xs)) \ rec_exec (rec_sigma f) (xs @ [Suc x]) = rec_exec (rec_sigma f) (xs @ [x]) + rec_exec f (xs @ [Suc x])" apply(induct x) apply(auto simp: rec_sigma.simps Let_def rec_pr_Suc_simp_rewrite rec_exec.simps get_fstn_args_take) done text \ The correctness of \rec_sigma\ with respect to its specification. \ lemma sigma_lemma: "primerec rg (Suc (length xs)) \ rec_exec (rec_sigma rg) (xs @ [x]) = Sigma (rec_exec rg) (xs @ [x])" apply(induct x) apply(auto simp: rec_exec.simps rec_sigma.simps Let_def get_fstn_args_take Sigma_0_simp_rewrite Sigma_Suc_simp_rewrite) done text \ \rec_accum f (x1, x2, \, xn, k) = f(x1, x2, \, xn, 0) * f(x1, x2, \, xn, 1) * \ f(x1, x2, \, xn, k)\ \ fun rec_accum :: "recf \ recf" where "rec_accum rf = (let vl = arity rf in Pr (vl - 1) (Cn (vl - 1) rf (get_fstn_args (vl - 1) (vl - 1) @ [Cn (vl - 1) (constn 0) [id (vl - 1) 0]])) (Cn (Suc vl) rec_mult [id (Suc vl) (vl), Cn (Suc vl) rf (get_fstn_args (Suc vl) (vl - 1) @ [Cn (Suc vl) s [id (Suc vl) (vl - 1)]])]))" text \ \Accum\ is the formal specification of \rec_accum\. \ function Accum :: "(nat list \ nat) \ nat list \ nat" where "Accum f xs = (if last xs = 0 then f xs else (Accum f (butlast xs @ [last xs - 1]) * f xs))" by pat_completeness auto termination proof show "wf (measure (\ (f, xs). last xs))" by auto next fix f xs assume "last xs \ (0::nat)" thus "((f, butlast xs @ [last xs - 1]), f, xs) \ measure (\(f, xs). last xs)" by auto qed lemma rec_accum_Suc_simp_rewrite: "primerec f (Suc (length xs)) \ rec_exec (rec_accum f) (xs @ [Suc x]) = rec_exec (rec_accum f) (xs @ [x]) * rec_exec f (xs @ [Suc x])" apply(induct x) apply(auto simp: rec_sigma.simps Let_def rec_pr_Suc_simp_rewrite rec_exec.simps get_fstn_args_take) done text \ The correctness of \rec_accum\ with respect to its specification. \ lemma accum_lemma : "primerec rg (Suc (length xs)) \ rec_exec (rec_accum rg) (xs @ [x]) = Accum (rec_exec rg) (xs @ [x])" apply(induct x) apply(auto simp: rec_exec.simps rec_sigma.simps Let_def get_fstn_args_take) done declare rec_accum.simps [simp del] text \ \rec_all t f (x1, x2, \, xn)\ computes the charactrization function of the following FOL formula: \(\ x \ t(x1, x2, \, xn). (f(x1, x2, \, xn, x) > 0))\ \ fun rec_all :: "recf \ recf \ recf" where "rec_all rt rf = (let vl = arity rf in Cn (vl - 1) rec_sg [Cn (vl - 1) (rec_accum rf) (get_fstn_args (vl - 1) (vl - 1) @ [rt])])" lemma rec_accum_ex: assumes "primerec rf (Suc (length xs))" shows "(rec_exec (rec_accum rf) (xs @ [x]) = 0) = (\ t \ x. rec_exec rf (xs @ [t]) = 0)" proof(induct x) case (Suc x) with assms show ?case apply(auto simp add: rec_exec.simps rec_accum.simps get_fstn_args_take) apply(rename_tac t ta) apply(rule_tac x = ta in exI, simp) apply(case_tac "t = Suc x", simp_all) apply(rule_tac x = t in exI, simp) done qed (insert assms,auto simp add: rec_exec.simps rec_accum.simps get_fstn_args_take) text \ The correctness of \rec_all\. \ lemma all_lemma: "\primerec rf (Suc (length xs)); primerec rt (length xs)\ \ rec_exec (rec_all rt rf) xs = (if (\ x \ (rec_exec rt xs). 0 < rec_exec rf (xs @ [x])) then 1 else 0)" apply(auto simp: rec_all.simps) apply(simp add: rec_exec.simps map_append get_fstn_args_take split: if_splits) apply(drule_tac x = "rec_exec rt xs" in rec_accum_ex) apply(cases "rec_exec (rec_accum rf) (xs @ [rec_exec rt xs]) = 0", simp_all) apply force apply(simp add: rec_exec.simps map_append get_fstn_args_take) apply(drule_tac x = "rec_exec rt xs" in rec_accum_ex) apply(cases "rec_exec (rec_accum rf) (xs @ [rec_exec rt xs]) = 0") apply force+ done text \ \rec_ex t f (x1, x2, \, xn)\ computes the charactrization function of the following FOL formula: \(\ x \ t(x1, x2, \, xn). (f(x1, x2, \, xn, x) > 0))\ \ fun rec_ex :: "recf \ recf \ recf" where "rec_ex rt rf = (let vl = arity rf in Cn (vl - 1) rec_sg [Cn (vl - 1) (rec_sigma rf) (get_fstn_args (vl - 1) (vl - 1) @ [rt])])" lemma rec_sigma_ex: assumes "primerec rf (Suc (length xs))" shows "(rec_exec (rec_sigma rf) (xs @ [x]) = 0) = (\ t \ x. rec_exec rf (xs @ [t]) = 0)" proof(induct x) case (Suc x) from Suc assms show ?case by(auto simp add: rec_exec.simps rec_sigma.simps get_fstn_args_take elim:le_SucE) qed (insert assms,auto simp: get_fstn_args_take rec_exec.simps rec_sigma.simps) text \ The correctness of \ex_lemma\. \ lemma ex_lemma:" \primerec rf (Suc (length xs)); primerec rt (length xs)\ \ (rec_exec (rec_ex rt rf) xs = (if (\ x \ (rec_exec rt xs). 0 Definition of \Min[R]\ on page 77 of Boolos's book. \ fun Minr :: "(nat list \ bool) \ nat list \ nat \ nat" where "Minr Rr xs w = (let setx = {y | y. (y \ w) \ Rr (xs @ [y])} in if (setx = {}) then (Suc w) else (Min setx))" declare Minr.simps[simp del] rec_all.simps[simp del] text \ The following is a set of auxilliary lemmas about \Minr\. \ lemma Minr_range: "Minr Rr xs w \ w \ Minr Rr xs w = Suc w" apply(auto simp: Minr.simps) apply(subgoal_tac "Min {x. x \ w \ Rr (xs @ [x])} \ x") apply(erule_tac order_trans, simp) apply(rule_tac Min_le, auto) done lemma expand_conj_in_set: "{x. x \ Suc w \ Rr (xs @ [x])} = (if Rr (xs @ [Suc w]) then insert (Suc w) {x. x \ w \ Rr (xs @ [x])} else {x. x \ w \ Rr (xs @ [x])})" by (auto elim:le_SucE) lemma Minr_strip_Suc[simp]: "Minr Rr xs w \ w \ Minr Rr xs (Suc w) = Minr Rr xs w" by(cases "\x\w. \ Rr (xs @ [x])",auto simp add: Minr.simps expand_conj_in_set) lemma x_empty_set[simp]: "\x\w. \ Rr (xs @ [x]) \ {x. x \ w \ Rr (xs @ [x])} = {} " by auto lemma Minr_is_Suc[simp]: "\Minr Rr xs w = Suc w; Rr (xs @ [Suc w])\ \ Minr Rr xs (Suc w) = Suc w" apply(simp add: Minr.simps expand_conj_in_set) apply(cases "\x\w. \ Rr (xs @ [x])", auto) done lemma Minr_is_Suc_Suc[simp]: "\Minr Rr xs w = Suc w; \ Rr (xs @ [Suc w])\ \ Minr Rr xs (Suc w) = Suc (Suc w)" apply(simp add: Minr.simps expand_conj_in_set) apply(cases "\x\w. \ Rr (xs @ [x])", auto) apply(subgoal_tac "Min {x. x \ w \ Rr (xs @ [x])} \ {x. x \ w \ Rr (xs @ [x])}", simp) apply(rule_tac Min_in, auto) done lemma Minr_Suc_simp: "Minr Rr xs (Suc w) = (if Minr Rr xs w \ w then Minr Rr xs w else if (Rr (xs @ [Suc w])) then (Suc w) else Suc (Suc w))" by(insert Minr_range[of Rr xs w], auto) text \ \rec_Minr\ is the recursive function used to implement \Minr\: if \Rr\ is implemented by a recursive function \recf\, then \rec_Minr recf\ is the recursive function used to implement \Minr Rr\ \ fun rec_Minr :: "recf \ recf" where "rec_Minr rf = (let vl = arity rf in let rq = rec_all (id vl (vl - 1)) (Cn (Suc vl) rec_not [Cn (Suc vl) rf (get_fstn_args (Suc vl) (vl - 1) @ [id (Suc vl) (vl)])]) in rec_sigma rq)" lemma length_getpren_params[simp]: "length (get_fstn_args m n) = n" by(induct n, auto simp: get_fstn_args.simps) lemma length_app: "(length (get_fstn_args (arity rf - Suc 0) (arity rf - Suc 0) @ [Cn (arity rf - Suc 0) (constn 0) [recf.id (arity rf - Suc 0) 0]])) = (Suc (arity rf - Suc 0))" apply(simp) done lemma primerec_accum: "primerec (rec_accum rf) n \ primerec rf n" apply(auto simp: rec_accum.simps Let_def) apply(erule_tac prime_pr_reverse, simp) apply(erule_tac prime_cn_reverse, simp only: length_app) done lemma primerec_all: "primerec (rec_all rt rf) n \ primerec rt n \ primerec rf (Suc n)" apply(simp add: rec_all.simps Let_def) apply(erule_tac prime_cn_reverse, simp) apply(erule_tac prime_cn_reverse, simp) apply(erule_tac x = n in allE, simp add: nth_append primerec_accum) done declare numeral_3_eq_3[simp] lemma primerec_rec_pred_1[intro]: "primerec rec_pred (Suc 0)" apply(simp add: rec_pred_def) apply(rule_tac prime_cn, auto dest:less_2_cases[unfolded numeral One_nat_def]) done lemma primerec_rec_minus_2[intro]: "primerec rec_minus (Suc (Suc 0))" apply(auto simp: rec_minus_def) done lemma primerec_constn_1[intro]: "primerec (constn n) (Suc 0)" apply(induct n) apply(auto simp: constn.simps) done lemma primerec_rec_sg_1[intro]: "primerec rec_sg (Suc 0)" apply(simp add: rec_sg_def) apply(rule_tac k = "Suc (Suc 0)" in prime_cn) apply(auto) apply(auto dest!:less_2_cases[unfolded numeral One_nat_def]) apply( auto) done lemma primerec_getpren[elim]: "\i < n; n \ m\ \ primerec (get_fstn_args m n ! i) m" apply(induct n, auto simp: get_fstn_args.simps) apply(cases "i = n", auto simp: nth_append intro: prime_id) done lemma primerec_rec_add_2[intro]: "primerec rec_add (Suc (Suc 0))" apply(simp add: rec_add_def) apply(rule_tac prime_pr, auto) done lemma primerec_rec_mult_2[intro]:"primerec rec_mult (Suc (Suc 0))" apply(simp add: rec_mult_def ) apply(rule_tac prime_pr, auto) using less_2_cases numeral_2_eq_2 by fastforce lemma primerec_ge_2_elim[elim]: "\primerec rf n; n \ Suc (Suc 0)\ \ primerec (rec_accum rf) n" apply(auto simp: rec_accum.simps) apply(simp add: nth_append, auto dest!:less_2_cases[unfolded numeral One_nat_def]) apply force apply force apply(auto simp: nth_append) done lemma primerec_all_iff: "\primerec rt n; primerec rf (Suc n); n > 0\ \ primerec (rec_all rt rf) n" apply(simp add: rec_all.simps, auto) apply(auto, simp add: nth_append, auto) done lemma primerec_rec_not_1[intro]: "primerec rec_not (Suc 0)" apply(simp add: rec_not_def) apply(rule prime_cn, auto dest!:less_2_cases[unfolded numeral One_nat_def]) done lemma Min_false1[simp]: "\\ Min {uu. uu \ w \ 0 < rec_exec rf (xs @ [uu])} \ w; x \ w; 0 < rec_exec rf (xs @ [x])\ \ False" apply(subgoal_tac "finite {uu. uu \ w \ 0 < rec_exec rf (xs @ [uu])}") apply(subgoal_tac "{uu. uu \ w \ 0 < rec_exec rf (xs @ [uu])} \ {}") apply(simp add: Min_le_iff, simp) apply(rule_tac x = x in exI, simp) apply(simp) done lemma sigma_minr_lemma: assumes prrf: "primerec rf (Suc (length xs))" shows "UF.Sigma (rec_exec (rec_all (recf.id (Suc (length xs)) (length xs)) (Cn (Suc (Suc (length xs))) rec_not [Cn (Suc (Suc (length xs))) rf (get_fstn_args (Suc (Suc (length xs))) (length xs) @ [recf.id (Suc (Suc (length xs))) (Suc (length xs))])]))) (xs @ [w]) = Minr (\args. 0 < rec_exec rf args) xs w" proof(induct w) let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" let ?rf = "(Cn (Suc (Suc (length xs))) rec_not [Cn (Suc (Suc (length xs))) rf (get_fstn_args (Suc (Suc (length xs))) (length xs) @ [recf.id (Suc (Suc (length xs))) (Suc ((length xs)))])])" let ?rq = "(rec_all ?rt ?rf)" have prrf: "primerec ?rf (Suc (length (xs @ [0]))) \ primerec ?rt (length (xs @ [0]))" apply(auto simp: prrf nth_append)+ done show "Sigma (rec_exec (rec_all ?rt ?rf)) (xs @ [0]) = Minr (\args. 0 < rec_exec rf args) xs 0" apply(simp add: Sigma.simps) apply(simp only: prrf all_lemma, auto simp: rec_exec.simps get_fstn_args_take Minr.simps) apply(rule_tac Min_eqI, auto) done next fix w let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" let ?rf = "(Cn (Suc (Suc (length xs))) rec_not [Cn (Suc (Suc (length xs))) rf (get_fstn_args (Suc (Suc (length xs))) (length xs) @ [recf.id (Suc (Suc (length xs))) (Suc ((length xs)))])])" let ?rq = "(rec_all ?rt ?rf)" assume ind: "Sigma (rec_exec (rec_all ?rt ?rf)) (xs @ [w]) = Minr (\args. 0 < rec_exec rf args) xs w" have prrf: "primerec ?rf (Suc (length (xs @ [Suc w]))) \ primerec ?rt (length (xs @ [Suc w]))" apply(auto simp: prrf nth_append)+ done show "UF.Sigma (rec_exec (rec_all ?rt ?rf)) (xs @ [Suc w]) = Minr (\args. 0 < rec_exec rf args) xs (Suc w)" apply(auto simp: Sigma_Suc_simp_rewrite ind Minr_Suc_simp) apply(simp_all only: prrf all_lemma) apply(auto simp: rec_exec.simps get_fstn_args_take Let_def Minr.simps split: if_splits) apply(drule_tac Min_false1, simp, simp, simp) apply (metis le_SucE neq0_conv) apply(drule_tac Min_false1, simp, simp, simp) apply(drule_tac Min_false1, simp, simp, simp) done qed text \ The correctness of \rec_Minr\. \ lemma Minr_lemma: " \primerec rf (Suc (length xs))\ \ rec_exec (rec_Minr rf) (xs @ [w]) = Minr (\ args. (0 < rec_exec rf args)) xs w" proof - let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" let ?rf = "(Cn (Suc (Suc (length xs))) rec_not [Cn (Suc (Suc (length xs))) rf (get_fstn_args (Suc (Suc (length xs))) (length xs) @ [recf.id (Suc (Suc (length xs))) (Suc ((length xs)))])])" let ?rq = "(rec_all ?rt ?rf)" assume h: "primerec rf (Suc (length xs))" have h1: "primerec ?rq (Suc (length xs))" apply(rule_tac primerec_all_iff) apply(auto simp: h nth_append)+ done moreover have "arity rf = Suc (length xs)" using h by auto ultimately show "rec_exec (rec_Minr rf) (xs @ [w]) = Minr (\ args. (0 < rec_exec rf args)) xs w" apply(simp add: arity.simps Let_def sigma_lemma all_lemma) apply(rule_tac sigma_minr_lemma) apply(simp add: h) done qed text \ \rec_le\ is the comparasion function which compares its two arguments, testing whether the first is less or equal to the second. \ definition rec_le :: "recf" where "rec_le = Cn (Suc (Suc 0)) rec_disj [rec_less, rec_eq]" text \ The correctness of \rec_le\. \ lemma le_lemma: "\x y. rec_exec rec_le [x, y] = (if (x \ y) then 1 else 0)" by(auto simp: rec_le_def rec_exec.simps) text \ Definition of \Max[Rr]\ on page 77 of Boolos's book. \ fun Maxr :: "(nat list \ bool) \ nat list \ nat \ nat" where "Maxr Rr xs w = (let setx = {y. y \ w \ Rr (xs @[y])} in if setx = {} then 0 else Max setx)" text \ \rec_maxr\ is the recursive function used to implementation \Maxr\. \ fun rec_maxr :: "recf \ recf" where "rec_maxr rr = (let vl = arity rr in let rt = id (Suc vl) (vl - 1) in let rf1 = Cn (Suc (Suc vl)) rec_le [id (Suc (Suc vl)) ((Suc vl)), id (Suc (Suc vl)) (vl)] in let rf2 = Cn (Suc (Suc vl)) rec_not [Cn (Suc (Suc vl)) rr (get_fstn_args (Suc (Suc vl)) (vl - 1) @ [id (Suc (Suc vl)) ((Suc vl))])] in let rf = Cn (Suc (Suc vl)) rec_disj [rf1, rf2] in let Qf = Cn (Suc vl) rec_not [rec_all rt rf] in Cn vl (rec_sigma Qf) (get_fstn_args vl vl @ [id vl (vl - 1)]))" declare rec_maxr.simps[simp del] Maxr.simps[simp del] declare le_lemma[simp] declare numeral_2_eq_2[simp] lemma primerec_rec_disj_2[intro]: "primerec rec_disj (Suc (Suc 0))" apply(simp add: rec_disj_def, auto) apply(auto dest!:less_2_cases[unfolded numeral One_nat_def]) done lemma primerec_rec_less_2[intro]: "primerec rec_less (Suc (Suc 0))" apply(simp add: rec_less_def, auto) apply(auto dest!:less_2_cases[unfolded numeral One_nat_def]) done lemma primerec_rec_eq_2[intro]: "primerec rec_eq (Suc (Suc 0))" apply(simp add: rec_eq_def) apply(rule_tac prime_cn, auto dest!:less_2_cases[unfolded numeral One_nat_def]) apply force+ done lemma primerec_rec_le_2[intro]: "primerec rec_le (Suc (Suc 0))" apply(simp add: rec_le_def) apply(rule_tac prime_cn, auto dest!:less_2_cases[unfolded numeral One_nat_def]) done lemma Sigma_0: "\ i \ n. (f (xs @ [i]) = 0) \ Sigma f (xs @ [n]) = 0" apply(induct n, simp add: Sigma.simps) apply(simp add: Sigma_Suc_simp_rewrite) done lemma Sigma_Suc[elim]: "\k Sigma f (xs @ [w]) = Suc w" apply(induct w) apply(simp add: Sigma.simps, simp) apply(simp add: Sigma.simps) done lemma Sigma_max_point: "\\ k < ma. f (xs @ [k]) = 1; \ k \ ma. f (xs @ [k]) = 0; ma \ w\ \ Sigma f (xs @ [w]) = ma" apply(induct w, auto) apply(rule_tac Sigma_0, simp) apply(simp add: Sigma_Suc_simp_rewrite) using Sigma_Suc by fastforce lemma Sigma_Max_lemma: assumes prrf: "primerec rf (Suc (length xs))" shows "UF.Sigma (rec_exec (Cn (Suc (Suc (length xs))) rec_not [rec_all (recf.id (Suc (Suc (length xs))) (length xs)) (Cn (Suc (Suc (Suc (length xs)))) rec_disj [Cn (Suc (Suc (Suc (length xs)))) rec_le [recf.id (Suc (Suc (Suc (length xs)))) (Suc (Suc (length xs))), recf.id (Suc (Suc (Suc (length xs)))) (Suc (length xs))], Cn (Suc (Suc (Suc (length xs)))) rec_not [Cn (Suc (Suc (Suc (length xs)))) rf (get_fstn_args (Suc (Suc (Suc (length xs)))) (length xs) @ [recf.id (Suc (Suc (Suc (length xs)))) (Suc (Suc (length xs)))])]])])) ((xs @ [w]) @ [w]) = Maxr (\args. 0 < rec_exec rf args) xs w" proof - let ?rt = "(recf.id (Suc (Suc (length xs))) ((length xs)))" let ?rf1 = "Cn (Suc (Suc (Suc (length xs)))) rec_le [recf.id (Suc (Suc (Suc (length xs)))) ((Suc (Suc (length xs)))), recf.id (Suc (Suc (Suc (length xs)))) ((Suc (length xs)))]" let ?rf2 = "Cn (Suc (Suc (Suc (length xs)))) rf (get_fstn_args (Suc (Suc (Suc (length xs)))) (length xs) @ [recf.id (Suc (Suc (Suc (length xs)))) ((Suc (Suc (length xs))))])" let ?rf3 = "Cn (Suc (Suc (Suc (length xs)))) rec_not [?rf2]" let ?rf = "Cn (Suc (Suc (Suc (length xs)))) rec_disj [?rf1, ?rf3]" let ?rq = "rec_all ?rt ?rf" let ?notrq = "Cn (Suc (Suc (length xs))) rec_not [?rq]" show "?thesis" proof(auto simp: Maxr.simps) assume h: "\x\w. rec_exec rf (xs @ [x]) = 0" have "primerec ?rf (Suc (length (xs @ [w, i]))) \ primerec ?rt (length (xs @ [w, i]))" using prrf apply(auto dest!:less_2_cases[unfolded numeral One_nat_def]) apply force+ apply(case_tac ia, auto simp: h nth_append primerec_getpren) done hence "Sigma (rec_exec ?notrq) ((xs@[w])@[w]) = 0" apply(rule_tac Sigma_0) apply(auto simp: rec_exec.simps all_lemma get_fstn_args_take nth_append h) done thus "UF.Sigma (rec_exec ?notrq) (xs @ [w, w]) = 0" by simp next fix x assume h: "x \ w" "0 < rec_exec rf (xs @ [x])" hence "\ ma. Max {y. y \ w \ 0 < rec_exec rf (xs @ [y])} = ma" by auto from this obtain ma where k1: "Max {y. y \ w \ 0 < rec_exec rf (xs @ [y])} = ma" .. hence k2: "ma \ w \ 0 < rec_exec rf (xs @ [ma])" using h apply(subgoal_tac "Max {y. y \ w \ 0 < rec_exec rf (xs @ [y])} \ {y. y \ w \ 0 < rec_exec rf (xs @ [y])}") apply(erule_tac CollectE, simp) apply(rule_tac Max_in, auto) done hence k3: "\ k < ma. (rec_exec ?notrq (xs @ [w, k]) = 1)" apply(auto simp: nth_append) apply(subgoal_tac "primerec ?rf (Suc (length (xs @ [w, k]))) \ primerec ?rt (length (xs @ [w, k]))") apply(auto simp: rec_exec.simps all_lemma get_fstn_args_take nth_append dest!:less_2_cases[unfolded numeral One_nat_def]) using prrf apply force+ done have k4: "\ k \ ma. (rec_exec ?notrq (xs @ [w, k]) = 0)" apply(auto) apply(subgoal_tac "primerec ?rf (Suc (length (xs @ [w, k]))) \ primerec ?rt (length (xs @ [w, k]))") apply(auto simp: rec_exec.simps all_lemma get_fstn_args_take nth_append) apply(subgoal_tac "x \ Max {y. y \ w \ 0 < rec_exec rf (xs @ [y])}", simp add: k1) apply(rule_tac Max_ge, auto dest!:less_2_cases[unfolded numeral One_nat_def]) using prrf apply force+ apply(auto simp: h nth_append) done from k3 k4 k1 have "Sigma (rec_exec ?notrq) ((xs @ [w]) @ [w]) = ma" apply(rule_tac Sigma_max_point, simp, simp, simp add: k2) done from k1 and this show "Sigma (rec_exec ?notrq) (xs @ [w, w]) = Max {y. y \ w \ 0 < rec_exec rf (xs @ [y])}" by simp qed qed text \ The correctness of \rec_maxr\. \ lemma Maxr_lemma: assumes h: "primerec rf (Suc (length xs))" shows "rec_exec (rec_maxr rf) (xs @ [w]) = Maxr (\ args. 0 < rec_exec rf args) xs w" proof - from h have "arity rf = Suc (length xs)" by auto thus "?thesis" proof(simp add: rec_exec.simps rec_maxr.simps nth_append get_fstn_args_take) let ?rt = "(recf.id (Suc (Suc (length xs))) ((length xs)))" let ?rf1 = "Cn (Suc (Suc (Suc (length xs)))) rec_le [recf.id (Suc (Suc (Suc (length xs)))) ((Suc (Suc (length xs)))), recf.id (Suc (Suc (Suc (length xs)))) ((Suc (length xs)))]" let ?rf2 = "Cn (Suc (Suc (Suc (length xs)))) rf (get_fstn_args (Suc (Suc (Suc (length xs)))) (length xs) @ [recf.id (Suc (Suc (Suc (length xs)))) ((Suc (Suc (length xs))))])" let ?rf3 = "Cn (Suc (Suc (Suc (length xs)))) rec_not [?rf2]" let ?rf = "Cn (Suc (Suc (Suc (length xs)))) rec_disj [?rf1, ?rf3]" let ?rq = "rec_all ?rt ?rf" let ?notrq = "Cn (Suc (Suc (length xs))) rec_not [?rq]" have prt: "primerec ?rt (Suc (Suc (length xs)))" by(auto intro: prime_id) have prrf: "primerec ?rf (Suc (Suc (Suc (length xs))))" apply(auto dest!:less_2_cases[unfolded numeral One_nat_def]) apply force+ apply(auto intro: prime_id) apply(simp add: h) apply(auto simp add: nth_append) done from prt and prrf have prrq: "primerec ?rq (Suc (Suc (length xs)))" by(erule_tac primerec_all_iff, auto) hence prnotrp: "primerec ?notrq (Suc (length ((xs @ [w]))))" by(rule_tac prime_cn, auto) have g1: "rec_exec (rec_sigma ?notrq) ((xs @ [w]) @ [w]) = Maxr (\args. 0 < rec_exec rf args) xs w" using prnotrp using sigma_lemma apply(simp only: sigma_lemma) apply(rule_tac Sigma_Max_lemma) apply(simp add: h) done thus "rec_exec (rec_sigma ?notrq) (xs @ [w, w]) = Maxr (\args. 0 < rec_exec rf args) xs w" apply(simp) done qed qed text \ \quo\ is the formal specification of division. \ fun quo :: "nat list \ nat" where "quo [x, y] = (let Rr = (\ zs. ((zs ! (Suc 0) * zs ! (Suc (Suc 0)) \ zs ! 0) \ zs ! Suc 0 \ (0::nat))) in Maxr Rr [x, y] x)" declare quo.simps[simp del] text \ The following lemmas shows more directly the menaing of \quo\: \ lemma quo_is_div: "y > 0 \ quo [x, y] = x div y" proof - { fix xa ya assume h: "y * ya \ x" "y > 0" hence "(y * ya) div y \ x div y" by(insert div_le_mono[of "y * ya" x y], simp) from this and h have "ya \ x div y" by simp} thus ?thesis by(simp add: quo.simps Maxr.simps, auto, rule_tac Max_eqI, simp, auto) qed lemma quo_zero[intro]: "quo [x, 0] = 0" by(simp add: quo.simps Maxr.simps) lemma quo_div: "quo [x, y] = x div y" by(cases "y=0", auto elim!:quo_is_div) text \ \rec_noteq\ is the recursive function testing whether its two arguments are not equal. \ definition rec_noteq:: "recf" where "rec_noteq = Cn (Suc (Suc 0)) rec_not [Cn (Suc (Suc 0)) rec_eq [id (Suc (Suc 0)) (0), id (Suc (Suc 0)) ((Suc 0))]]" text \ The correctness of \rec_noteq\. \ lemma noteq_lemma: "\ x y. rec_exec rec_noteq [x, y] = (if x \ y then 1 else 0)" by(simp add: rec_exec.simps rec_noteq_def) declare noteq_lemma[simp] text \ \rec_quo\ is the recursive function used to implement \quo\ \ definition rec_quo :: "recf" where "rec_quo = (let rR = Cn (Suc (Suc (Suc 0))) rec_conj [Cn (Suc (Suc (Suc 0))) rec_le [Cn (Suc (Suc (Suc 0))) rec_mult [id (Suc (Suc (Suc 0))) (Suc 0), id (Suc (Suc (Suc 0))) ((Suc (Suc 0)))], id (Suc (Suc (Suc 0))) (0)], Cn (Suc (Suc (Suc 0))) rec_noteq [id (Suc (Suc (Suc 0))) (Suc (0)), Cn (Suc (Suc (Suc 0))) (constn 0) [id (Suc (Suc (Suc 0))) (0)]]] in Cn (Suc (Suc 0)) (rec_maxr rR)) [id (Suc (Suc 0)) (0),id (Suc (Suc 0)) (Suc (0)), id (Suc (Suc 0)) (0)]" lemma primerec_rec_conj_2[intro]: "primerec rec_conj (Suc (Suc 0))" apply(simp add: rec_conj_def) apply(rule_tac prime_cn, auto dest!:less_2_cases[unfolded numeral One_nat_def]) done lemma primerec_rec_noteq_2[intro]: "primerec rec_noteq (Suc (Suc 0))" apply(simp add: rec_noteq_def) apply(rule_tac prime_cn, auto dest!:less_2_cases[unfolded numeral One_nat_def]) done lemma quo_lemma1: "rec_exec rec_quo [x, y] = quo [x, y]" proof(simp add: rec_exec.simps rec_quo_def) let ?rR = "(Cn (Suc (Suc (Suc 0))) rec_conj [Cn (Suc (Suc (Suc 0))) rec_le [Cn (Suc (Suc (Suc 0))) rec_mult [recf.id (Suc (Suc (Suc 0))) (Suc (0)), recf.id (Suc (Suc (Suc 0))) (Suc (Suc (0)))], recf.id (Suc (Suc (Suc 0))) (0)], Cn (Suc (Suc (Suc 0))) rec_noteq [recf.id (Suc (Suc (Suc 0))) (Suc (0)), Cn (Suc (Suc (Suc 0))) (constn 0) [recf.id (Suc (Suc (Suc 0))) (0)]]])" have "rec_exec (rec_maxr ?rR) ([x, y]@ [ x]) = Maxr (\ args. 0 < rec_exec ?rR args) [x, y] x" proof(rule_tac Maxr_lemma, simp) show "primerec ?rR (Suc (Suc (Suc 0)))" apply(auto dest!:less_2_cases[unfolded numeral One_nat_def]) apply force+ done qed hence g1: "rec_exec (rec_maxr ?rR) ([x, y, x]) = Maxr (\ args. if rec_exec ?rR args = 0 then False else True) [x, y] x" by simp have g2: "Maxr (\ args. if rec_exec ?rR args = 0 then False else True) [x, y] x = quo [x, y]" apply(simp add: rec_exec.simps) apply(simp add: Maxr.simps quo.simps, auto) done from g1 and g2 show "rec_exec (rec_maxr ?rR) ([x, y, x]) = quo [x, y]" by simp qed text \ The correctness of \quo\. \ lemma quo_lemma2: "rec_exec rec_quo [x, y] = x div y" using quo_lemma1[of x y] quo_div[of x y] by simp text \ \rec_mod\ is the recursive function used to implement the reminder function. \ definition rec_mod :: "recf" where "rec_mod = Cn (Suc (Suc 0)) rec_minus [id (Suc (Suc 0)) (0), Cn (Suc (Suc 0)) rec_mult [rec_quo, id (Suc (Suc 0)) (Suc (0))]]" text \ The correctness of \rec_mod\: \ lemma mod_lemma: "\ x y. rec_exec rec_mod [x, y] = (x mod y)" by(simp add: rec_exec.simps rec_mod_def quo_lemma2 minus_div_mult_eq_mod) text\lemmas for embranch function\ type_synonym ftype = "nat list \ nat" type_synonym rtype = "nat list \ bool" text \ The specifation of the mutli-way branching statement on page 79 of Boolos's book. \ fun Embranch :: "(ftype * rtype) list \ nat list \ nat" where "Embranch [] xs = 0" | "Embranch (gc # gcs) xs = ( let (g, c) = gc in if c xs then g xs else Embranch gcs xs)" fun rec_embranch' :: "(recf * recf) list \ nat \ recf" where "rec_embranch' [] vl = Cn vl z [id vl (vl - 1)]" | "rec_embranch' ((rg, rc) # rgcs) vl = Cn vl rec_add [Cn vl rec_mult [rg, rc], rec_embranch' rgcs vl]" text \ \rec_embrach\ is the recursive function used to implement \Embranch\. \ fun rec_embranch :: "(recf * recf) list \ recf" where "rec_embranch ((rg, rc) # rgcs) = (let vl = arity rg in rec_embranch' ((rg, rc) # rgcs) vl)" declare Embranch.simps[simp del] rec_embranch.simps[simp del] lemma embranch_all0: "\\ j < length rcs. rec_exec (rcs ! j) xs = 0; length rgs = length rcs; rcs \ []; list_all (\ rf. primerec rf (length xs)) (rgs @ rcs)\ \ rec_exec (rec_embranch (zip rgs rcs)) xs = 0" proof(induct rcs arbitrary: rgs) case (Cons a rcs) then show ?case proof(cases rgs, simp) fix a rcs rgs aa list assume ind: "\rgs. \\j []; list_all (\rf. primerec rf (length xs)) (rgs @ rcs)\ \ rec_exec (rec_embranch (zip rgs rcs)) xs = 0" and h: "\j []" "list_all (\rf. primerec rf (length xs)) (rgs @ a # rcs)" "rgs = aa # list" have g: "rcs \ [] \ rec_exec (rec_embranch (zip list rcs)) xs = 0" using h by(rule_tac ind, auto) show "rec_exec (rec_embranch (zip rgs (a # rcs))) xs = 0" proof(cases "rcs = []", simp) show "rec_exec (rec_embranch (zip rgs [a])) xs = 0" using h by (auto simp add: rec_embranch.simps rec_exec.simps) next assume "rcs \ []" hence "rec_exec (rec_embranch (zip list rcs)) xs = 0" using g by simp thus "rec_exec (rec_embranch (zip rgs (a # rcs))) xs = 0" using h by(cases rcs;cases list, auto simp add: rec_embranch.simps rec_exec.simps) qed qed qed simp lemma embranch_exec_0: "\rec_exec aa xs = 0; zip rgs list \ []; list_all (\ rf. primerec rf (length xs)) ([a, aa] @ rgs @ list)\ \ rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = rec_exec (rec_embranch (zip rgs list)) xs" apply(auto simp add: rec_exec.simps rec_embranch.simps) apply(cases "zip rgs list", force) apply(cases "hd (zip rgs list)", simp add: rec_embranch.simps rec_exec.simps) apply(subgoal_tac "arity a = length xs") apply(cases rgs;cases list;force) by force lemma zip_null_iff: "\length xs = k; length ys = k; zip xs ys = []\ \ xs = [] \ ys = []" - apply(cases xs, simp, simp) - apply(cases ys, simp, simp) - done +by(cases xs, simp, simp) lemma zip_null_gr: "\length xs = k; length ys = k; zip xs ys \ []\ \ 0 < k" - apply(cases xs, simp, simp) - done +by(cases xs, simp, simp) lemma Embranch_0: "\length rgs = k; length rcs = k; k > 0; \ j < k. rec_exec (rcs ! j) xs = 0\ \ Embranch (zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) rcs)) xs = 0" proof(induct rgs arbitrary: rcs k) case (Cons a rgs rcs k) then show ?case apply(cases rcs, simp, cases "rgs = []") apply(simp add: Embranch.simps) apply(erule_tac x = 0 in allE) apply (auto simp add: Embranch.simps intro!: Cons(1)). qed simp text \ The correctness of \rec_embranch\. \ lemma embranch_lemma: assumes branch_num: "length rgs = n" "length rcs = n" "n > 0" and partition: "(\ i < n. (rec_exec (rcs ! i) xs = 1 \ (\ j < n. j \ i \ rec_exec (rcs ! j) xs = 0)))" and prime_all: "list_all (\ rf. primerec rf (length xs)) (rgs @ rcs)" shows "rec_exec (rec_embranch (zip rgs rcs)) xs = Embranch (zip (map rec_exec rgs) (map (\ r args. 0 < rec_exec r args) rcs)) xs" using branch_num partition prime_all proof(induct rgs arbitrary: rcs n, simp) fix a rgs rcs n assume ind: "\rcs n. \length rgs = n; length rcs = n; 0 < n; \i (\j i \ rec_exec (rcs ! j) xs = 0); list_all (\rf. primerec rf (length xs)) (rgs @ rcs)\ \ rec_exec (rec_embranch (zip rgs rcs)) xs = Embranch (zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) rcs)) xs" and h: "length (a # rgs) = n" "length (rcs::recf list) = n" "0 < n" " \i (\j i \ rec_exec (rcs ! j) xs = 0)" "list_all (\rf. primerec rf (length xs)) ((a # rgs) @ rcs)" from h show "rec_exec (rec_embranch (zip (a # rgs) rcs)) xs = Embranch (zip (map rec_exec (a # rgs)) (map (\r args. 0 < rec_exec r args) rcs)) xs" apply(cases rcs, simp, simp) apply(cases "rec_exec (hd rcs) xs = 0") apply(case_tac [!] "zip rgs (tl rcs) = []", simp) apply(subgoal_tac "rgs = [] \ (tl rcs) = []", simp add: Embranch.simps rec_exec.simps rec_embranch.simps) apply(rule_tac zip_null_iff, simp, simp, simp) proof - fix aa list assume "rcs = aa # list" assume g: "Suc (length rgs) = n" "Suc (length list) = n" "\i (\j i \ rec_exec ((aa # list) ! j) xs = 0)" "primerec a (length xs) \ list_all (\rf. primerec rf (length xs)) rgs \ primerec aa (length xs) \ list_all (\rf. primerec rf (length xs)) list" "rec_exec (hd rcs) xs = 0" "rcs = aa # list" "zip rgs (tl rcs) \ []" hence "rec_exec aa xs = 0" "zip rgs list \ []" by auto - note g = g(1,2,3,4,6) this - have "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs + note g = g(1,2,3,4,6,7) this + hence "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = rec_exec (rec_embranch (zip rgs list)) xs" - apply(rule embranch_exec_0, simp_all add: g) - done + by(simp add: embranch_exec_0) from g and this show "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = Embranch ((rec_exec a, \args. 0 < rec_exec aa args) # zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) list)) xs" apply(simp add: Embranch.simps) apply(rule_tac n = "n - Suc 0" in ind) apply(cases n;force) apply(cases n;force) apply(cases n;force simp add: zip_null_gr) apply(auto) apply(rename_tac i) apply(case_tac i, force, simp) apply(rule_tac x = "i - 1" in exI, simp) by auto next fix aa list assume g: "Suc (length rgs) = n" "Suc (length list) = n" "\i (\j i \ rec_exec ((aa # list) ! j) xs = 0)" "primerec a (length xs) \ list_all (\rf. primerec rf (length xs)) rgs \ primerec aa (length xs) \ list_all (\rf. primerec rf (length xs)) list" "rcs = aa # list" "rec_exec (hd rcs) xs \ 0" "zip rgs (tl rcs) = []" thus "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = Embranch ((rec_exec a, \args. 0 < rec_exec aa args) # zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) list)) xs" apply(subgoal_tac "rgs = [] \ list = []", simp) prefer 2 apply(rule_tac zip_null_iff, simp, simp, simp) apply(simp add: rec_exec.simps rec_embranch.simps Embranch.simps, auto) done next fix aa list assume g: "Suc (length rgs) = n" "Suc (length list) = n" "\i (\j i \ rec_exec ((aa # list) ! j) xs = 0)" "primerec a (length xs) \ list_all (\rf. primerec rf (length xs)) rgs \ primerec aa (length xs) \ list_all (\rf. primerec rf (length xs)) list" "rcs = aa # list" "rec_exec (hd rcs) xs \ 0" "zip rgs (tl rcs) \ []" have "rec_exec aa xs = Suc 0" using g apply(cases "rec_exec aa xs", simp, auto) done moreover have "rec_exec (rec_embranch' (zip rgs list) (length xs)) xs = 0" proof - have "rec_embranch' (zip rgs list) (length xs) = rec_embranch (zip rgs list)" using g apply(cases "zip rgs list", force) apply(cases "hd (zip rgs list)") apply(simp add: rec_embranch.simps) apply(cases rgs, simp, simp, cases list, simp, auto) done moreover have "rec_exec (rec_embranch (zip rgs list)) xs = 0" proof(rule embranch_all0) show " \j []" using g by(cases list; force) next show "list_all (\rf. primerec rf (length xs)) (rgs @ list)" using g by auto qed ultimately show "rec_exec (rec_embranch' (zip rgs list) (length xs)) xs = 0" by simp qed moreover have "Embranch (zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) list)) xs = 0" using g apply(rule_tac k = "length rgs" in Embranch_0) apply(simp, cases n, simp, simp) apply(cases rgs, simp, simp) apply(auto) apply(rename_tac i j) apply(case_tac i, simp) apply(erule_tac x = "Suc j" in allE, simp) apply(simp) apply(rule_tac x = 0 in allE, auto) done moreover have "arity a = length xs" using g apply(auto) done ultimately show "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = Embranch ((rec_exec a, \args. 0 < rec_exec aa args) # zip (map rec_exec rgs) (map (\r args. 0 < rec_exec r args) list)) xs" apply(simp add: rec_exec.simps rec_embranch.simps Embranch.simps) done qed qed text\ \prime n\ means \n\ is a prime number. \ fun Prime :: "nat \ bool" where "Prime x = (1 < x \ (\ u < x. (\ v < x. u * v \ x)))" declare Prime.simps [simp del] lemma primerec_all1: "primerec (rec_all rt rf) n \ primerec rt n" by (simp add: primerec_all) lemma primerec_all2: "primerec (rec_all rt rf) n \ primerec rf (Suc n)" by(insert primerec_all[of rt rf n], simp) text \ \rec_prime\ is the recursive function used to implement \Prime\. \ definition rec_prime :: "recf" where "rec_prime = Cn (Suc 0) rec_conj [Cn (Suc 0) rec_less [constn 1, id (Suc 0) (0)], rec_all (Cn 1 rec_minus [id 1 0, constn 1]) (rec_all (Cn 2 rec_minus [id 2 0, Cn 2 (constn 1) [id 2 0]]) (Cn 3 rec_noteq [Cn 3 rec_mult [id 3 1, id 3 2], id 3 0]))]" declare numeral_2_eq_2[simp del] numeral_3_eq_3[simp del] lemma exec_tmp: "rec_exec (rec_all (Cn 2 rec_minus [recf.id 2 0, Cn 2 (constn (Suc 0)) [recf.id 2 0]]) (Cn 3 rec_noteq [Cn 3 rec_mult [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0])) [x, k] = ((if (\w\rec_exec (Cn 2 rec_minus [recf.id 2 0, Cn 2 (constn (Suc 0)) [recf.id 2 0]]) ([x, k]). 0 < rec_exec (Cn 3 rec_noteq [Cn 3 rec_mult [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0]) ([x, k] @ [w])) then 1 else 0))" apply(rule_tac all_lemma) apply(auto simp:numeral) apply (metis (no_types, lifting) Suc_mono length_Cons less_2_cases list.size(3) nth_Cons_0 nth_Cons_Suc numeral_2_eq_2 prime_cn prime_id primerec_rec_mult_2 zero_less_Suc) by (metis (no_types, lifting) One_nat_def length_Cons less_2_cases nth_Cons_0 nth_Cons_Suc prime_cn_reverse primerec_rec_eq_2 rec_eq_def zero_less_Suc) text \ The correctness of \Prime\. \ lemma prime_lemma: "rec_exec rec_prime [x] = (if Prime x then 1 else 0)" proof(simp add: rec_exec.simps rec_prime_def) let ?rt1 = "(Cn 2 rec_minus [recf.id 2 0, Cn 2 (constn (Suc 0)) [recf.id 2 0]])" let ?rf1 = "(Cn 3 rec_noteq [Cn 3 rec_mult [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 (0)])" let ?rt2 = "(Cn (Suc 0) rec_minus [recf.id (Suc 0) 0, constn (Suc 0)])" let ?rf2 = "rec_all ?rt1 ?rf1" have h1: "rec_exec (rec_all ?rt2 ?rf2) ([x]) = (if (\k\rec_exec ?rt2 ([x]). 0 < rec_exec ?rf2 ([x] @ [k])) then 1 else 0)" proof(rule_tac all_lemma, simp_all) show "primerec ?rf2 (Suc (Suc 0))" apply(rule_tac primerec_all_iff) apply(auto simp: numeral) apply (metis (no_types, lifting) One_nat_def length_Cons less_2_cases nth_Cons_0 nth_Cons_Suc prime_cn_reverse primerec_rec_eq_2 rec_eq_def zero_less_Suc) by (metis (no_types, lifting) Suc_mono length_Cons less_2_cases list.size(3) nth_Cons_0 nth_Cons_Suc numeral_2_eq_2 prime_cn prime_id primerec_rec_mult_2 zero_less_Suc) next show "primerec (Cn (Suc 0) rec_minus [recf.id (Suc 0) 0, constn (Suc 0)]) (Suc 0)" using less_2_cases numeral by fastforce qed from h1 show "(Suc 0 < x \ (rec_exec (rec_all ?rt2 ?rf2) [x] = 0 \ \ Prime x) \ (0 < rec_exec (rec_all ?rt2 ?rf2) [x] \ Prime x)) \ (\ Suc 0 < x \ \ Prime x \ (rec_exec (rec_all ?rt2 ?rf2) [x] = 0 \ \ Prime x))" apply(auto simp:rec_exec.simps) apply(simp add: exec_tmp rec_exec.simps) proof - assume *:"\k\x - Suc 0. (0::nat) < (if \w\x - Suc 0. 0 < (if k * w \ x then 1 else (0 :: nat)) then 1 else 0)" "Suc 0 < x" thus "Prime x" apply(simp add: rec_exec.simps split: if_splits) apply(simp add: Prime.simps, auto) apply(rename_tac u v) apply(erule_tac x = u in allE, auto) apply(case_tac u, simp) apply(case_tac "u - 1", simp, simp) apply(case_tac v, simp) apply(case_tac "v - 1", simp, simp) done next assume "\ Suc 0 < x" "Prime x" thus "False" apply(simp add: Prime.simps) done next fix k assume "rec_exec (rec_all ?rt1 ?rf1) [x, k] = 0" "k \ x - Suc 0" "Prime x" thus "False" apply(simp add: exec_tmp rec_exec.simps Prime.simps split: if_splits) done next fix k assume "rec_exec (rec_all ?rt1 ?rf1) [x, k] = 0" "k \ x - Suc 0" "Prime x" thus "False" apply(simp add: exec_tmp rec_exec.simps Prime.simps split: if_splits) done qed qed definition rec_dummyfac :: "recf" where "rec_dummyfac = Pr 1 (constn 1) (Cn 3 rec_mult [id 3 2, Cn 3 s [id 3 1]])" text \ The recursive function used to implment factorization. \ definition rec_fac :: "recf" where "rec_fac = Cn 1 rec_dummyfac [id 1 0, id 1 0]" text \ Formal specification of factorization. \ fun fac :: "nat \ nat" ("_!" [100] 99) where "fac 0 = 1" | "fac (Suc x) = (Suc x) * fac x" lemma fac_dummy: "rec_exec rec_dummyfac [x, y] = y !" apply(induct y) apply(auto simp: rec_dummyfac_def rec_exec.simps) done text \ The correctness of \rec_fac\. \ lemma fac_lemma: "rec_exec rec_fac [x] = x!" apply(simp add: rec_fac_def rec_exec.simps fac_dummy) done declare fac.simps[simp del] text \ \Np x\ returns the first prime number after \x\. \ fun Np ::"nat \ nat" where "Np x = Min {y. y \ Suc (x!) \ x < y \ Prime y}" declare Np.simps[simp del] rec_Minr.simps[simp del] text \ \rec_np\ is the recursive function used to implement \Np\. \ definition rec_np :: "recf" where "rec_np = (let Rr = Cn 2 rec_conj [Cn 2 rec_less [id 2 0, id 2 1], Cn 2 rec_prime [id 2 1]] in Cn 1 (rec_Minr Rr) [id 1 0, Cn 1 s [rec_fac]])" lemma n_le_fact[simp]: "n < Suc (n!)" proof(induct n) case (Suc n) then show ?case apply(simp add: fac.simps) apply(cases n, auto simp: fac.simps) done qed simp lemma divsor_ex: "\\ Prime x; x > Suc 0\ \ (\ u > Suc 0. (\ v > Suc 0. u * v = x))" by(auto simp: Prime.simps) lemma divsor_prime_ex: "\\ Prime x; x > Suc 0\ \ \ p. Prime p \ p dvd x" apply(induct x rule: wf_induct[where r = "measure (\ y. y)"], simp) apply(drule_tac divsor_ex, simp, auto) apply(rename_tac u v) apply(erule_tac x = u in allE, simp) apply(case_tac "Prime u", simp) apply(rule_tac x = u in exI, simp, auto) done lemma fact_pos[intro]: "0 < n!" apply(induct n) apply(auto simp: fac.simps) done lemma fac_Suc: "Suc n! = (Suc n) * (n!)" by(simp add: fac.simps) lemma fac_dvd: "\0 < q; q \ n\ \ q dvd n!" proof(induct n) case (Suc n) then show ?case apply(cases "q \ n", simp add: fac_Suc) apply(subgoal_tac "q = Suc n", simp only: fac_Suc) apply(rule_tac dvd_mult2, simp, simp) done qed simp lemma fac_dvd2: "\Suc 0 < q; q dvd n!; q \ n\ \ \ q dvd Suc (n!)" proof(auto simp: dvd_def) fix k ka assume h1: "Suc 0 < q" "q \ n" and h2: "Suc (q * k) = q * ka" have "k < ka" proof - have "q * k < q * ka" using h2 by arith thus "k < ka" using h1 by(auto) qed hence "\d. d > 0 \ ka = d + k" by(rule_tac x = "ka - k" in exI, simp) from this obtain d where "d > 0 \ ka = d + k" .. from h2 and this and h1 show "False" by(simp add: add_mult_distrib2) qed lemma prime_ex: "\ p. n < p \ p \ Suc (n!) \ Prime p" proof(cases "Prime (n! + 1)") case True thus "?thesis" by(rule_tac x = "Suc (n!)" in exI, simp) next assume h: "\ Prime (n! + 1)" hence "\ p. Prime p \ p dvd (n! + 1)" by(erule_tac divsor_prime_ex, auto) from this obtain q where k: "Prime q \ q dvd (n! + 1)" .. thus "?thesis" proof(cases "q > n") case True thus "?thesis" using k by(auto intro:dvd_imp_le) next case False thus "?thesis" proof - assume g: "\ n < q" have j: "q > Suc 0" using k by(cases q, auto simp: Prime.simps) hence "q dvd n!" using g apply(rule_tac fac_dvd, auto) done hence "\ q dvd Suc (n!)" using g j by(rule_tac fac_dvd2, auto) thus "?thesis" using k by simp qed qed qed lemma Suc_Suc_induct[elim!]: "\i < Suc (Suc 0); primerec (ys ! 0) n; primerec (ys ! 1) n\ \ primerec (ys ! i) n" by(cases i, auto) lemma primerec_rec_prime_1[intro]: "primerec rec_prime (Suc 0)" apply(auto simp: rec_prime_def, auto) apply(rule_tac primerec_all_iff, auto, auto) apply(rule_tac primerec_all_iff, auto, auto simp: numeral_2_eq_2 numeral_3_eq_3) done text \ The correctness of \rec_np\. \ lemma np_lemma: "rec_exec rec_np [x] = Np x" proof(auto simp: rec_np_def rec_exec.simps Let_def fac_lemma) let ?rr = "(Cn 2 rec_conj [Cn 2 rec_less [recf.id 2 0, recf.id 2 (Suc 0)], Cn 2 rec_prime [recf.id 2 (Suc 0)]])" let ?R = "\ zs. zs ! 0 < zs ! 1 \ Prime (zs ! 1)" have g1: "rec_exec (rec_Minr ?rr) ([x] @ [Suc (x!)]) = Minr (\ args. 0 < rec_exec ?rr args) [x] (Suc (x!))" by(rule_tac Minr_lemma, auto simp: rec_exec.simps prime_lemma, auto simp: numeral_2_eq_2 numeral_3_eq_3) have g2: "Minr (\ args. 0 < rec_exec ?rr args) [x] (Suc (x!)) = Np x" using prime_ex[of x] apply(auto simp: Minr.simps Np.simps rec_exec.simps prime_lemma) apply(subgoal_tac "{uu. (Prime uu \ (x < uu \ uu \ Suc (x!)) \ x < uu) \ Prime uu} = {y. y \ Suc (x!) \ x < y \ Prime y}", auto) done from g1 and g2 show "rec_exec (rec_Minr ?rr) ([x, Suc (x!)]) = Np x" by simp qed text \ \rec_power\ is the recursive function used to implement power function. \ definition rec_power :: "recf" where "rec_power = Pr 1 (constn 1) (Cn 3 rec_mult [id 3 0, id 3 2])" text \ The correctness of \rec_power\. \ lemma power_lemma: "rec_exec rec_power [x, y] = x^y" by(induct y, auto simp: rec_exec.simps rec_power_def) text\ \Pi k\ returns the \k\-th prime number. \ fun Pi :: "nat \ nat" where "Pi 0 = 2" | "Pi (Suc x) = Np (Pi x)" definition rec_dummy_pi :: "recf" where "rec_dummy_pi = Pr 1 (constn 2) (Cn 3 rec_np [id 3 2])" text \ \rec_pi\ is the recursive function used to implement \Pi\. \ definition rec_pi :: "recf" where "rec_pi = Cn 1 rec_dummy_pi [id 1 0, id 1 0]" lemma pi_dummy_lemma: "rec_exec rec_dummy_pi [x, y] = Pi y" apply(induct y) by(auto simp: rec_exec.simps rec_dummy_pi_def Pi.simps np_lemma) text \ The correctness of \rec_pi\. \ lemma pi_lemma: "rec_exec rec_pi [x] = Pi x" apply(simp add: rec_pi_def rec_exec.simps pi_dummy_lemma) done fun loR :: "nat list \ bool" where "loR [x, y, u] = (x mod (y^u) = 0)" declare loR.simps[simp del] text \ \Lo\ specifies the \lo\ function given on page 79 of Boolos's book. It is one of the two notions of integeral logarithmetic operation on that page. The other is \lg\. \ fun lo :: " nat \ nat \ nat" where "lo x y = (if x > 1 \ y > 1 \ {u. loR [x, y, u]} \ {} then Max {u. loR [x, y, u]} else 0)" declare lo.simps[simp del] lemma primerec_sigma[intro!]: "\n > Suc 0; primerec rf n\ \ primerec (rec_sigma rf) n" apply(simp add: rec_sigma.simps) apply(auto, auto simp: nth_append) done lemma primerec_rec_maxr[intro!]: "\primerec rf n; n > 0\ \ primerec (rec_maxr rf) n" apply(simp add: rec_maxr.simps) apply(rule_tac prime_cn, auto) apply(rule_tac primerec_all_iff, auto, auto simp: nth_append) done lemma Suc_Suc_Suc_induct[elim!]: "\i < Suc (Suc (Suc (0::nat))); primerec (ys ! 0) n; primerec (ys ! 1) n; primerec (ys ! 2) n\ \ primerec (ys ! i) n" apply(cases i, auto) apply(cases "i-1", simp, simp add: numeral_2_eq_2) done lemma primerec_2[intro]: "primerec rec_quo (Suc (Suc 0))" "primerec rec_mod (Suc (Suc 0))" "primerec rec_power (Suc (Suc 0))" by(force simp: prime_cn prime_id rec_mod_def rec_quo_def rec_power_def prime_pr numeral)+ text \ \rec_lo\ is the recursive function used to implement \Lo\. \ definition rec_lo :: "recf" where "rec_lo = (let rR = Cn 3 rec_eq [Cn 3 rec_mod [id 3 0, Cn 3 rec_power [id 3 1, id 3 2]], Cn 3 (constn 0) [id 3 1]] in let rb = Cn 2 (rec_maxr rR) [id 2 0, id 2 1, id 2 0] in let rcond = Cn 2 rec_conj [Cn 2 rec_less [Cn 2 (constn 1) [id 2 0], id 2 0], Cn 2 rec_less [Cn 2 (constn 1) [id 2 0], id 2 1]] in let rcond2 = Cn 2 rec_minus [Cn 2 (constn 1) [id 2 0], rcond] in Cn 2 rec_add [Cn 2 rec_mult [rb, rcond], Cn 2 rec_mult [Cn 2 (constn 0) [id 2 0], rcond2]])" lemma rec_lo_Maxr_lor: "\Suc 0 < x; Suc 0 < y\ \ rec_exec rec_lo [x, y] = Maxr loR [x, y] x" proof(auto simp: rec_exec.simps rec_lo_def Let_def numeral_2_eq_2 numeral_3_eq_3) let ?rR = "(Cn (Suc (Suc (Suc 0))) rec_eq [Cn (Suc (Suc (Suc 0))) rec_mod [recf.id (Suc (Suc (Suc 0))) 0, Cn (Suc (Suc (Suc 0))) rec_power [recf.id (Suc (Suc (Suc 0))) (Suc 0), recf.id (Suc (Suc (Suc 0))) (Suc (Suc 0))]], Cn (Suc (Suc (Suc 0))) (constn 0) [recf.id (Suc (Suc (Suc 0))) (Suc 0)]])" have h: "rec_exec (rec_maxr ?rR) ([x, y] @ [x]) = Maxr (\ args. 0 < rec_exec ?rR args) [x, y] x" by(rule_tac Maxr_lemma, auto simp: rec_exec.simps mod_lemma power_lemma, auto simp: numeral_2_eq_2 numeral_3_eq_3) have "Maxr loR [x, y] x = Maxr (\ args. 0 < rec_exec ?rR args) [x, y] x" apply(simp add: rec_exec.simps mod_lemma power_lemma) apply(simp add: Maxr.simps loR.simps) done from h and this show "rec_exec (rec_maxr ?rR) [x, y, x] = Maxr loR [x, y] x" apply(simp) done qed lemma x_less_exp: "\y > Suc 0\ \ x < y^x" proof(induct x) case (Suc x) then show ?case apply(cases x, simp, auto) apply(rule_tac y = "y* y^(x-1)" in le_less_trans, auto) done qed simp lemma uplimit_loR: assumes "Suc 0 < x" "Suc 0 < y" "loR [x, y, xa]" shows "xa \ x" proof - have "Suc 0 < x \ Suc 0 < y \ y ^ xa dvd x \ xa \ x" by (meson Suc_lessD le_less_trans nat_dvd_not_less nat_le_linear x_less_exp) thus ?thesis using assms by(auto simp: loR.simps) qed lemma loR_set_strengthen[simp]: "\xa \ x; loR [x, y, xa]; Suc 0 < x; Suc 0 < y\ \ {u. loR [x, y, u]} = {ya. ya \ x \ loR [x, y, ya]}" apply(rule_tac Collect_cong, auto) apply(erule_tac uplimit_loR, simp, simp) done lemma Maxr_lo: "\Suc 0 < x; Suc 0 < y\ \ Maxr loR [x, y] x = lo x y" apply(simp add: Maxr.simps lo.simps, auto simp: uplimit_loR) by (meson uplimit_loR)+ lemma lo_lemma': "\Suc 0 < x; Suc 0 < y\ \ rec_exec rec_lo [x, y] = lo x y" by(simp add: Maxr_lo rec_lo_Maxr_lor) lemma lo_lemma'': "\\ Suc 0 < x\ \ rec_exec rec_lo [x, y] = lo x y" apply(cases x, auto simp: rec_exec.simps rec_lo_def Let_def lo.simps) done lemma lo_lemma''': "\\ Suc 0 < y\ \ rec_exec rec_lo [x, y] = lo x y" apply(cases y, auto simp: rec_exec.simps rec_lo_def Let_def lo.simps) done text \ The correctness of \rec_lo\: \ lemma lo_lemma: "rec_exec rec_lo [x, y] = lo x y" apply(cases "Suc 0 < x \ Suc 0 < y") apply(auto simp: lo_lemma' lo_lemma'' lo_lemma''') done fun lgR :: "nat list \ bool" where "lgR [x, y, u] = (y^u \ x)" text \ \lg\ specifies the \lg\ function given on page 79 of Boolos's book. It is one of the two notions of integeral logarithmetic operation on that page. The other is \lo\. \ fun lg :: "nat \ nat \ nat" where "lg x y = (if x > 1 \ y > 1 \ {u. lgR [x, y, u]} \ {} then Max {u. lgR [x, y, u]} else 0)" declare lg.simps[simp del] lgR.simps[simp del] text \ \rec_lg\ is the recursive function used to implement \lg\. \ definition rec_lg :: "recf" where "rec_lg = (let rec_lgR = Cn 3 rec_le [Cn 3 rec_power [id 3 1, id 3 2], id 3 0] in let conR1 = Cn 2 rec_conj [Cn 2 rec_less [Cn 2 (constn 1) [id 2 0], id 2 0], Cn 2 rec_less [Cn 2 (constn 1) [id 2 0], id 2 1]] in let conR2 = Cn 2 rec_not [conR1] in Cn 2 rec_add [Cn 2 rec_mult [conR1, Cn 2 (rec_maxr rec_lgR) [id 2 0, id 2 1, id 2 0]], Cn 2 rec_mult [conR2, Cn 2 (constn 0) [id 2 0]]])" lemma lg_maxr: "\Suc 0 < x; Suc 0 < y\ \ rec_exec rec_lg [x, y] = Maxr lgR [x, y] x" proof(simp add: rec_exec.simps rec_lg_def Let_def) assume h: "Suc 0 < x" "Suc 0 < y" let ?rR = "(Cn 3 rec_le [Cn 3 rec_power [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0])" have "rec_exec (rec_maxr ?rR) ([x, y] @ [x]) = Maxr ((\ args. 0 < rec_exec ?rR args)) [x, y] x" proof(rule Maxr_lemma) show "primerec (Cn 3 rec_le [Cn 3 rec_power [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0]) (Suc (length [x, y]))" apply(auto simp: numeral_3_eq_3)+ done qed moreover have "Maxr lgR [x, y] x = Maxr ((\ args. 0 < rec_exec ?rR args)) [x, y] x" apply(simp add: rec_exec.simps power_lemma) apply(simp add: Maxr.simps lgR.simps) done ultimately show "rec_exec (rec_maxr ?rR) [x, y, x] = Maxr lgR [x, y] x" by simp qed lemma lgR_ok: "\Suc 0 < y; lgR [x, y, xa]\ \ xa \ x" apply(auto simp add: lgR.simps) apply(subgoal_tac "y^xa > xa", simp) apply(erule x_less_exp) done lemma lgR_set_strengthen[simp]: "\Suc 0 < x; Suc 0 < y; lgR [x, y, xa]\ \ {u. lgR [x, y, u]} = {ya. ya \ x \ lgR [x, y, ya]}" apply(rule_tac Collect_cong, auto simp:lgR_ok) done lemma maxr_lg: "\Suc 0 < x; Suc 0 < y\ \ Maxr lgR [x, y] x = lg x y" apply(auto simp add: lg.simps Maxr.simps) using lgR_ok by blast lemma lg_lemma': "\Suc 0 < x; Suc 0 < y\ \ rec_exec rec_lg [x, y] = lg x y" apply(simp add: maxr_lg lg_maxr) done lemma lg_lemma'': "\ Suc 0 < x \ rec_exec rec_lg [x, y] = lg x y" apply(simp add: rec_exec.simps rec_lg_def Let_def lg.simps) done lemma lg_lemma''': "\ Suc 0 < y \ rec_exec rec_lg [x, y] = lg x y" apply(simp add: rec_exec.simps rec_lg_def Let_def lg.simps) done text \ The correctness of \rec_lg\. \ lemma lg_lemma: "rec_exec rec_lg [x, y] = lg x y" apply(cases "Suc 0 < x \ Suc 0 < y", auto simp: lg_lemma' lg_lemma'' lg_lemma''') done text \ \Entry sr i\ returns the \i\-th entry of a list of natural numbers encoded by number \sr\ using Godel's coding. \ fun Entry :: "nat \ nat \ nat" where "Entry sr i = lo sr (Pi (Suc i))" text \ \rec_entry\ is the recursive function used to implement \Entry\. \ definition rec_entry:: "recf" where "rec_entry = Cn 2 rec_lo [id 2 0, Cn 2 rec_pi [Cn 2 s [id 2 1]]]" declare Pi.simps[simp del] text \ The correctness of \rec_entry\. \ lemma entry_lemma: "rec_exec rec_entry [str, i] = Entry str i" by(simp add: rec_entry_def rec_exec.simps lo_lemma pi_lemma) subsection \The construction of F\ text \ Using the auxilliary functions obtained in last section, we are going to contruct the function \F\, which is an interpreter of Turing Machines. \ fun listsum2 :: "nat list \ nat \ nat" where "listsum2 xs 0 = 0" | "listsum2 xs (Suc n) = listsum2 xs n + xs ! n" fun rec_listsum2 :: "nat \ nat \ recf" where "rec_listsum2 vl 0 = Cn vl z [id vl 0]" | "rec_listsum2 vl (Suc n) = Cn vl rec_add [rec_listsum2 vl n, id vl n]" declare listsum2.simps[simp del] rec_listsum2.simps[simp del] lemma listsum2_lemma: "\length xs = vl; n \ vl\ \ rec_exec (rec_listsum2 vl n) xs = listsum2 xs n" apply(induct n, simp_all) apply(simp_all add: rec_exec.simps rec_listsum2.simps listsum2.simps) done fun strt' :: "nat list \ nat \ nat" where "strt' xs 0 = 0" | "strt' xs (Suc n) = (let dbound = listsum2 xs n + n in strt' xs n + (2^(xs ! n + dbound) - 2^dbound))" fun rec_strt' :: "nat \ nat \ recf" where "rec_strt' vl 0 = Cn vl z [id vl 0]" | "rec_strt' vl (Suc n) = (let rec_dbound = Cn vl rec_add [rec_listsum2 vl n, Cn vl (constn n) [id vl 0]] in Cn vl rec_add [rec_strt' vl n, Cn vl rec_minus [Cn vl rec_power [Cn vl (constn 2) [id vl 0], Cn vl rec_add [id vl (n), rec_dbound]], Cn vl rec_power [Cn vl (constn 2) [id vl 0], rec_dbound]]])" declare strt'.simps[simp del] rec_strt'.simps[simp del] lemma strt'_lemma: "\length xs = vl; n \ vl\ \ rec_exec (rec_strt' vl n) xs = strt' xs n" apply(induct n) apply(simp_all add: rec_exec.simps rec_strt'.simps strt'.simps Let_def power_lemma listsum2_lemma) done text \ \strt\ corresponds to the \strt\ function on page 90 of B book, but this definition generalises the original one to deal with multiple input arguments. \ fun strt :: "nat list \ nat" where "strt xs = (let ys = map Suc xs in strt' ys (length ys))" fun rec_map :: "recf \ nat \ recf list" where "rec_map rf vl = map (\ i. Cn vl rf [id vl i]) [0.. \rec_strt\ is the recursive function used to implement \strt\. \ fun rec_strt :: "nat \ recf" where "rec_strt vl = Cn vl (rec_strt' vl vl) (rec_map s vl)" lemma map_s_lemma: "length xs = vl \ map ((\a. rec_exec a xs) \ (\i. Cn vl s [recf.id vl i])) [0.. ys y. xs = ys @ [y]", auto) proof - fix ys y assume ind: "\xs. length xs = length (ys::nat list) \ map ((\a. rec_exec a xs) \ (\i. Cn (length ys) s [recf.id (length ys) (i)])) [0..a. rec_exec a (ys @ [y])) \ (\i. Cn (Suc (length ys)) s [recf.id (Suc (length ys)) (i)])) [0..a. rec_exec a ys) \ (\i. Cn (length ys) s [recf.id (length ys) (i)])) [0..a. rec_exec a (ys @ [y])) \ (\i. Cn (Suc (length ys)) s [recf.id (Suc (length ys)) (i)])) [0..a. rec_exec a ys) \ (\i. Cn (length ys) s [recf.id (length ys) (i)])) [0..ys y. xs = ys @ [y]" apply(rule_tac x = "butlast xs" in exI, rule_tac x = "last xs" in exI) apply(subgoal_tac "xs \ []", auto) done qed text \ The correctness of \rec_strt\. \ lemma strt_lemma: "length xs = vl \ rec_exec (rec_strt vl) xs = strt xs" apply(simp add: strt.simps rec_exec.simps strt'_lemma) apply(subgoal_tac "(map ((\a. rec_exec a xs) \ (\i. Cn vl s [recf.id vl (i)])) [0.. The \scan\ function on page 90 of B book. \ fun scan :: "nat \ nat" where "scan r = r mod 2" text \ \rec_scan\ is the implemention of \scan\. \ definition rec_scan :: "recf" where "rec_scan = Cn 1 rec_mod [id 1 0, constn 2]" text \ The correctness of \scan\. \ lemma scan_lemma: "rec_exec rec_scan [r] = r mod 2" by(simp add: rec_exec.simps rec_scan_def mod_lemma) fun newleft0 :: "nat list \ nat" where "newleft0 [p, r] = p" definition rec_newleft0 :: "recf" where "rec_newleft0 = id 2 0" fun newrgt0 :: "nat list \ nat" where "newrgt0 [p, r] = r - scan r" definition rec_newrgt0 :: "recf" where "rec_newrgt0 = Cn 2 rec_minus [id 2 1, Cn 2 rec_scan [id 2 1]]" (*newleft1, newrgt1: left rgt number after execute on step*) fun newleft1 :: "nat list \ nat" where "newleft1 [p, r] = p" definition rec_newleft1 :: "recf" where "rec_newleft1 = id 2 0" fun newrgt1 :: "nat list \ nat" where "newrgt1 [p, r] = r + 1 - scan r" definition rec_newrgt1 :: "recf" where "rec_newrgt1 = Cn 2 rec_minus [Cn 2 rec_add [id 2 1, Cn 2 (constn 1) [id 2 0]], Cn 2 rec_scan [id 2 1]]" fun newleft2 :: "nat list \ nat" where "newleft2 [p, r] = p div 2" definition rec_newleft2 :: "recf" where "rec_newleft2 = Cn 2 rec_quo [id 2 0, Cn 2 (constn 2) [id 2 0]]" fun newrgt2 :: "nat list \ nat" where "newrgt2 [p, r] = 2 * r + p mod 2" definition rec_newrgt2 :: "recf" where "rec_newrgt2 = Cn 2 rec_add [Cn 2 rec_mult [Cn 2 (constn 2) [id 2 0], id 2 1], Cn 2 rec_mod [id 2 0, Cn 2 (constn 2) [id 2 0]]]" fun newleft3 :: "nat list \ nat" where "newleft3 [p, r] = 2 * p + r mod 2" definition rec_newleft3 :: "recf" where "rec_newleft3 = Cn 2 rec_add [Cn 2 rec_mult [Cn 2 (constn 2) [id 2 0], id 2 0], Cn 2 rec_mod [id 2 1, Cn 2 (constn 2) [id 2 0]]]" fun newrgt3 :: "nat list \ nat" where "newrgt3 [p, r] = r div 2" definition rec_newrgt3 :: "recf" where "rec_newrgt3 = Cn 2 rec_quo [id 2 1, Cn 2 (constn 2) [id 2 0]]" text \ The \new_left\ function on page 91 of B book. \ fun newleft :: "nat \ nat \ nat \ nat" where "newleft p r a = (if a = 0 \ a = 1 then newleft0 [p, r] else if a = 2 then newleft2 [p, r] else if a = 3 then newleft3 [p, r] else p)" text \ \rec_newleft\ is the recursive function used to implement \newleft\. \ definition rec_newleft :: "recf" where "rec_newleft = (let g0 = Cn 3 rec_newleft0 [id 3 0, id 3 1] in let g1 = Cn 3 rec_newleft2 [id 3 0, id 3 1] in let g2 = Cn 3 rec_newleft3 [id 3 0, id 3 1] in let g3 = id 3 0 in let r0 = Cn 3 rec_disj [Cn 3 rec_eq [id 3 2, Cn 3 (constn 0) [id 3 0]], Cn 3 rec_eq [id 3 2, Cn 3 (constn 1) [id 3 0]]] in let r1 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 2) [id 3 0]] in let r2 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 3) [id 3 0]] in let r3 = Cn 3 rec_less [Cn 3 (constn 3) [id 3 0], id 3 2] in let gs = [g0, g1, g2, g3] in let rs = [r0, r1, r2, r3] in rec_embranch (zip gs rs))" declare newleft.simps[simp del] lemma Suc_Suc_Suc_Suc_induct: "\i < Suc (Suc (Suc (Suc 0))); i = 0 \ P i; i = 1 \ P i; i =2 \ P i; i =3 \ P i\ \ P i" apply(cases i, force) apply(cases "i - 1", force) apply(cases "i - 1 - 1", force) by(cases "i - 1 - 1 - 1", auto simp:numeral) declare quo_lemma2[simp] mod_lemma[simp] text \ The correctness of \rec_newleft\. \ lemma newleft_lemma: "rec_exec rec_newleft [p, r, a] = newleft p r a" proof(simp only: rec_newleft_def Let_def) let ?rgs = "[Cn 3 rec_newleft0 [recf.id 3 0, recf.id 3 1], Cn 3 rec_newleft2 [recf.id 3 0, recf.id 3 1], Cn 3 rec_newleft3 [recf.id 3 0, recf.id 3 1], recf.id 3 0]" let ?rrs = "[Cn 3 rec_disj [Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 0) [recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 1) [recf.id 3 0]]], Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 2) [recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 3) [recf.id 3 0]], Cn 3 rec_less [Cn 3 (constn 3) [recf.id 3 0], recf.id 3 2]]" have k1: "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] = Embranch (zip (map rec_exec ?rgs) (map (\r args. 0 < rec_exec r args) ?rrs)) [p, r, a]" apply(rule_tac embranch_lemma ) apply(auto simp: numeral_3_eq_3 numeral_2_eq_2 rec_newleft0_def rec_newleft1_def rec_newleft2_def rec_newleft3_def)+ apply(cases "a = 0 \ a = 1", rule_tac x = 0 in exI) prefer 2 apply(cases "a = 2", rule_tac x = "Suc 0" in exI) prefer 2 apply(cases "a = 3", rule_tac x = "2" in exI) prefer 2 apply(cases "a > 3", rule_tac x = "3" in exI, auto) apply(auto simp: rec_exec.simps) apply(erule_tac [!] Suc_Suc_Suc_Suc_induct, auto simp: rec_exec.simps) done have k2: "Embranch (zip (map rec_exec ?rgs) (map (\r args. 0 < rec_exec r args) ?rrs)) [p, r, a] = newleft p r a" apply(simp add: Embranch.simps) apply(simp add: rec_exec.simps) apply(auto simp: newleft.simps rec_newleft0_def rec_exec.simps rec_newleft1_def rec_newleft2_def rec_newleft3_def) done from k1 and k2 show "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] = newleft p r a" by simp qed text \ The \newrght\ function is one similar to \newleft\, but used to compute the right number. \ fun newrght :: "nat \ nat \ nat \ nat" where "newrght p r a = (if a = 0 then newrgt0 [p, r] else if a = 1 then newrgt1 [p, r] else if a = 2 then newrgt2 [p, r] else if a = 3 then newrgt3 [p, r] else r)" text \ \rec_newrght\ is the recursive function used to implement \newrgth\. \ definition rec_newrght :: "recf" where "rec_newrght = (let g0 = Cn 3 rec_newrgt0 [id 3 0, id 3 1] in let g1 = Cn 3 rec_newrgt1 [id 3 0, id 3 1] in let g2 = Cn 3 rec_newrgt2 [id 3 0, id 3 1] in let g3 = Cn 3 rec_newrgt3 [id 3 0, id 3 1] in let g4 = id 3 1 in let r0 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 0) [id 3 0]] in let r1 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 1) [id 3 0]] in let r2 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 2) [id 3 0]] in let r3 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 3) [id 3 0]] in let r4 = Cn 3 rec_less [Cn 3 (constn 3) [id 3 0], id 3 2] in let gs = [g0, g1, g2, g3, g4] in let rs = [r0, r1, r2, r3, r4] in rec_embranch (zip gs rs))" declare newrght.simps[simp del] lemma numeral_4_eq_4: "4 = Suc 3" by auto lemma Suc_5_induct: "\i < Suc (Suc (Suc (Suc (Suc 0)))); i = 0 \ P 0; i = 1 \ P 1; i = 2 \ P 2; i = 3 \ P 3; i = 4 \ P 4\ \ P i" apply(cases i, force) apply(cases "i-1", force) apply(cases "i-1-1") using less_2_cases numeral by auto lemma primerec_rec_scan_1[intro]: "primerec rec_scan (Suc 0)" apply(auto simp: rec_scan_def, auto) done text \ The correctness of \rec_newrght\. \ lemma newrght_lemma: "rec_exec rec_newrght [p, r, a] = newrght p r a" proof(simp only: rec_newrght_def Let_def) let ?gs' = "[newrgt0, newrgt1, newrgt2, newrgt3, \ zs. zs ! 1]" let ?r0 = "\ zs. zs ! 2 = 0" let ?r1 = "\ zs. zs ! 2 = 1" let ?r2 = "\ zs. zs ! 2 = 2" let ?r3 = "\ zs. zs ! 2 = 3" let ?r4 = "\ zs. zs ! 2 > 3" let ?gs = "map (\ g. (\ zs. g [zs ! 0, zs ! 1])) ?gs'" let ?rs = "[?r0, ?r1, ?r2, ?r3, ?r4]" let ?rgs = "[Cn 3 rec_newrgt0 [recf.id 3 0, recf.id 3 1], Cn 3 rec_newrgt1 [recf.id 3 0, recf.id 3 1], Cn 3 rec_newrgt2 [recf.id 3 0, recf.id 3 1], Cn 3 rec_newrgt3 [recf.id 3 0, recf.id 3 1], recf.id 3 1]" let ?rrs = "[Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 0) [recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 1) [recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 2) [recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 3) [recf.id 3 0]], Cn 3 rec_less [Cn 3 (constn 3) [recf.id 3 0], recf.id 3 2]]" have k1: "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] = Embranch (zip (map rec_exec ?rgs) (map (\r args. 0 < rec_exec r args) ?rrs)) [p, r, a]" apply(rule_tac embranch_lemma) apply(auto simp: numeral_3_eq_3 numeral_2_eq_2 rec_newrgt0_def rec_newrgt1_def rec_newrgt2_def rec_newrgt3_def)+ apply(cases "a = 0", rule_tac x = 0 in exI) prefer 2 apply(cases "a = 1", rule_tac x = "Suc 0" in exI) prefer 2 apply(cases "a = 2", rule_tac x = "2" in exI) prefer 2 apply(cases "a = 3", rule_tac x = "3" in exI) prefer 2 apply(cases "a > 3", rule_tac x = "4" in exI, auto simp: rec_exec.simps) apply(erule_tac [!] Suc_5_induct, auto simp: rec_exec.simps) done have k2: "Embranch (zip (map rec_exec ?rgs) (map (\r args. 0 < rec_exec r args) ?rrs)) [p, r, a] = newrght p r a" apply(auto simp:Embranch.simps rec_exec.simps) apply(auto simp: newrght.simps rec_newrgt3_def rec_newrgt2_def rec_newrgt1_def rec_newrgt0_def rec_exec.simps scan_lemma) done from k1 and k2 show "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] = newrght p r a" by simp qed declare Entry.simps[simp del] text \ The \actn\ function given on page 92 of B book, which is used to fetch Turing Machine intructions. In \actn m q r\, \m\ is the Godel coding of a Turing Machine, \q\ is the current state of Turing Machine, \r\ is the right number of Turing Machine tape. \ fun actn :: "nat \ nat \ nat \ nat" where "actn m q r = (if q \ 0 then Entry m (4*(q - 1) + 2 * scan r) else 4)" text \ \rec_actn\ is the recursive function used to implement \actn\ \ definition rec_actn :: "recf" where "rec_actn = Cn 3 rec_add [Cn 3 rec_mult [Cn 3 rec_entry [id 3 0, Cn 3 rec_add [Cn 3 rec_mult [Cn 3 (constn 4) [id 3 0], Cn 3 rec_minus [id 3 1, Cn 3 (constn 1) [id 3 0]]], Cn 3 rec_mult [Cn 3 (constn 2) [id 3 0], Cn 3 rec_scan [id 3 2]]]], Cn 3 rec_noteq [id 3 1, Cn 3 (constn 0) [id 3 0]]], Cn 3 rec_mult [Cn 3 (constn 4) [id 3 0], Cn 3 rec_eq [id 3 1, Cn 3 (constn 0) [id 3 0]]]] " text \ The correctness of \actn\. \ lemma actn_lemma: "rec_exec rec_actn [m, q, r] = actn m q r" by(auto simp: rec_actn_def rec_exec.simps entry_lemma scan_lemma) fun newstat :: "nat \ nat \ nat \ nat" where "newstat m q r = (if q \ 0 then Entry m (4*(q - 1) + 2*scan r + 1) else 0)" definition rec_newstat :: "recf" where "rec_newstat = Cn 3 rec_add [Cn 3 rec_mult [Cn 3 rec_entry [id 3 0, Cn 3 rec_add [Cn 3 rec_mult [Cn 3 (constn 4) [id 3 0], Cn 3 rec_minus [id 3 1, Cn 3 (constn 1) [id 3 0]]], Cn 3 rec_add [Cn 3 rec_mult [Cn 3 (constn 2) [id 3 0], Cn 3 rec_scan [id 3 2]], Cn 3 (constn 1) [id 3 0]]]], Cn 3 rec_noteq [id 3 1, Cn 3 (constn 0) [id 3 0]]], Cn 3 rec_mult [Cn 3 (constn 0) [id 3 0], Cn 3 rec_eq [id 3 1, Cn 3 (constn 0) [id 3 0]]]] " lemma newstat_lemma: "rec_exec rec_newstat [m, q, r] = newstat m q r" by(auto simp: rec_exec.simps entry_lemma scan_lemma rec_newstat_def) declare newstat.simps[simp del] actn.simps[simp del] text\code the configuration\ fun trpl :: "nat \ nat \ nat \ nat" where "trpl p q r = (Pi 0)^p * (Pi 1)^q * (Pi 2)^r" definition rec_trpl :: "recf" where "rec_trpl = Cn 3 rec_mult [Cn 3 rec_mult [Cn 3 rec_power [Cn 3 (constn (Pi 0)) [id 3 0], id 3 0], Cn 3 rec_power [Cn 3 (constn (Pi 1)) [id 3 0], id 3 1]], Cn 3 rec_power [Cn 3 (constn (Pi 2)) [id 3 0], id 3 2]]" declare trpl.simps[simp del] lemma trpl_lemma: "rec_exec rec_trpl [p, q, r] = trpl p q r" by(auto simp: rec_trpl_def rec_exec.simps power_lemma trpl.simps) text\left, stat, rght: decode func\ fun left :: "nat \ nat" where "left c = lo c (Pi 0)" fun stat :: "nat \ nat" where "stat c = lo c (Pi 1)" fun rght :: "nat \ nat" where "rght c = lo c (Pi 2)" fun inpt :: "nat \ nat list \ nat" where "inpt m xs = trpl 0 1 (strt xs)" fun newconf :: "nat \ nat \ nat" where "newconf m c = trpl (newleft (left c) (rght c) (actn m (stat c) (rght c))) (newstat m (stat c) (rght c)) (newrght (left c) (rght c) (actn m (stat c) (rght c)))" declare left.simps[simp del] stat.simps[simp del] rght.simps[simp del] inpt.simps[simp del] newconf.simps[simp del] definition rec_left :: "recf" where "rec_left = Cn 1 rec_lo [id 1 0, constn (Pi 0)]" definition rec_right :: "recf" where "rec_right = Cn 1 rec_lo [id 1 0, constn (Pi 2)]" definition rec_stat :: "recf" where "rec_stat = Cn 1 rec_lo [id 1 0, constn (Pi 1)]" definition rec_inpt :: "nat \ recf" where "rec_inpt vl = Cn vl rec_trpl [Cn vl (constn 0) [id vl 0], Cn vl (constn 1) [id vl 0], Cn vl (rec_strt (vl - 1)) (map (\ i. id vl (i)) [1..a. rec_exec a (m # xs)) \ (\i. recf.id (Suc (length xs)) (i))) [Suc 0.. i. xs ! (i - 1)) [Suc 0.. map (\ i. xs ! (i - 1)) [Suc 0.. ys y. xs = ys @ [y]", auto) proof - fix ys y assume ind: "\xs. length (ys::nat list) = length (xs::nat list) \ map (\i. xs ! (i - Suc 0)) [Suc 0.. length (ys::nat list)" have "map (\i. ys ! (i - Suc 0)) [Suc 0..i. (ys @ [y]) ! (i - Suc 0)) [Suc 0..i. ys ! (i - Suc 0)) [Suc 0..i. (ys @ [y]) ! (i - Suc 0)) [Suc 0..ys y. xs = ys @ [y]" apply(rule_tac x = "butlast xs" in exI, rule_tac x = "last xs" in exI) apply(cases "xs \ []", auto) done qed qed simp lemma nonempty_listE: "Suc 0 \ length xs \ (map ((\a. rec_exec a (m # xs)) \ (\i. recf.id (Suc (length xs)) (i))) [Suc 0..Suc (length xs) = vl\ \ rec_exec (rec_inpt vl) (m # xs) = inpt m xs" apply(auto simp: rec_exec.simps rec_inpt_def trpl_lemma inpt.simps strt_lemma) apply(subgoal_tac "(map ((\a. rec_exec a (m # xs)) \ (\i. recf.id (Suc (length xs)) (i))) [Suc 0.. \conf m r k\ computes the TM configuration after \k\ steps of execution of TM coded as \m\ starting from the initial configuration where the left number equals \0\, right number equals \r\. \ fun conf :: "nat \ nat \ nat \ nat" where "conf m r 0 = trpl 0 (Suc 0) r" | "conf m r (Suc t) = newconf m (conf m r t)" declare conf.simps[simp del] text \ \conf\ is implemented by the following recursive function \rec_conf\. \ definition rec_conf :: "recf" where "rec_conf = Pr 2 (Cn 2 rec_trpl [Cn 2 (constn 0) [id 2 0], Cn 2 (constn (Suc 0)) [id 2 0], id 2 1]) (Cn 4 rec_newconf [id 4 0, id 4 3])" lemma conf_step: "rec_exec rec_conf [m, r, Suc t] = rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" proof - have "rec_exec rec_conf ([m, r] @ [Suc t]) = rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" by(simp only: rec_conf_def rec_pr_Suc_simp_rewrite, simp add: rec_exec.simps) thus "rec_exec rec_conf [m, r, Suc t] = rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" by simp qed text \ The correctness of \rec_conf\. \ lemma conf_lemma: "rec_exec rec_conf [m, r, t] = conf m r t" by (induct t) (auto simp add: rec_conf_def rec_exec.simps conf.simps inpt_lemma trpl_lemma) text \ \NSTD c\ returns true if the configureation coded by \c\ is no a stardard final configuration. \ fun NSTD :: "nat \ bool" where "NSTD c = (stat c \ 0 \ left c \ 0 \ rght c \ 2^(lg (rght c + 1) 2) - 1 \ rght c = 0)" text \ \rec_NSTD\ is the recursive function implementing \NSTD\. \ definition rec_NSTD :: "recf" where "rec_NSTD = Cn 1 rec_disj [ Cn 1 rec_disj [ Cn 1 rec_disj [Cn 1 rec_noteq [rec_stat, constn 0], Cn 1 rec_noteq [rec_left, constn 0]] , Cn 1 rec_noteq [rec_right, Cn 1 rec_minus [Cn 1 rec_power [constn 2, Cn 1 rec_lg [Cn 1 rec_add [rec_right, constn 1], constn 2]], constn 1]]], Cn 1 rec_eq [rec_right, constn 0]]" lemma NSTD_lemma1: "rec_exec rec_NSTD [c] = Suc 0 \ rec_exec rec_NSTD [c] = 0" by(simp add: rec_exec.simps rec_NSTD_def) declare NSTD.simps[simp del] lemma NSTD_lemma2': "(rec_exec rec_NSTD [c] = Suc 0) \ NSTD c" apply(simp add: rec_exec.simps rec_NSTD_def stat_lemma left_lemma lg_lemma right_lemma power_lemma NSTD.simps) apply(auto) apply(cases "0 < left c", simp, simp) done lemma NSTD_lemma2'': "NSTD c \ (rec_exec rec_NSTD [c] = Suc 0)" apply(simp add: rec_exec.simps rec_NSTD_def stat_lemma left_lemma lg_lemma right_lemma power_lemma NSTD.simps) apply(auto split: if_splits) done text \ The correctness of \NSTD\. \ lemma NSTD_lemma2: "(rec_exec rec_NSTD [c] = Suc 0) = NSTD c" using NSTD_lemma1 apply(auto intro: NSTD_lemma2' NSTD_lemma2'') done fun nstd :: "nat \ nat" where "nstd c = (if NSTD c then 1 else 0)" lemma nstd_lemma: "rec_exec rec_NSTD [c] = nstd c" using NSTD_lemma1 apply(simp add: NSTD_lemma2, auto) done text\ \nonstep m r t\ means afer \t\ steps of execution, the TM coded by \m\ is not at a stardard final configuration. \ fun nonstop :: "nat \ nat \ nat \ nat" where "nonstop m r t = nstd (conf m r t)" text \ \rec_nonstop\ is the recursive function implementing \nonstop\. \ definition rec_nonstop :: "recf" where "rec_nonstop = Cn 3 rec_NSTD [rec_conf]" text \ The correctness of \rec_nonstop\. \ lemma nonstop_lemma: "rec_exec rec_nonstop [m, r, t] = nonstop m r t" apply(simp add: rec_exec.simps rec_nonstop_def nstd_lemma conf_lemma) done text\ \rec_halt\ is the recursive function calculating the steps a TM needs to execute before to reach a stardard final configuration. This recursive function is the only one using \Mn\ combinator. So it is the only non-primitive recursive function needs to be used in the construction of the universal function \F\. \ definition rec_halt :: "recf" where "rec_halt = Mn (Suc (Suc 0)) (rec_nonstop)" declare nonstop.simps[simp del] text \ The lemma relates the interpreter of primitive functions with the calculation relation of general recursive functions. \ declare numeral_2_eq_2[simp] numeral_3_eq_3[simp] lemma primerec_rec_right_1[intro]: "primerec rec_right (Suc 0)" by(auto simp: rec_right_def rec_lo_def Let_def;force) lemma primerec_rec_pi_helper: "\iiresolve_tac @{context} [@{thm prime_cn}, @{thm prime_pr}] 1\ ;(simp add:primerec_rec_pi_helpers primrec_dummyfac)?)+ by fastforce+ lemma primerec_recs[intro]: "primerec rec_trpl (Suc (Suc (Suc 0)))" "primerec rec_newleft0 (Suc (Suc 0))" "primerec rec_newleft1 (Suc (Suc 0))" "primerec rec_newleft2 (Suc (Suc 0))" "primerec rec_newleft3 (Suc (Suc 0))" "primerec rec_newleft (Suc (Suc (Suc 0)))" "primerec rec_left (Suc 0)" "primerec rec_actn (Suc (Suc (Suc 0)))" "primerec rec_stat (Suc 0)" "primerec rec_newstat (Suc (Suc (Suc 0)))" apply(simp_all add: rec_newleft_def rec_embranch.simps rec_left_def rec_lo_def rec_entry_def rec_actn_def Let_def arity.simps rec_newleft0_def rec_stat_def rec_newstat_def rec_newleft1_def rec_newleft2_def rec_newleft3_def rec_trpl_def) apply(tactic \resolve_tac @{context} [@{thm prime_cn}, @{thm prime_id}, @{thm prime_pr}] 1\;force)+ done lemma primerec_rec_newrght[intro]: "primerec rec_newrght (Suc (Suc (Suc 0)))" apply(simp add: rec_newrght_def rec_embranch.simps Let_def arity.simps rec_newrgt0_def rec_newrgt1_def rec_newrgt2_def rec_newrgt3_def) apply(tactic \resolve_tac @{context} [@{thm prime_cn}, @{thm prime_id}, @{thm prime_pr}] 1\;force)+ done lemma primerec_rec_newconf[intro]: "primerec rec_newconf (Suc (Suc 0))" apply(simp add: rec_newconf_def) by(tactic \resolve_tac @{context} [@{thm prime_cn}, @{thm prime_id}, @{thm prime_pr}] 1\;force) lemma primerec_rec_conf[intro]: "primerec rec_conf (Suc (Suc (Suc 0)))" apply(simp add: rec_conf_def) by(tactic \resolve_tac @{context} [@{thm prime_cn}, @{thm prime_id}, @{thm prime_pr}] 1\;force simp: numeral) lemma primerec_recs2[intro]: "primerec rec_lg (Suc (Suc 0))" "primerec rec_nonstop (Suc (Suc (Suc 0)))" apply(simp_all add: rec_lg_def rec_nonstop_def rec_NSTD_def rec_stat_def rec_lo_def Let_def rec_left_def rec_right_def rec_newconf_def rec_newstat_def) by(tactic \resolve_tac @{context} [@{thm prime_cn}, @{thm prime_id}, @{thm prime_pr}] 1\;fastforce)+ lemma primerec_terminate: "\primerec f x; length xs = x\ \ terminate f xs" proof(induct arbitrary: xs rule: primerec.induct) fix xs assume "length (xs::nat list) = Suc 0" thus "terminate z xs" by(cases xs, auto intro: termi_z) next fix xs assume "length (xs::nat list) = Suc 0" thus "terminate s xs" by(cases xs, auto intro: termi_s) next fix n m xs assume "n < m" "length (xs::nat list) = m" thus "terminate (id m n) xs" by(erule_tac termi_id, simp) next fix f k gs m n xs assume ind: "\i (\x. length x = m \ terminate (gs ! i) x)" and ind2: "\ xs. length xs = k \ terminate f xs" and h: "primerec f k" "length gs = k" "m = n" "length (xs::nat list) = m" have "terminate f (map (\g. rec_exec g xs) gs)" using ind2[of "(map (\g. rec_exec g xs) gs)"] h by simp moreover have "\g\set gs. terminate g xs" using ind h by(auto simp: set_conv_nth) ultimately show "terminate (Cn n f gs) xs" using h by(rule_tac termi_cn, auto) next fix f n g m xs assume ind1: "\xs. length xs = n \ terminate f xs" and ind2: "\xs. length xs = Suc (Suc n) \ terminate g xs" and h: "primerec f n" " primerec g (Suc (Suc n))" " m = Suc n" "length (xs::nat list) = m" have "\y The following lemma gives the correctness of \rec_halt\. It says: if \rec_halt\ calculates that the TM coded by \m\ will reach a standard final configuration after \t\ steps of execution, then it is indeed so. \ text \F: universal machine\ text \ \valu r\ extracts computing result out of the right number \r\. \ fun valu :: "nat \ nat" where "valu r = (lg (r + 1) 2) - 1" text \ \rec_valu\ is the recursive function implementing \valu\. \ definition rec_valu :: "recf" where "rec_valu = Cn 1 rec_minus [Cn 1 rec_lg [s, constn 2], constn 1]" text \ The correctness of \rec_valu\. \ lemma value_lemma: "rec_exec rec_valu [r] = valu r" by(simp add: rec_exec.simps rec_valu_def lg_lemma) lemma primerec_rec_valu_1[intro]: "primerec rec_valu (Suc 0)" unfolding rec_valu_def apply(rule prime_cn[of _ "Suc (Suc 0)"]) by auto auto declare valu.simps[simp del] text \ The definition of the universal function \rec_F\. \ definition rec_F :: "recf" where "rec_F = Cn (Suc (Suc 0)) rec_valu [Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])]]" lemma terminate_halt_lemma: "\rec_exec rec_nonstop ([m, r] @ [t]) = 0; \i \ terminate rec_halt [m, r]" apply(simp add: rec_halt_def) apply(rule termi_mn, auto) by(rule primerec_terminate; auto)+ text \ The correctness of \rec_F\, halt case. \ lemma F_lemma: "rec_exec rec_halt [m, r] = t \ rec_exec rec_F [m, r] = (valu (rght (conf m r t)))" by(simp add: rec_F_def rec_exec.simps value_lemma right_lemma conf_lemma halt_lemma) lemma terminate_F_lemma: "terminate rec_halt [m, r] \ terminate rec_F [m, r]" apply(simp add: rec_F_def) apply(rule termi_cn, auto) apply(rule primerec_terminate, auto) apply(rule termi_cn, auto) apply(rule primerec_terminate, auto) apply(rule termi_cn, auto) apply(rule primerec_terminate, auto) apply(rule termi_id;force) apply(rule termi_id;force) done text \ The correctness of \rec_F\, nonhalt case. \ subsection \Coding function of TMs\ text \ The purpose of this section is to get the coding function of Turing Machine, which is going to be named \code\. \ fun bl2nat :: "cell list \ nat \ nat" where "bl2nat [] n = 0" | "bl2nat (Bk#bl) n = bl2nat bl (Suc n)" | "bl2nat (Oc#bl) n = 2^n + bl2nat bl (Suc n)" fun bl2wc :: "cell list \ nat" where "bl2wc xs = bl2nat xs 0" fun trpl_code :: "config \ nat" where "trpl_code (st, l, r) = trpl (bl2wc l) st (bl2wc r)" declare bl2nat.simps[simp del] bl2wc.simps[simp del] trpl_code.simps[simp del] fun action_map :: "action \ nat" where "action_map W0 = 0" | "action_map W1 = 1" | "action_map L = 2" | "action_map R = 3" | "action_map Nop = 4" fun action_map_iff :: "nat \ action" where "action_map_iff (0::nat) = W0" | "action_map_iff (Suc 0) = W1" | "action_map_iff (Suc (Suc 0)) = L" | "action_map_iff (Suc (Suc (Suc 0))) = R" | "action_map_iff n = Nop" fun block_map :: "cell \ nat" where "block_map Bk = 0" | "block_map Oc = 1" fun godel_code' :: "nat list \ nat \ nat" where "godel_code' [] n = 1" | "godel_code' (x#xs) n = (Pi n)^x * godel_code' xs (Suc n) " fun godel_code :: "nat list \ nat" where "godel_code xs = (let lh = length xs in 2^lh * (godel_code' xs (Suc 0)))" fun modify_tprog :: "instr list \ nat list" where "modify_tprog [] = []" | "modify_tprog ((ac, ns)#nl) = action_map ac # ns # modify_tprog nl" text \ \code tp\ gives the Godel coding of TM program \tp\. \ fun code :: "instr list \ nat" where "code tp = (let nl = modify_tprog tp in godel_code nl)" subsection \Relating interperter functions to the execution of TMs\ lemma bl2wc_0[simp]: "bl2wc [] = 0" by(simp add: bl2wc.simps bl2nat.simps) lemma fetch_action_map_4[simp]: "\fetch tp 0 b = (nact, ns)\ \ action_map nact = 4" apply(simp add: fetch.simps) done lemma Pi_gr_1[simp]: "Pi n > Suc 0" proof(induct n, auto simp: Pi.simps Np.simps) fix n let ?setx = "{y. y \ Suc (Pi n!) \ Pi n < y \ Prime y}" have "finite ?setx" by auto moreover have "?setx \ {}" using prime_ex[of "Pi n"] apply(auto) done ultimately show "Suc 0 < Min ?setx" apply(simp add: Min_gr_iff) apply(auto simp: Prime.simps) done qed lemma Pi_not_0[simp]: "Pi n > 0" using Pi_gr_1[of n] by arith declare godel_code.simps[simp del] lemma godel_code'_nonzero[simp]: "0 < godel_code' nl n" apply(induct nl arbitrary: n) apply(auto simp: godel_code'.simps) done lemma godel_code_great: "godel_code nl > 0" apply(simp add: godel_code.simps) done lemma godel_code_eq_1: "(godel_code nl = 1) = (nl = [])" apply(auto simp: godel_code.simps) done lemma godel_code_1_iff[elim]: "\i < length nl; \ Suc 0 < godel_code nl\ \ nl ! i = 0" using godel_code_great[of nl] godel_code_eq_1[of nl] apply(simp) done lemma prime_coprime: "\Prime x; Prime y; x\y\ \ coprime x y" proof (simp only: Prime.simps coprime_def, auto simp: dvd_def, rule_tac classical, simp) fix d k ka assume case_ka: "\uv d * ka" and case_k: "\uv d * k" and h: "(0::nat) < d" "d \ Suc 0" "Suc 0 < d * ka" "ka \ k" "Suc 0 < d * k" from h have "k > Suc 0 \ ka >Suc 0" by (cases ka;cases k;force+) from this show "False" proof(erule_tac disjE) assume "(Suc 0::nat) < k" hence "k < d*k \ d < d*k" using h by(auto) thus "?thesis" using case_k apply(erule_tac x = d in allE) apply(simp) apply(erule_tac x = k in allE) apply(simp) done next assume "(Suc 0::nat) < ka" hence "ka < d * ka \ d < d*ka" using h by auto thus "?thesis" using case_ka apply(erule_tac x = d in allE) apply(simp) apply(erule_tac x = ka in allE) apply(simp) done qed qed lemma Pi_inc: "Pi (Suc i) > Pi i" proof(simp add: Pi.simps Np.simps) let ?setx = "{y. y \ Suc (Pi i!) \ Pi i < y \ Prime y}" have "finite ?setx" by simp moreover have "?setx \ {}" using prime_ex[of "Pi i"] apply(auto) done ultimately show "Pi i < Min ?setx" apply(simp) done qed lemma Pi_inc_gr: "i < j \ Pi i < Pi j" proof(induct j, simp) fix j assume ind: "i < j \ Pi i < Pi j" and h: "i < Suc j" from h show "Pi i < Pi (Suc j)" proof(cases "i < j") case True thus "?thesis" proof - assume "i < j" hence "Pi i < Pi j" by(erule_tac ind) moreover have "Pi j < Pi (Suc j)" apply(simp add: Pi_inc) done ultimately show "?thesis" by simp qed next assume "i < Suc j" "\ i < j" hence "i = j" by arith thus "Pi i < Pi (Suc j)" apply(simp add: Pi_inc) done qed qed lemma Pi_notEq: "i \ j \ Pi i \ Pi j" apply(cases "i < j") using Pi_inc_gr[of i j] apply(simp) using Pi_inc_gr[of j i] apply(simp) done lemma prime_2[intro]: "Prime (Suc (Suc 0))" apply(auto simp: Prime.simps) using less_2_cases by fastforce lemma Prime_Pi[intro]: "Prime (Pi n)" proof(induct n, auto simp: Pi.simps Np.simps) fix n let ?setx = "{y. y \ Suc (Pi n!) \ Pi n < y \ Prime y}" show "Prime (Min ?setx)" proof - have "finite ?setx" by simp moreover have "?setx \ {}" using prime_ex[of "Pi n"] apply(simp) done ultimately show "?thesis" apply(drule_tac Min_in, simp, simp) done qed qed lemma Pi_coprime: "i \ j \ coprime (Pi i) (Pi j)" using Prime_Pi[of i] using Prime_Pi[of j] apply(rule_tac prime_coprime, simp_all add: Pi_notEq) done lemma Pi_power_coprime: "i \ j \ coprime ((Pi i)^m) ((Pi j)^n)" unfolding coprime_power_right_iff coprime_power_left_iff using Pi_coprime by auto lemma coprime_dvd_mult_nat2: "\coprime (k::nat) n; k dvd n * m\ \ k dvd m" unfolding coprime_dvd_mult_right_iff. declare godel_code'.simps[simp del] lemma godel_code'_butlast_last_id' : "godel_code' (ys @ [y]) (Suc j) = godel_code' ys (Suc j) * Pi (Suc (length ys + j)) ^ y" proof(induct ys arbitrary: j, simp_all add: godel_code'.simps) qed lemma godel_code'_butlast_last_id: "xs \ [] \ godel_code' xs (Suc j) = godel_code' (butlast xs) (Suc j) * Pi (length xs + j)^(last xs)" apply(subgoal_tac "\ ys y. xs = ys @ [y]") apply(erule_tac exE, erule_tac exE, simp add: godel_code'_butlast_last_id') apply(rule_tac x = "butlast xs" in exI) apply(rule_tac x = "last xs" in exI, auto) done lemma godel_code'_not0: "godel_code' xs n \ 0" apply(induct xs, auto simp: godel_code'.simps) done lemma godel_code_append_cons: "length xs = i \ godel_code' (xs@y#ys) (Suc 0) = godel_code' xs (Suc 0) * Pi (Suc i)^y * godel_code' ys (i + 2)" proof(induct "length xs" arbitrary: i y ys xs, simp add: godel_code'.simps,simp) fix x xs i y ys assume ind: "\xs i y ys. \x = i; length xs = i\ \ godel_code' (xs @ y # ys) (Suc 0) = godel_code' xs (Suc 0) * Pi (Suc i) ^ y * godel_code' ys (Suc (Suc i))" and h: "Suc x = i" "length (xs::nat list) = i" have "godel_code' (butlast xs @ last xs # ((y::nat)#ys)) (Suc 0) = godel_code' (butlast xs) (Suc 0) * Pi (Suc (i - 1))^(last xs) * godel_code' (y#ys) (Suc (Suc (i - 1)))" apply(rule_tac ind) using h by(auto) moreover have "godel_code' xs (Suc 0)= godel_code' (butlast xs) (Suc 0) * Pi (i)^(last xs)" using godel_code'_butlast_last_id[of xs] h apply(cases "xs = []", simp, simp) done moreover have "butlast xs @ last xs # y # ys = xs @ y # ys" using h apply(cases xs, auto) done ultimately show "godel_code' (xs @ y # ys) (Suc 0) = godel_code' xs (Suc 0) * Pi (Suc i) ^ y * godel_code' ys (Suc (Suc i))" using h apply(simp add: godel_code'_not0 Pi_not_0) apply(simp add: godel_code'.simps) done qed lemma Pi_coprime_pre: "length ps \ i \ coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" proof(induct "length ps" arbitrary: ps) fix x ps assume ind: "\ps. \x = length ps; length ps \ i\ \ coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" and h: "Suc x = length ps" "length (ps::nat list) \ i" have g: "coprime (Pi (Suc i)) (godel_code' (butlast ps) (Suc 0))" apply(rule_tac ind) using h by auto have k: "godel_code' ps (Suc 0) = godel_code' (butlast ps) (Suc 0) * Pi (length ps)^(last ps)" using godel_code'_butlast_last_id[of ps 0] h by(cases ps, simp, simp) from g have "coprime (Pi (Suc i)) (Pi (length ps) ^ last ps)" unfolding coprime_power_right_iff using Pi_coprime h(2) by auto with g have "coprime (Pi (Suc i)) (godel_code' (butlast ps) (Suc 0) * Pi (length ps)^(last ps)) " unfolding coprime_mult_right_iff coprime_power_right_iff by auto from this and k show "coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" by simp qed (auto simp add: godel_code'.simps) lemma Pi_coprime_suf: "i < j \ coprime (Pi i) (godel_code' ps j)" proof(induct "length ps" arbitrary: ps) fix x ps assume ind: "\ps. \x = length ps; i < j\ \ coprime (Pi i) (godel_code' ps j)" and h: "Suc x = length (ps::nat list)" "i < j" have g: "coprime (Pi i) (godel_code' (butlast ps) j)" apply(rule ind) using h by auto have k: "(godel_code' ps j) = godel_code' (butlast ps) j * Pi (length ps + j - 1)^last ps" using h godel_code'_butlast_last_id[of ps "j - 1"] apply(cases "ps = []", simp, simp) done from g have "coprime (Pi i) (godel_code' (butlast ps) j * Pi (length ps + j - 1)^last ps)" using Pi_power_coprime[of i "length ps + j - 1" 1 "last ps"] h by(auto) from k and this show "coprime (Pi i) (godel_code' ps j)" by auto qed (simp add: godel_code'.simps) lemma godel_finite: "finite {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" proof(rule bounded_nat_set_is_finite[of _ "godel_code' nl (Suc 0)",rule_format],goal_cases) case (1 ia) then show ?case proof(cases "ia < godel_code' nl (Suc 0)") case False hence g1: "Pi (Suc i) ^ ia dvd godel_code' nl (Suc 0)" and g2: "\ ia < godel_code' nl (Suc 0)" and "Pi (Suc i)^ia \ godel_code' nl (Suc 0)" using godel_code'_not0[of nl "Suc 0"] using 1 by (auto elim:dvd_imp_le) moreover have "ia < Pi (Suc i)^ia" by(rule x_less_exp[OF Pi_gr_1]) ultimately show ?thesis using g2 by(auto) qed auto qed lemma godel_code_in: "i < length nl \ nl ! i \ {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" proof - assume h: "i {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" by(simp) qed lemma godel_code'_get_nth: "i < length nl \ Max {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)} = nl ! i" proof(rule_tac Max_eqI) let ?gc = "godel_code' nl (Suc 0)" assume h: "i < length nl" thus "finite {u. Pi (Suc i) ^ u dvd ?gc}" by (simp add: godel_finite) next fix y let ?suf ="godel_code' (drop (Suc i) nl) (i + 2)" let ?pref = "godel_code' (take i nl) (Suc 0)" assume h: "i < length nl" "y \ {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" moreover hence "godel_code' (take i nl@(nl!i)#drop (Suc i) nl) (Suc 0) = ?pref * Pi (Suc i)^(nl!i) * ?suf" by(rule_tac godel_code_append_cons, simp) moreover from h have "take i nl @ (nl!i) # drop (Suc i) nl = nl" using upd_conv_take_nth_drop[of i nl "nl!i"] by simp ultimately show "y\nl!i" proof(simp) let ?suf' = "godel_code' (drop (Suc i) nl) (Suc (Suc i))" assume mult_dvd: "Pi (Suc i) ^ y dvd ?pref * Pi (Suc i) ^ nl ! i * ?suf'" hence "Pi (Suc i) ^ y dvd ?pref * Pi (Suc i) ^ nl ! i" proof - have "coprime (Pi (Suc i)^y) ?suf'" by (simp add: Pi_coprime_suf) thus ?thesis using coprime_dvd_mult_left_iff mult_dvd by blast qed hence "Pi (Suc i) ^ y dvd Pi (Suc i) ^ nl ! i" proof(rule_tac coprime_dvd_mult_nat2) have "coprime (Pi (Suc i)^y) (?pref^Suc 0)" using Pi_coprime_pre by simp thus "coprime (Pi (Suc i) ^ y) ?pref" by simp qed hence "Pi (Suc i) ^ y \ Pi (Suc i) ^ nl ! i " apply(rule_tac dvd_imp_le, auto) done thus "y \ nl ! i" apply(rule_tac power_le_imp_le_exp, auto) done qed next assume h: "i {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" by(rule_tac godel_code_in, simp) qed lemma godel_code'_set[simp]: "{u. Pi (Suc i) ^ u dvd (Suc (Suc 0)) ^ length nl * godel_code' nl (Suc 0)} = {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" apply(rule_tac Collect_cong, auto) apply(rule_tac n = " (Suc (Suc 0)) ^ length nl" in coprime_dvd_mult_nat2) proof - have "Pi 0 = (2::nat)" by(simp add: Pi.simps) show "coprime (Pi (Suc i) ^ u) ((Suc (Suc 0)) ^ length nl)" for u using Pi_coprime Pi.simps(1) by force qed lemma godel_code_get_nth: "i < length nl \ Max {u. Pi (Suc i) ^ u dvd godel_code nl} = nl ! i" by(simp add: godel_code.simps godel_code'_get_nth) lemma mod_dvd_simp: "(x mod y = (0::nat)) = (y dvd x)" by(simp add: dvd_def, auto) lemma dvd_power_le: "\a > Suc 0; a ^ y dvd a ^ l\ \ y \ l" apply(cases "y \ l", simp, simp) apply(subgoal_tac "\ d. y = l + d", auto simp: power_add) apply(rule_tac x = "y - l" in exI, simp) done lemma Pi_nonzeroE[elim]: "Pi n = 0 \ RR" using Pi_not_0[of n] by simp lemma Pi_not_oneE[elim]: "Pi n = Suc 0 \ RR" using Pi_gr_1[of n] by simp lemma finite_power_dvd: "\(a::nat) > Suc 0; y \ 0\ \ finite {u. a^u dvd y}" apply(auto simp: dvd_def simp:gr0_conv_Suc intro!:bounded_nat_set_is_finite[of _ y]) by (metis le_less_trans mod_less mod_mult_self1_is_0 not_le Suc_lessD less_trans_Suc mult.right_neutral n_less_n_mult_m x_less_exp zero_less_Suc zero_less_mult_pos) lemma conf_decode1: "\m \ n; m \ k; k \ n\ \ Max {u. Pi m ^ u dvd Pi m ^ l * Pi n ^ st * Pi k ^ r} = l" proof - let ?setx = "{u. Pi m ^ u dvd Pi m ^ l * Pi n ^ st * Pi k ^ r}" assume g: "m \ n" "m \ k" "k \ n" show "Max ?setx = l" proof(rule_tac Max_eqI) show "finite ?setx" apply(rule_tac finite_power_dvd, auto) done next fix y assume h: "y \ ?setx" have "Pi m ^ y dvd Pi m ^ l" proof - have "Pi m ^ y dvd Pi m ^ l * Pi n ^ st" using h g Pi_power_coprime by (simp add: coprime_dvd_mult_left_iff) thus "Pi m^y dvd Pi m^l" using g Pi_power_coprime coprime_dvd_mult_left_iff by blast qed thus "y \ (l::nat)" apply(rule_tac a = "Pi m" in power_le_imp_le_exp) apply(simp_all) apply(rule_tac dvd_power_le, auto) done next show "l \ ?setx" by simp qed qed lemma left_trpl_fst[simp]: "left (trpl l st r) = l" apply(simp add: left.simps trpl.simps lo.simps loR.simps mod_dvd_simp) apply(auto simp: conf_decode1) apply(cases "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r") apply(auto) apply(erule_tac x = l in allE, auto) done lemma stat_trpl_snd[simp]: "stat (trpl l st r) = st" apply(simp add: stat.simps trpl.simps lo.simps loR.simps mod_dvd_simp, auto) apply(subgoal_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r = Pi (Suc 0)^st * Pi 0 ^ l * Pi (Suc (Suc 0)) ^ r") apply(simp (no_asm_simp) add: conf_decode1, simp) apply(cases "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r", auto) apply(erule_tac x = st in allE, auto) done lemma rght_trpl_trd[simp]: "rght (trpl l st r) = r" apply(simp add: rght.simps trpl.simps lo.simps loR.simps mod_dvd_simp, auto) apply(subgoal_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r = Pi (Suc (Suc 0))^r * Pi 0 ^ l * Pi (Suc 0) ^ st") apply(simp (no_asm_simp) add: conf_decode1, simp) apply(cases "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r", auto) apply(erule_tac x = r in allE, auto) done lemma max_lor: "i < length nl \ Max {u. loR [godel_code nl, Pi (Suc i), u]} = nl ! i" apply(simp add: loR.simps godel_code_get_nth mod_dvd_simp) done lemma godel_decode: "i < length nl \ Entry (godel_code nl) i = nl ! i" apply(auto simp: Entry.simps lo.simps max_lor) apply(erule_tac x = "nl!i" in allE) using max_lor[of i nl] godel_finite[of i nl] apply(simp) apply(drule_tac Max_in, auto simp: loR.simps godel_code.simps mod_dvd_simp) using godel_code_in[of i nl] apply(simp) done lemma Four_Suc: "4 = Suc (Suc (Suc (Suc 0)))" by auto declare numeral_2_eq_2[simp del] lemma modify_tprog_fetch_even: "\st \ length tp div 2; st > 0\ \ modify_tprog tp ! (4 * (st - Suc 0) ) = action_map (fst (tp ! (2 * (st - Suc 0))))" proof(induct st arbitrary: tp, simp) fix tp st assume ind: "\tp. \st \ length tp div 2; 0 < st\ \ modify_tprog tp ! (4 * (st - Suc 0)) = action_map (fst ((tp::instr list) ! (2 * (st - Suc 0))))" and h: "Suc st \ length (tp::instr list) div 2" "0 < Suc st" thus "modify_tprog tp ! (4 * (Suc st - Suc 0)) = action_map (fst (tp ! (2 * (Suc st - Suc 0))))" proof(cases "st = 0") case True thus "?thesis" using h by(cases tp, auto) next case False assume g: "st \ 0" hence "\ aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" using h by(cases tp; cases "tl tp", auto) from this obtain aa ab ba bb tp' where g1: "tp = (aa, ab) # (ba, bb) # tp'" by blast hence g2: "modify_tprog tp' ! (4 * (st - Suc 0)) = action_map (fst ((tp'::instr list) ! (2 * (st - Suc 0))))" using h g by (auto intro:ind) thus "?thesis" using g1 g by(cases st, auto simp add: Four_Suc) qed qed lemma modify_tprog_fetch_odd: "\st \ length tp div 2; st > 0\ \ modify_tprog tp ! (Suc (Suc (4 * (st - Suc 0)))) = action_map (fst (tp ! (Suc (2 * (st - Suc 0)))))" proof(induct st arbitrary: tp, simp) fix tp st assume ind: "\tp. \st \ length tp div 2; 0 < st\ \ modify_tprog tp ! Suc (Suc (4 * (st - Suc 0))) = action_map (fst (tp ! Suc (2 * (st - Suc 0))))" and h: "Suc st \ length (tp::instr list) div 2" "0 < Suc st" thus "modify_tprog tp ! Suc (Suc (4 * (Suc st - Suc 0))) = action_map (fst (tp ! Suc (2 * (Suc st - Suc 0))))" proof(cases "st = 0") case True thus "?thesis" using h apply(cases tp, force) by(cases "tl tp", auto) next case False assume g: "st \ 0" hence "\ aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" using h apply(cases tp, simp, cases "tl tp", simp, simp) done from this obtain aa ab ba bb tp' where g1: "tp = (aa, ab) # (ba, bb) # tp'" by blast hence g2: "modify_tprog tp' ! Suc (Suc (4 * (st - Suc 0))) = action_map (fst (tp' ! Suc (2 * (st - Suc 0))))" apply(rule_tac ind) using h g by auto thus "?thesis" using g1 g apply(cases st, simp, simp add: Four_Suc) done qed qed lemma modify_tprog_fetch_action: "\st \ length tp div 2; st > 0; b = 1 \ b = 0\ \ modify_tprog tp ! (4 * (st - Suc 0) + 2* b) = action_map (fst (tp ! ((2 * (st - Suc 0)) + b)))" apply(erule_tac disjE, auto elim: modify_tprog_fetch_odd modify_tprog_fetch_even) done lemma length_modify: "length (modify_tprog tp) = 2 * length tp" apply(induct tp, auto) done declare fetch.simps[simp del] lemma fetch_action_eq: "\block_map b = scan r; fetch tp st b = (nact, ns); st \ length tp div 2\ \ actn (code tp) st r = action_map nact" proof(simp add: actn.simps, auto) let ?i = "4 * (st - Suc 0) + 2 * (r mod 2)" assume h: "block_map b = r mod 2" "fetch tp st b = (nact, ns)" "st \ length tp div 2" "0 < st" have "?i < length (modify_tprog tp)" proof - have "length (modify_tprog tp) = 2 * length tp" by(simp add: length_modify) thus "?thesis" using h by(auto) qed hence "Entry (godel_code (modify_tprog tp))?i = (modify_tprog tp) ! ?i" by(erule_tac godel_decode) moreover have "modify_tprog tp ! ?i = action_map (fst (tp ! (2 * (st - Suc 0) + r mod 2)))" apply(rule_tac modify_tprog_fetch_action) using h by(auto) moreover have "(fst (tp ! (2 * (st - Suc 0) + r mod 2))) = nact" using h apply(cases st, simp_all add: fetch.simps nth_of.simps) apply(cases b, auto simp: block_map.simps nth_of.simps fetch.simps split: if_splits) apply(cases "r mod 2", simp, simp) done ultimately show "Entry (godel_code (modify_tprog tp)) (4 * (st - Suc 0) + 2 * (r mod 2)) = action_map nact" by simp qed lemma fetch_zero_zero[simp]: "fetch tp 0 b = (nact, ns) \ ns = 0" by(simp add: fetch.simps) lemma modify_tprog_fetch_state: "\st \ length tp div 2; st > 0; b = 1 \ b = 0\ \ modify_tprog tp ! Suc (4 * (st - Suc 0) + 2 * b) = (snd (tp ! (2 * (st - Suc 0) + b)))" proof(induct st arbitrary: tp, simp) fix st tp assume ind: "\tp. \st \ length tp div 2; 0 < st; b = 1 \ b = 0\ \ modify_tprog tp ! Suc (4 * (st - Suc 0) + 2 * b) = snd (tp ! (2 * (st - Suc 0) + b))" and h: "Suc st \ length (tp::instr list) div 2" "0 < Suc st" "b = 1 \ b = 0" show "modify_tprog tp ! Suc (4 * (Suc st - Suc 0) + 2 * b) = snd (tp ! (2 * (Suc st - Suc 0) + b))" proof(cases "st = 0") case True thus "?thesis" using h apply(cases tp, force) apply(cases "tl tp", auto) done next case False assume g: "st \ 0" hence "\ aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" using h by(cases tp, force, cases "tl tp", auto) from this obtain aa ab ba bb tp' where g1: "tp = (aa, ab) # (ba, bb) # tp'" by blast hence g2: "modify_tprog tp' ! Suc (4 * (st - Suc 0) + 2 * b) = snd (tp' ! (2 * (st - Suc 0) + b))" apply(intro ind) using h g by auto thus "?thesis" using g1 g by(cases st;force) qed qed lemma fetch_state_eq: "\block_map b = scan r; fetch tp st b = (nact, ns); st \ length tp div 2\ \ newstat (code tp) st r = ns" proof(simp add: newstat.simps, auto) let ?i = "Suc (4 * (st - Suc 0) + 2 * (r mod 2))" assume h: "block_map b = r mod 2" "fetch tp st b = (nact, ns)" "st \ length tp div 2" "0 < st" have "?i < length (modify_tprog tp)" proof - have "length (modify_tprog tp) = 2 * length tp" by(simp add: length_modify) thus "?thesis" using h by(auto) qed hence "Entry (godel_code (modify_tprog tp)) (?i) = (modify_tprog tp) ! ?i" by(erule_tac godel_decode) moreover have "modify_tprog tp ! ?i = (snd (tp ! (2 * (st - Suc 0) + r mod 2)))" apply(rule_tac modify_tprog_fetch_state) using h by(auto) moreover have "(snd (tp ! (2 * (st - Suc 0) + r mod 2))) = ns" using h apply(cases st, simp) apply(cases b, auto simp: fetch.simps split: if_splits) apply(cases "(2 * (st - r mod 2) + r mod 2) = (2 * (st - 1) + r mod 2)";auto) by (metis diff_Suc_Suc diff_zero prod.sel(2)) ultimately show "Entry (godel_code (modify_tprog tp)) (?i) = ns" by simp qed lemma tpl_eqI[intro!]: "\a = a'; b = b'; c = c'\ \ trpl a b c = trpl a' b' c'" by simp lemma bl2nat_double: "bl2nat xs (Suc n) = 2 * bl2nat xs n" proof(induct xs arbitrary: n) case Nil thus "?case" by(simp add: bl2nat.simps) next case (Cons x xs) thus "?case" proof - assume ind: "\n. bl2nat xs (Suc n) = 2 * bl2nat xs n " show "bl2nat (x # xs) (Suc n) = 2 * bl2nat (x # xs) n" proof(cases x) case Bk thus "?thesis" apply(simp add: bl2nat.simps) using ind[of "Suc n"] by simp next case Oc thus "?thesis" apply(simp add: bl2nat.simps) using ind[of "Suc n"] by simp qed qed qed lemma bl2wc_simps[simp]: "bl2wc (Oc # tl c) = Suc (bl2wc c) - bl2wc c mod 2 " "bl2wc (Bk # c) = 2*bl2wc (c)" "2 * bl2wc (tl c) = bl2wc c - bl2wc c mod 2 " "bl2wc [Oc] = Suc 0" "c \ [] \ bl2wc (tl c) = bl2wc c div 2" "c \ [] \ bl2wc [hd c] = bl2wc c mod 2" "c \ [] \ bl2wc (hd c # d) = 2 * bl2wc d + bl2wc c mod 2" "2 * (bl2wc c div 2) = bl2wc c - bl2wc c mod 2" "bl2wc (Oc # list) mod 2 = Suc 0" by(cases c;cases "hd c";force simp: bl2wc.simps bl2nat.simps bl2nat_double)+ declare code.simps[simp del] declare nth_of.simps[simp del] text \ The lemma relates the one step execution of TMs with the interpreter function \rec_newconf\. \ lemma rec_t_eq_step: "(\ (s, l, r). s \ length tp div 2) c \ trpl_code (step0 c tp) = rec_exec rec_newconf [code tp, trpl_code c]" proof(cases c) case (fields s l r) assume "case c of (s, l, r) \ s \ length tp div 2" with fields have "s \ length tp div 2" by auto thus ?thesis unfolding fields proof(cases "fetch tp s (read r)", simp add: newconf.simps trpl_code.simps step.simps) fix a b ca aa ba assume h: "(a::nat) \ length tp div 2" "fetch tp a (read ca) = (aa, ba)" moreover hence "actn (code tp) a (bl2wc ca) = action_map aa" apply(rule_tac b = "read ca" in fetch_action_eq, auto) apply(cases "hd ca";cases ca;force) done moreover from h have "(newstat (code tp) a (bl2wc ca)) = ba" apply(rule_tac b = "read ca" in fetch_state_eq, auto split: list.splits) apply(cases "hd ca";cases ca;force) done ultimately show "trpl_code (ba, update aa (b, ca)) = trpl (newleft (bl2wc b) (bl2wc ca) (actn (code tp) a (bl2wc ca))) (newstat (code tp) a (bl2wc ca)) (newrght (bl2wc b) (bl2wc ca) (actn (code tp) a (bl2wc ca)))" apply(cases aa) apply(auto simp: trpl_code.simps newleft.simps newrght.simps split: action.splits) done qed qed lemma bl2nat_simps[simp]: "bl2nat (Oc # Oc\x) 0 = (2 * 2 ^ x - Suc 0)" "bl2nat (Bk\x) n = 0" by(induct x;force simp: bl2nat.simps bl2nat_double exp_ind)+ lemma bl2nat_exp_zero[simp]: "bl2nat (Oc\y) 0 = 2^y - Suc 0" proof(induct y) case (Suc y) then show ?case by(cases "(2::nat)^y", auto) qed (auto simp: bl2nat.simps bl2nat_double) lemma bl2nat_cons_bk: "bl2nat (ks @ [Bk]) 0 = bl2nat ks 0" proof(induct ks) case (Cons a ks) then show ?case by (cases a, auto simp: bl2nat.simps bl2nat_double) qed (auto simp: bl2nat.simps) lemma bl2nat_cons_oc: "bl2nat (ks @ [Oc]) 0 = bl2nat ks 0 + 2 ^ length ks" proof(induct ks) case (Cons a ks) then show ?case by(cases a, auto simp: bl2nat.simps bl2nat_double) qed (auto simp: bl2nat.simps) lemma bl2nat_append: "bl2nat (xs @ ys) 0 = bl2nat xs 0 + bl2nat ys (length xs) " proof(induct "length xs" arbitrary: xs ys, simp add: bl2nat.simps) fix x xs ys assume ind: "\xs ys. x = length xs \ bl2nat (xs @ ys) 0 = bl2nat xs 0 + bl2nat ys (length xs)" and h: "Suc x = length (xs::cell list)" have "\ ks k. xs = ks @ [k]" apply(rule_tac x = "butlast xs" in exI, rule_tac x = "last xs" in exI) using h apply(cases xs, auto) done from this obtain ks k where "xs = ks @ [k]" by blast moreover hence "bl2nat (ks @ (k # ys)) 0 = bl2nat ks 0 + bl2nat (k # ys) (length ks)" apply(rule_tac ind) using h by simp ultimately show "bl2nat (xs @ ys) 0 = bl2nat xs 0 + bl2nat ys (length xs)" apply(cases k, simp_all add: bl2nat.simps) apply(simp_all only: bl2nat_cons_bk bl2nat_cons_oc) done qed lemma trpl_code_simp[simp]: "trpl_code (steps0 (Suc 0, Bk\l, ) tp 0) = rec_exec rec_conf [code tp, bl2wc (), 0]" apply(simp add: steps.simps rec_exec.simps conf_lemma conf.simps inpt.simps trpl_code.simps bl2wc.simps) done text \ The following lemma relates the multi-step interpreter function \rec_conf\ with the multi-step execution of TMs. \ lemma state_in_range_step : "\a \ length A div 2; step0 (a, b, c) A = (st, l, r); tm_wf (A,0)\ \ st \ length A div 2" apply(simp add: step.simps fetch.simps tm_wf.simps split: if_splits list.splits) apply(case_tac [!] a, auto simp: list_all_length fetch.simps nth_of.simps) apply(erule_tac x = "A ! (2*nat) " in ballE, auto) apply(cases "hd c", auto simp: fetch.simps nth_of.simps) apply(erule_tac x = "A !(2 * nat)" in ballE, auto) apply(erule_tac x = "A !Suc (2 * nat)" in ballE, auto) done lemma state_in_range: "\steps0 (Suc 0, tp) A stp = (st, l, r); tm_wf (A, 0)\ \ st \ length A div 2" proof(induct stp arbitrary: st l r) case (Suc stp st l r) from Suc.prems show ?case proof(simp add: step_red, cases "(steps0 (Suc 0, tp) A stp)", simp) fix a b c assume h3: "step0 (a, b, c) A = (st, l, r)" and h4: "steps0 (Suc 0, tp) A stp = (a, b, c)" have "a \ length A div 2" using Suc.prems h4 by (auto intro: Suc.hyps) thus "?thesis" using h3 Suc.prems by (auto elim: state_in_range_step) qed qed(auto simp: tm_wf.simps steps.simps) lemma rec_t_eq_steps: "tm_wf (tp,0) \ trpl_code (steps0 (Suc 0, Bk\l, ) tp stp) = rec_exec rec_conf [code tp, bl2wc (), stp]" proof(induct stp) case 0 thus "?case" by(simp) next case (Suc n) thus "?case" proof - assume ind: "tm_wf (tp,0) \ trpl_code (steps0 (Suc 0, Bk\ l, ) tp n) = rec_exec rec_conf [code tp, bl2wc (), n]" and h: "tm_wf (tp, 0)" show "trpl_code (steps0 (Suc 0, Bk\ l, ) tp (Suc n)) = rec_exec rec_conf [code tp, bl2wc (), Suc n]" proof(cases "steps0 (Suc 0, Bk\ l, ) tp n", simp only: step_red conf_lemma conf.simps) fix a b c assume g: "steps0 (Suc 0, Bk\ l, ) tp n = (a, b, c) " hence "conf (code tp) (bl2wc ()) n= trpl_code (a, b, c)" using ind h apply(simp add: conf_lemma) done moreover hence "trpl_code (step0 (a, b, c) tp) = rec_exec rec_newconf [code tp, trpl_code (a, b, c)]" apply(rule_tac rec_t_eq_step) using h g apply(simp add: state_in_range) done ultimately show "trpl_code (step0 (a, b, c) tp) = newconf (code tp) (conf (code tp) (bl2wc ()) n)" by(simp) qed qed qed lemma bl2wc_Bk_0[simp]: "bl2wc (Bk\ m) = 0" apply(induct m) apply(simp, simp) done lemma bl2wc_Oc_then_Bk[simp]: "bl2wc (Oc\ rs@Bk\ n) = bl2wc (Oc\ rs)" apply(induct rs, simp, simp add: bl2wc.simps bl2nat.simps bl2nat_double) done lemma lg_power: "x > Suc 0 \ lg (x ^ rs) x = rs" proof(simp add: lg.simps, auto) fix xa assume h: "Suc 0 < x" show "Max {ya. ya \ x ^ rs \ lgR [x ^ rs, x, ya]} = rs" apply(rule_tac Max_eqI, simp_all add: lgR.simps) apply(simp add: h) using x_less_exp[of x rs] h apply(simp) done next assume "\ Suc 0 < x ^ rs" "Suc 0 < x" thus "rs = 0" apply(cases "x ^ rs", simp, simp) done next assume "Suc 0 < x" "\xa. \ lgR [x ^ rs, x, xa]" thus "rs = 0" apply(simp only:lgR.simps) apply(erule_tac x = rs in allE, simp) done qed text \ The following lemma relates execution of TMs with the multi-step interpreter function \rec_nonstop\. Note, \rec_nonstop\ is constructed using \rec_conf\. \ declare tm_wf.simps[simp del] lemma nonstop_t_eq: "\steps0 (Suc 0, Bk\l, ) tp stp = (0, Bk\ m, Oc\ rs @ Bk\ n); tm_wf (tp, 0); rs > 0\ \ rec_exec rec_nonstop [code tp, bl2wc (), stp] = 0" proof(simp add: nonstop_lemma nonstop.simps ) assume h: "steps0 (Suc 0, Bk\l, ) tp stp = (0, Bk\ m, Oc\ rs @ Bk\ n)" and tc_t: "tm_wf (tp, 0)" "rs > 0" have g: "rec_exec rec_conf [code tp, bl2wc (), stp] = trpl_code (0, Bk\ m, Oc\ rs@Bk\ n)" using rec_t_eq_steps[of tp l lm stp] tc_t h by(simp) thus "\ NSTD (conf (code tp) (bl2wc ()) stp)" proof(auto simp: NSTD.simps) show "stat (conf (code tp) (bl2wc ()) stp) = 0" using g by(auto simp: conf_lemma trpl_code.simps) next show "left (conf (code tp) (bl2wc ()) stp) = 0" using g by(simp add: conf_lemma trpl_code.simps) next show "rght (conf (code tp) (bl2wc ()) stp) = 2 ^ lg (Suc (rght (conf (code tp) (bl2wc ()) stp))) 2 - Suc 0" using g h proof(simp add: conf_lemma trpl_code.simps) have "2 ^ lg (Suc (bl2wc (Oc\ rs))) 2 = Suc (bl2wc (Oc\ rs))" apply(simp add: bl2wc.simps lg_power) done thus "bl2wc (Oc\ rs) = 2 ^ lg (Suc (bl2wc (Oc\ rs))) 2 - Suc 0" apply(simp) done qed next show "0 < rght (conf (code tp) (bl2wc ()) stp)" using g h tc_t apply(simp add: conf_lemma trpl_code.simps bl2wc.simps bl2nat.simps) apply(cases rs, simp, simp add: bl2nat.simps) done qed qed lemma actn_0_is_4[simp]: "actn m 0 r = 4" by(simp add: actn.simps) lemma newstat_0_0[simp]: "newstat m 0 r = 0" by(simp add: newstat.simps) declare step_red[simp del] lemma halt_least_step: "\steps0 (Suc 0, Bk\l, ) tp stp = (0, Bk\ m, Oc\rs @ Bk\n); tm_wf (tp, 0); 0 \ \ stp. (nonstop (code tp) (bl2wc ()) stp = 0 \ (\ stp'. nonstop (code tp) (bl2wc ()) stp' = 0 \ stp \ stp'))" proof(induct stp) case 0 then show ?case by (simp add: steps.simps(1)) next case (Suc stp) hence ind: "steps0 (Suc 0, Bk\ l, ) tp stp = (0, Bk\ m, Oc\ rs @ Bk\ n) \ \stp. nonstop (code tp) (bl2wc ()) stp = 0 \ (\stp'. nonstop (code tp) (bl2wc ()) stp' = 0 \ stp \ stp')" and h: "steps0 (Suc 0, Bk\ l, ) tp (Suc stp) = (0, Bk\ m, Oc\ rs @ Bk\ n)" "tm_wf (tp, 0::nat)" "0 < rs" by simp+ { fix a b c nat assume "steps0 (Suc 0, Bk\ l, ) tp stp = (a, b, c)" "a = Suc nat" hence "\stp. nonstop (code tp) (bl2wc ()) stp = 0 \ (\stp'. nonstop (code tp) (bl2wc ()) stp' = 0 \ stp \ stp')" using h apply(rule_tac x = "Suc stp" in exI, auto) apply(drule_tac nonstop_t_eq, simp_all add: nonstop_lemma) proof - fix stp' assume g:"steps0 (Suc 0, Bk\ l, ) tp stp = (Suc nat, b, c)" "nonstop (code tp) (bl2wc ()) stp' = 0" thus "Suc stp \ stp'" proof(cases "Suc stp \ stp'", simp, simp) assume "\ Suc stp \ stp'" hence "stp' \ stp" by simp hence "\ is_final (steps0 (Suc 0, Bk\ l, ) tp stp')" using g apply(cases "steps0 (Suc 0, Bk\ l, ) tp stp'",auto, simp) apply(subgoal_tac "\ n. stp = stp' + n", auto) apply(cases "fst (steps0 (Suc 0, Bk \ l, ) tp stp')", simp_all add: steps.simps) apply(rule_tac x = "stp - stp'" in exI, simp) done hence "nonstop (code tp) (bl2wc ()) stp' = 1" proof(cases "steps0 (Suc 0, Bk\ l, ) tp stp'", simp add: nonstop.simps) fix a b c assume k: "0 < a" "steps0 (Suc 0, Bk\ l, ) tp stp' = (a, b, c)" thus " NSTD (conf (code tp) (bl2wc ()) stp')" using rec_t_eq_steps[of tp l lm stp'] h proof(simp add: conf_lemma) assume "trpl_code (a, b, c) = conf (code tp) (bl2wc ()) stp'" moreover have "NSTD (trpl_code (a, b, c))" using k apply(auto simp: trpl_code.simps NSTD.simps) done ultimately show "NSTD (conf (code tp) (bl2wc ()) stp')" by simp qed qed thus "False" using g by simp qed qed } note [intro] = this from h show "\stp. nonstop (code tp) (bl2wc ()) stp = 0 \ (\stp'. nonstop (code tp) (bl2wc ()) stp' = 0 \ stp \ stp')" by(simp add: step_red, cases "steps0 (Suc 0, Bk\ l, ) tp stp", simp, cases "fst (steps0 (Suc 0, Bk\ l, ) tp stp)", auto simp add: nonstop_t_eq intro:ind dest:nonstop_t_eq) qed lemma conf_trpl_ex: "\ p q r. conf m (bl2wc ()) stp = trpl p q r" apply(induct stp, auto simp: conf.simps inpt.simps trpl.simps newconf.simps) apply(rule_tac x = 0 in exI, rule_tac x = 1 in exI, rule_tac x = "bl2wc ()" in exI) apply(simp) done lemma nonstop_rgt_ex: "nonstop m (bl2wc ()) stpa = 0 \ \ r. conf m (bl2wc ()) stpa = trpl 0 0 r" apply(auto simp: nonstop.simps NSTD.simps split: if_splits) using conf_trpl_ex[of m lm stpa] apply(auto) done lemma max_divisors: "x > Suc 0 \ Max {u. x ^ u dvd x ^ r} = r" proof(rule_tac Max_eqI) assume "x > Suc 0" thus "finite {u. x ^ u dvd x ^ r}" apply(rule_tac finite_power_dvd, auto) done next fix y assume "Suc 0 < x" "y \ {u. x ^ u dvd x ^ r}" thus "y \ r" apply(cases "y\ r", simp) apply(subgoal_tac "\ d. y = r + d") apply(auto simp: power_add) apply(rule_tac x = "y - r" in exI, simp) done next show "r \ {u. x ^ u dvd x ^ r}" by simp qed lemma lo_power: assumes "x > Suc 0" shows "lo (x ^ r) x = r" proof - have "\ Suc 0 < x ^ r \ r = 0" using assms by (metis Suc_lessD Suc_lessI nat_power_eq_Suc_0_iff zero_less_power) moreover have "\xa. \ x ^ xa dvd x ^ r \ r = 0" using dvd_refl assms by(cases "x^r";blast) ultimately show ?thesis using assms by(auto simp: lo.simps loR.simps mod_dvd_simp elim:max_divisors) qed lemma lo_rgt: "lo (trpl 0 0 r) (Pi 2) = r" apply(simp add: trpl.simps lo_power) done lemma conf_keep: "conf m lm stp = trpl 0 0 r \ conf m lm (stp + n) = trpl 0 0 r" apply(induct n) apply(auto simp: conf.simps newconf.simps newleft.simps newrght.simps rght.simps lo_rgt) done lemma halt_state_keep_steps_add: "\nonstop m (bl2wc ()) stpa = 0\ \ conf m (bl2wc ()) stpa = conf m (bl2wc ()) (stpa + n)" apply(drule_tac nonstop_rgt_ex, auto simp: conf_keep) done lemma halt_state_keep: "\nonstop m (bl2wc ()) stpa = 0; nonstop m (bl2wc ()) stpb = 0\ \ conf m (bl2wc ()) stpa = conf m (bl2wc ()) stpb" apply(cases "stpa > stpb") using halt_state_keep_steps_add[of m lm stpb "stpa - stpb"] apply simp using halt_state_keep_steps_add[of m lm stpa "stpb - stpa"] apply(simp) done text \ The correntess of \rec_F\ which relates the interpreter function \rec_F\ with the execution of of TMs. \ lemma terminate_halt: "\steps0 (Suc 0, Bk\l, ) tp stp = (0, Bk\m, Oc\rs@Bk\n); tm_wf (tp,0); 0 \ terminate rec_halt [code tp, (bl2wc ())]" by(frule_tac halt_least_step;force simp:nonstop_lemma intro:terminate_halt_lemma) lemma terminate_F: "\steps0 (Suc 0, Bk\l, ) tp stp = (0, Bk\m, Oc\rs@Bk\n); tm_wf (tp,0); 0 \ terminate rec_F [code tp, (bl2wc ())]" apply(drule_tac terminate_halt, simp_all) apply(erule_tac terminate_F_lemma) done lemma F_correct: "\steps0 (Suc 0, Bk\l, ) tp stp = (0, Bk\m, Oc\rs@Bk\n); tm_wf (tp,0); 0 \ rec_exec rec_F [code tp, (bl2wc ())] = (rs - Suc 0)" apply(frule_tac halt_least_step, auto) apply(frule_tac nonstop_t_eq, auto simp: nonstop_lemma) using rec_t_eq_steps[of tp l lm stp] apply(simp add: conf_lemma) proof - fix stpa assume h: "nonstop (code tp) (bl2wc ()) stpa = 0" "\stp'. nonstop (code tp) (bl2wc ()) stp' = 0 \ stpa \ stp'" "nonstop (code tp) (bl2wc ()) stp = 0" "trpl_code (0, Bk\ m, Oc\ rs @ Bk\ n) = conf (code tp) (bl2wc ()) stp" "steps0 (Suc 0, Bk\ l, ) tp stp = (0, Bk\ m, Oc\ rs @ Bk\ n)" hence g1: "conf (code tp) (bl2wc ()) stpa = trpl_code (0, Bk\ m, Oc\ rs @ Bk\n)" using halt_state_keep[of "code tp" lm stpa stp] by(simp) moreover have g2: "rec_exec rec_halt [code tp, (bl2wc ())] = stpa" using h by(auto simp: rec_exec.simps rec_halt_def nonstop_lemma intro!: Least_equality) show "rec_exec rec_F [code tp, (bl2wc ())] = (rs - Suc 0)" proof - have "valu (rght (conf (code tp) (bl2wc ()) stpa)) = rs - Suc 0" using g1 apply(simp add: valu.simps trpl_code.simps bl2wc.simps bl2nat_append lg_power) done thus "?thesis" by(simp add: rec_exec.simps F_lemma g2) qed qed end \ No newline at end of file