diff --git a/thys/Decreasing-Diagrams-II/Decreasing_Diagrams_II.thy b/thys/Decreasing-Diagrams-II/Decreasing_Diagrams_II.thy --- a/thys/Decreasing-Diagrams-II/Decreasing_Diagrams_II.thy +++ b/thys/Decreasing-Diagrams-II/Decreasing_Diagrams_II.thy @@ -1,962 +1,962 @@ (* * Title: Decreasing Diagrams II * Author: Bertram Felgenhauer (2015) * Maintainer: Bertram Felgenhauer * License: LGPL *) section \Decreasing Diagrams\ theory Decreasing_Diagrams_II imports Decreasing_Diagrams_II_Aux "HOL-Cardinals.Wellorder_Extension" "Abstract-Rewriting.Abstract_Rewriting" begin subsection \Greek accents\ datatype accent = Acute | Grave | Macron lemma UNIV_accent: "UNIV = { Acute, Grave, Macron }" using accent.nchotomy by blast lemma finite_accent: "finite (UNIV :: accent set)" by (simp add: UNIV_accent) type_synonym 'a letter = "accent \ 'a" definition letter_less :: "('a \ 'a) set \ ('a letter \ 'a letter) set" where [simp]: "letter_less R = {(a,b). (snd a, snd b) \ R}" lemma mono_letter_less: "mono letter_less" by (auto simp add: mono_def) subsection \Comparing Greek strings\ type_synonym 'a greek = "'a letter list" definition adj_msog :: "'a greek \ 'a greek \ ('a letter \ 'a greek) \ ('a letter \ 'a greek)" where "adj_msog xs zs l \ case l of (y,ys) \ (y, case fst y of Acute \ ys @ zs | Grave \ xs @ ys | Macron \ ys)" definition ms_of_greek :: "'a greek \ ('a letter \ 'a greek) multiset" where "ms_of_greek as = mset (map (\(xs, y, zs) \ adj_msog xs zs (y, [])) (list_splits as))" lemma adj_msog_adj_msog[simp]: "adj_msog xs zs (adj_msog xs' zs' y) = adj_msog (xs @ xs') (zs' @ zs) y" by (auto simp: adj_msog_def split: accent.splits prod.splits) lemma compose_adj_msog[simp]: "adj_msog xs zs \ adj_msog xs' zs' = adj_msog (xs @ xs') (zs' @ zs)" by (simp add: comp_def) lemma adj_msog_single: "adj_msog xs zs (x,[]) = (x, (case fst x of Grave \ xs | Acute \ zs | Macron \ []))" by (simp add: adj_msog_def split: accent.splits) lemma ms_of_greek_elem: assumes "(x,xs) \ set_mset (ms_of_greek ys)" shows "x \ set ys" using assms by (auto dest: elem_list_splits_elem simp: adj_msog_def ms_of_greek_def) lemma ms_of_greek_shorter: assumes "(x, t) \# ms_of_greek s" shows "length s > length t" using assms[unfolded ms_of_greek_def in_multiset_in_set] by (auto simp: elem_list_splits_length adj_msog_def split: accent.splits) lemma msog_append: "ms_of_greek (xs @ ys) = image_mset (adj_msog [] ys) (ms_of_greek xs) + image_mset (adj_msog xs []) (ms_of_greek ys)" by (auto simp: ms_of_greek_def list_splits_append multiset.map_comp comp_def prod.case_distrib) definition nest :: "('a \ 'a) set \ ('a greek \ 'a greek) set \ ('a greek \ 'a greek) set" where [simp]: "nest r s = {(a,b). (ms_of_greek a, ms_of_greek b) \ mult (letter_less r <*lex*> s)}" lemma mono_nest: "mono (nest r)" unfolding mono_def proof (intro allI impI subsetI) fix R S x assume 1: "R \ S" and 2: "x \ nest r R" from 1 have "mult (letter_less r <*lex*> R) \ mult (letter_less r <*lex*> S)" using mono_mult mono_lex2[of "letter_less r"] unfolding mono_def by blast with 2 show "x \ nest r S" by auto qed lemma nest_mono[mono_set]: "x \ y \ (a,b) \ nest r x \ (a,b) \ nest r y" using mono_nest[unfolded mono_def, rule_format, of x y r] by blast definition greek_less :: "('a \ 'a) set \ ('a greek \ 'a greek) set" where "greek_less r = lfp (nest r)" lemma greek_less_unfold: "greek_less r = nest r (greek_less r)" using mono_nest[of r] lfp_unfold[of "nest r"] by (simp add: greek_less_def) subsection \Preservation of strict partial orders\ lemma strict_order_letter_less: assumes "strict_order r" shows "strict_order (letter_less r)" using assms unfolding irrefl_def trans_def letter_less_def by fast lemma strict_order_nest: assumes r: "strict_order r" and R: "strict_order R" shows "strict_order (nest r R)" proof - have "strict_order (mult (letter_less r <*lex*> R))" using strict_order_letter_less[of r] irrefl_lex_prod[of "letter_less r" R] trans_lex_prod[of "letter_less r" R] strict_order_mult[of "letter_less r <*lex*> R"] assms by fast from this show "strict_order (nest r R)" unfolding nest_def trans_def irrefl_def by fast qed lemma strict_order_greek_less: assumes "strict_order r" shows "strict_order (greek_less r)" by (simp add: greek_less_def strict_order_lfp[OF mono_nest strict_order_nest[OF assms]]) lemma trans_letter_less: assumes "trans r" shows "trans (letter_less r)" using assms unfolding trans_def letter_less_def by fast lemma trans_order_nest: "trans (nest r R)" using trans_mult unfolding nest_def trans_def by fast lemma trans_greek_less[simp]: "trans (greek_less r)" by (subst greek_less_unfold) (rule trans_order_nest) lemma mono_greek_less: "mono greek_less" unfolding greek_less_def mono_def proof (intro allI impI lfp_mono) fix r s :: "('a \ 'a) set" and R :: "('a greek \ 'a greek) set" assume "r \ s" then have "letter_less r <*lex*> R \ letter_less s <*lex*> R" using mono_letter_less mono_lex1 unfolding mono_def by metis then show "nest r R \ nest s R" using mono_mult unfolding nest_def mono_def by blast qed subsection \Involution\ definition inv_letter :: "'a letter \ 'a letter" where "inv_letter l \ case l of (a, x) \ (case a of Grave \ Acute | Acute \ Grave | Macron \ Macron, x)" lemma inv_letter_pair[simp]: "inv_letter (a, x) = (case a of Grave \ Acute | Acute \ Grave | Macron \ Macron, x)" by (simp add: inv_letter_def) lemma snd_inv_letter[simp]: "snd (inv_letter x) = snd x" by (simp add: inv_letter_def split: prod.splits) lemma inv_letter_invol[simp]: "inv_letter (inv_letter x) = x" by (simp add: inv_letter_def split: prod.splits accent.splits) lemma inv_letter_mono[simp]: assumes "(x, y) \ letter_less r" shows "(inv_letter x, inv_letter y) \ letter_less r" using assms by simp definition inv_greek :: "'a greek \ 'a greek" where "inv_greek s = rev (map inv_letter s)" lemma inv_greek_invol[simp]: "inv_greek (inv_greek s) = s" by (simp add: inv_greek_def rev_map comp_def) lemma inv_greek_append: "inv_greek (s @ t) = inv_greek t @ inv_greek s" by (simp add: inv_greek_def) definition inv_msog :: "('a letter \ 'a greek) multiset \ ('a letter \ 'a greek) multiset" where "inv_msog M = image_mset (\(x, t). (inv_letter x, inv_greek t)) M" lemma inv_msog_invol[simp]: "inv_msog (inv_msog M) = M" by (simp add: inv_msog_def multiset.map_comp comp_def prod.case_distrib) lemma ms_of_greek_inv_greek: "ms_of_greek (inv_greek M) = inv_msog (ms_of_greek M)" unfolding inv_msog_def inv_greek_def ms_of_greek_def list_splits_rev list_splits_map mset_map multiset.map_comp mset_rev inv_letter_def adj_msog_def by (rule cong[OF cong[OF refl[of "image_mset"]] refl]) (auto split: accent.splits) lemma inv_greek_mono: assumes "trans r" and "(s, t) \ greek_less r" shows "(inv_greek s, inv_greek t) \ greek_less r" using assms(2) proof (induct "length s + length t" arbitrary: s t rule: less_induct) note * = trans_lex_prod[OF trans_letter_less[OF \trans r\] trans_greek_less[of r]] case (less s t) have "(inv_msog (ms_of_greek s), inv_msog (ms_of_greek t)) \ mult (letter_less r <*lex*> greek_less r)" unfolding inv_msog_def proof (induct rule: mult_of_image_mset[OF * *]) case (1 x y) thus ?case by (auto intro: less(1) split: prod.splits dest!: ms_of_greek_shorter) next case 2 thus ?case using less(2) by (subst(asm) greek_less_unfold) simp qed thus ?case by (subst greek_less_unfold) (auto simp: ms_of_greek_inv_greek) qed subsection \Monotonicity of @{term "greek_less r"}\ lemma greek_less_rempty[simp]: "(a,[]) \ greek_less r \ False" by (subst greek_less_unfold) (auto simp: ms_of_greek_def) lemma greek_less_nonempty: assumes "b \ []" shows "(a,b) \ greek_less r \ (a,b) \ nest r (greek_less r)" by (subst greek_less_unfold) simp lemma greek_less_lempty[simp]: "([],b) \ greek_less r \ b \ []" proof assume "([],b) \ greek_less r" then show "b \ []" using greek_less_rempty by fast next assume "b \ []" then show "([],b) \ greek_less r" unfolding greek_less_nonempty[OF \b \ []\] by (simp add: ms_of_greek_def) qed lemma greek_less_singleton: "(a, b) \ letter_less r \ ([a], [b]) \ greek_less r" by (subst greek_less_unfold) (auto split: accent.splits simp: adj_msog_def ms_of_greek_def) lemma ms_of_greek_cons: "ms_of_greek (x # s) = {# adj_msog [] s (x,[]) #} + image_mset (adj_msog [x] []) (ms_of_greek s)" using msog_append[of "[x]" s] by (auto simp add: adj_msog_def ms_of_greek_def accent.splits) lemma greek_less_cons_mono: assumes "trans r" shows "(s, t) \ greek_less r \ (x # s, x # t) \ greek_less r" proof (induct "length s + length t" arbitrary: s t rule: less_induct) note * = trans_lex_prod[OF trans_letter_less[OF \trans r\] trans_greek_less[of r]] case (less s t) { fix M have "(M + image_mset (adj_msog [x] []) (ms_of_greek s), M + image_mset (adj_msog [x] []) (ms_of_greek t)) \ mult (letter_less r <*lex*> greek_less r)" proof (rule mult_on_union, induct rule: mult_of_image_mset[OF * *]) case (1 x y) thus ?case unfolding adj_msog_def by (auto intro: less(1) split: prod.splits accent.splits dest!: ms_of_greek_shorter) next case 2 thus ?case using less(2) by (subst(asm) greek_less_unfold) simp qed } moreover { fix N have "({# adj_msog [] s (x,[]) #} + N,{# adj_msog [] t (x,[]) #} + N) \ (mult (letter_less r <*lex*> greek_less r))\<^sup>=" by (auto simp: adj_msog_def less split: accent.splits) } ultimately show ?case using transD[OF trans_mult] by (subst greek_less_unfold) (fastforce simp: ms_of_greek_cons) qed lemma greek_less_app_mono2: assumes "trans r" and "(s, t) \ greek_less r" shows "(p @ s, p @ t) \ greek_less r" using assms by (induct p) (auto simp add: greek_less_cons_mono) lemma greek_less_app_mono1: assumes "trans r" and "(s, t) \ greek_less r" shows "(s @ p, t @ p) \ greek_less r" using inv_greek_mono[of r "inv_greek p @ inv_greek s" "inv_greek p @ inv_greek t"] by (simp add: assms inv_greek_append inv_greek_mono greek_less_app_mono2) subsection \Well-founded-ness of @{term "greek_less r"}\ lemma greek_embed: assumes "trans r" shows "list_emb (\a b. (a, b): reflcl (letter_less r)) a b \ (a, b) \ reflcl (greek_less r)" proof (induct rule: list_emb.induct) case (list_emb_Cons a b y) thus ?case using trans_greek_less[unfolded trans_def] \trans r\ greek_less_app_mono1[of r "[]" "[y]" a] greek_less_app_mono2[of r a b "[y]"] by auto next case (list_emb_Cons2 x y a b) thus ?case using trans_greek_less[unfolded trans_def] \trans r\ greek_less_singleton[of x y r] greek_less_app_mono1[of r "[x]" "[y]" a] greek_less_app_mono2[of r a b "[y]"] by auto qed simp lemma wqo_letter_less: assumes t: "trans r" and w: "wqo_on (\a b. (a, b) \ r\<^sup>=) UNIV" shows "wqo_on (\a b. (a, b) \ (letter_less r)\<^sup>=) UNIV" proof (rule wqo_on_hom[of _ id _ "prod_le (=) (\a b. (a, b) \ r\<^sup>=)", unfolded image_id id_apply]) show "wqo_on (prod_le ((=) :: accent \ accent \ bool) (\a b. (a, b) \ r\<^sup>=)) UNIV" by (rule dickson[OF finite_eq_wqo_on[OF finite_accent] w, unfolded UNIV_Times_UNIV]) qed (insert t, auto simp: transp_on_def trans_def prod_le_def) lemma wf_greek_less: assumes "wf r" and "trans r" shows "wf (greek_less r)" proof - obtain q where "r \ q" and "well_order q" by (metis total_well_order_extension \wf r\) define q' where "q' = q - Id" from \well_order q\ have "reflcl q' = q" by (auto simp add: well_order_on_def linear_order_on_def partial_order_on_def preorder_on_def refl_on_def q'_def) from \well_order q\ have "trans q'" and "irrefl q'" unfolding well_order_on_def linear_order_on_def partial_order_on_def preorder_on_def antisym_def trans_def irrefl_def q'_def by blast+ from \r \ q\ \wf r\ have "r \ q'" by (auto simp add: q'_def) have "wqo_on (\a b. (a,b) \ (greek_less q')\<^sup>=) UNIV" proof (intro wqo_on_hom[of "(\a b. (a, b) \ (greek_less q')\<^sup>=)" "id" UNIV "list_emb (\a b. (a, b) \ (letter_less q')\<^sup>=)", unfolded surj_id]) show "transp_on (\a b. (a, b) \ (greek_less q')\<^sup>=) UNIV" using trans_greek_less[of q'] unfolding trans_def transp_on_def by blast next show "\x\UNIV. \y\UNIV. list_emb (\a b. (a, b) \ (letter_less q')\<^sup>=) x y \ (id x, id y) \ (greek_less q')\<^sup>=" using greek_embed[OF \trans q'\] by auto next show "wqo_on (list_emb (\a b. (a, b) \ (letter_less q')\<^sup>=)) UNIV" using higman[OF wqo_letter_less[OF \trans q'\]] \well_order q\ \reflcl q' = q\ by (auto simp: well_order_implies_wqo) qed with wqo_on_imp_wfp_on[OF this] strict_order_strict[OF strict_order_greek_less] \irrefl q'\ \trans q'\ have "wfp_on (\a b. (a, b) \ greek_less q') UNIV" by force then show ?thesis using mono_greek_less \r \ q'\ wf_subset unfolding wf_iff_wfp_on[symmetric] mono_def by metis qed subsection \Basic Comparisons\ lemma pairwise_imp_mult: assumes "N \ {#}" and "\x \ set_mset M. \y \ set_mset N. (x, y) \ r" shows "(M, N) \ mult r" using assms one_step_implies_mult[of _ _ _ "{#}"] by auto lemma singleton_greek_less: assumes as: "snd ` set as \ under r b" shows "(as, [(a,b)]) \ greek_less r" proof - { fix e assume "e \ set_mset (ms_of_greek as)" with as ms_of_greek_elem[of _ _ as] have "(e, ((a,b),[])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def under_def) } moreover have "ms_of_greek [(a,b)] = {# ((a,b),[]) #}" by (auto simp: ms_of_greek_def adj_msog_def split: accent.splits) ultimately show ?thesis by (subst greek_less_unfold) (auto intro!: pairwise_imp_mult) qed lemma peak_greek_less: assumes as: "snd ` set as \ under r a" and b': "b' \ {[(Grave,b)],[]}" and cs: "snd ` set cs \ under r a \ under r b" and a': "a' \ {[(Acute,a)],[]}" and bs: "snd ` set bs \ under r b" shows "(as @ b' @ cs @ a' @ bs, [(Acute,a),(Grave,b)]) \ greek_less r" proof - let ?A = "(Acute,a)" and ?B = "(Grave,b)" have "(ms_of_greek (as @ b' @ cs @ a' @ bs), ms_of_greek [?A,?B]) \ mult (letter_less r <*lex*> greek_less r)" proof (intro pairwise_imp_mult) (* we distinguish 5 cases depending on where in xs an element e originates *) { fix e assume "e \ set_mset (ms_of_greek as)" with as ms_of_greek_elem[of _ _ as] have "(adj_msog [] (b' @ cs @ a' @ bs) e, (?A,[?B])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def under_def) } moreover { fix e assume "e \ set_mset (ms_of_greek b')" with b' singleton_greek_less[OF as] ms_of_greek_elem[of _ _ b'] have "(adj_msog as (cs @ a' @ bs) e, (?B,[?A])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def ms_of_greek_def) } moreover { fix e assume "e \ set_mset (ms_of_greek cs)" with cs ms_of_greek_elem[of _ _ cs] have "(adj_msog (as @ b') (a' @ bs) e, (?A,[?B])) \ letter_less r <*lex*> greek_less r \ (adj_msog (as @ b') (a' @ bs) e, (?B,[?A])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def under_def) } moreover { fix e assume "e \ set_mset (ms_of_greek a')" with a' singleton_greek_less[OF bs] ms_of_greek_elem[of _ _ a'] have "(adj_msog (as @ b' @ cs) bs e, (?A,[?B])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def ms_of_greek_def) } moreover { fix e assume "e \ set_mset (ms_of_greek bs)" with bs ms_of_greek_elem[of _ _ bs] have "(adj_msog (as @ b' @ cs @ a') [] e, (?B,[?A])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def under_def) } moreover have "ms_of_greek [?A,?B] = {# (?B,[?A]), (?A,[?B]) #}" by (simp add: adj_msog_def ms_of_greek_def) ultimately show "\x\set_mset (ms_of_greek (as @ b' @ cs @ a' @ bs)). \y\set_mset (ms_of_greek [?A,?B]). (x, y) \ letter_less r <*lex*> greek_less r" by (auto simp: msog_append) blast qed (auto simp: ms_of_greek_def) then show ?thesis by (subst greek_less_unfold) auto qed lemma rcliff_greek_less1: assumes "trans r" (* unused assumption kept for symmetry with lcliff_greek_less1 *) and as: "snd ` set as \ under r a \ under r b" and b': "b' \ {[(Grave,b)],[]}" and cs: "snd ` set cs \ under r b" and a': "a' = [(Macron,a)]" and bs: "snd ` set bs \ under r b" shows "(as @ b' @ cs @ a' @ bs, [(Macron,a),(Grave,b)]) \ greek_less r" proof - let ?A = "(Macron,a)" and ?B = "(Grave,b)" have *: "ms_of_greek [?A,?B] = {#(?B,[?A]), (?A,[])#}" "ms_of_greek [?A] = {#(?A,[])#}" by (simp_all add: ms_of_greek_def adj_msog_def) then have **: "ms_of_greek [(Macron, a), (Grave, b)] - {#((Macron, a), [])#} \ {#}" by (auto) (* we distinguish 5 cases depending on where in xs an element e originates *) { fix e assume "e \ set_mset (ms_of_greek as)" with as ms_of_greek_elem[of _ _ as] have "(adj_msog [] (b' @ cs @ a' @ bs) e, (?B,[?A])) \ letter_less r <*lex*> greek_less r" by (cases e) (force simp: adj_msog_def under_def) } moreover { fix e assume "e \ set_mset (ms_of_greek b')" with b' singleton_greek_less as ms_of_greek_elem[of _ _ b'] have "(adj_msog as (cs @ a' @ bs) e, (?B,[?A])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def ms_of_greek_def) } moreover { fix e assume "e \ set_mset (ms_of_greek cs)" with cs ms_of_greek_elem[of _ _ cs] have "(adj_msog (as @ b') (a' @ bs) e, (?B,[?A])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def under_def) } moreover { fix e assume "e \ set_mset (ms_of_greek bs)" with bs ms_of_greek_elem[of _ _ bs] have "(adj_msog (as @ b' @ cs @ a') [] e, (?B,[?A])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def under_def) } moreover have "ms_of_greek [?A,?B] = {# (?B,[?A]), (?A,[]) #}" by (simp add: adj_msog_def ms_of_greek_def) ultimately have "\x\set_mset (ms_of_greek (as @ b' @ cs @ a' @ bs) - {#(?A,[])#}). \y\set_mset (ms_of_greek [?A,?B] - {#(?A,[])#}). (x, y) \ letter_less r <*lex*> greek_less r" unfolding msog_append by (auto simp: a' msog_append ac_simps * adj_msog_single) from one_step_implies_mult[OF ** this,of "{#(?A,[])#}"] have "(ms_of_greek (as @ b' @ cs @ a' @ bs), ms_of_greek [?A,?B]) \ mult (letter_less r <*lex*> greek_less r)" unfolding a' msog_append by (auto simp: a' ac_simps * adj_msog_single) then show ?thesis by (subst greek_less_unfold) auto qed lemma rcliff_greek_less2: assumes "trans r" (* unused assumption kept for symmetry with lcliff_greek_less2 *) and as: "snd ` set as \ under r a" and b': "b' \ {[(Grave,b)],[]}" and cs: "snd ` set cs \ under r a \ under r b" shows "(as @ b' @ cs, [(Macron,a),(Grave,b)]) \ greek_less r" proof - let ?A = "(Macron,a)" and ?B = "(Grave,b)" have "(ms_of_greek (as @ b' @ cs), ms_of_greek [?A,?B]) \ mult (letter_less r <*lex*> greek_less r)" proof (intro pairwise_imp_mult) (* we distinguish 3 cases depending on where in xs an element e originates *) { fix e assume "e \ set_mset (ms_of_greek as)" with as ms_of_greek_elem[of _ _ as] have "(adj_msog [] (b' @ cs) e, (?A,[])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def under_def) } moreover { fix e assume "e \ set_mset (ms_of_greek b')" with b' singleton_greek_less[OF as] ms_of_greek_elem[of _ _ b'] have "(adj_msog as (cs) e, (?B,[?A])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def ms_of_greek_def) } moreover { fix e assume "e \ set_mset (ms_of_greek cs)" with cs ms_of_greek_elem[of _ _ cs] have "(adj_msog (as @ b') [] e, (?A,[])) \ letter_less r <*lex*> greek_less r \ (adj_msog (as @ b') [] e, (?B,[?A])) \ letter_less r <*lex*> greek_less r" by (cases e) (fastforce simp: adj_msog_def under_def) } moreover have *: "ms_of_greek [?A,?B] = {# (?B,[?A]), (?A,[]) #}" by (simp add: adj_msog_def ms_of_greek_def) ultimately show "\x\set_mset (ms_of_greek (as @ b' @ cs)). \y\set_mset (ms_of_greek [?A,?B]). (x, y) \ letter_less r <*lex*> greek_less r" by (auto simp: msog_append adj_msog_single ac_simps *) blast qed (auto simp: ms_of_greek_def) then show ?thesis by (subst greek_less_unfold) auto qed lemma snd_inv_greek [simp]: "snd ` set (inv_greek as) = snd ` set as" by (force simp: inv_greek_def) lemma lcliff_greek_less1: assumes "trans r" and as: "snd ` set as \ under r a" and b': "b' = [(Macron,b)]" and cs: "snd ` set cs \ under r a" and a': "a' \ {[(Acute,a)],[]}" and bs: "snd ` set bs \ under r a \ under r b" shows "(as @ b' @ cs @ a' @ bs, [(Acute,a),(Macron,b)]) \ greek_less r" proof - have *: "inv_greek [(Acute,a),(Macron,b)] = [(Macron,b),(Grave,a)]" by (simp add: inv_greek_def) have "(inv_greek (inv_greek (as @ b' @ cs @ a' @ bs)), inv_greek (inv_greek ([(Acute,a),(Macron,b)]))) \ greek_less r" apply (rule inv_greek_mono[OF \trans r\]) apply (unfold inv_greek_append append_assoc *) apply (insert assms) apply (rule rcliff_greek_less1, auto simp: inv_greek_def) done then show ?thesis by simp qed lemma lcliff_greek_less2: assumes "trans r" and cs: "snd ` set cs \ under r a \ under r b" and a': "a' \ {[(Acute,a)],[]}" and bs: "snd ` set bs \ under r b" shows "(cs @ a' @ bs, [(Acute,a),(Macron,b)]) \ greek_less r" proof - have *: "inv_greek [(Acute,a),(Macron,b)] = [(Macron,b),(Grave,a)]" by (simp add: inv_greek_def) have "(inv_greek (inv_greek (cs @ a' @ bs)), inv_greek (inv_greek ([(Acute,a),(Macron,b)]))) \ greek_less r" apply (rule inv_greek_mono[OF \trans r\]) apply (unfold inv_greek_append append_assoc *) apply (insert assms) apply (rule rcliff_greek_less2, auto simp: inv_greek_def) done then show ?thesis by simp qed subsection \Labeled abstract rewriting\ context fixes L R E :: "'b \ 'a rel" begin definition lstep :: "'b letter \ 'a rel" where [simp]: "lstep x = (case x of (a, i) \ (case a of Acute \ (L i)\ | Grave \ R i | Macron \ E i))" fun lconv :: "'b greek \ 'a rel" where "lconv [] = Id" | "lconv (x # xs) = lstep x O lconv xs" lemma lconv_append[simp]: "lconv (xs @ ys) = lconv xs O lconv ys" by (induct xs) auto lemma conversion_join_or_peak_or_cliff: obtains (join) as bs cs where "set as \ {Grave}" and "set bs \ {Macron}" and "set cs \ {Acute}" and "ds = as @ bs @ cs" | (peak) as bs where "ds = as @ ([Acute] @ [Grave]) @ bs" | (lcliff) as bs where "ds = as @ ([Acute] @ [Macron]) @ bs" | (rcliff) as bs where "ds = as @ ([Macron] @ [Grave]) @ bs" proof (induct ds arbitrary: thesis) case (Cons d ds thesis) note IH = this show ?case proof (rule IH(1)) fix as bs assume "ds = as @ ([Acute] @ [Grave]) @ bs" then show ?case using IH(3)[of "d # as" bs] by simp next fix as bs assume "ds = as @ ([Acute] @ [Macron]) @ bs" then show ?case using IH(4)[of "d # as" bs] by simp next fix as bs assume "ds = as @ ([Macron] @ [Grave]) @ bs" then show ?case using IH(5)[of "d # as" bs] by simp next fix as bs cs assume *: "set as \ {Grave}" "set bs \ {Macron}" "set cs \ {Acute}" "ds = as @ bs @ cs" show ?case proof (cases d) case Grave thus ?thesis using * IH(2)[of "d # as" bs cs] by simp next case Macron show ?thesis proof (cases as) case Nil thus ?thesis using * Macron IH(2)[of as "d # bs" cs] by simp next case (Cons a as) thus ?thesis using * Macron IH(5)[of "[]" "as @ bs @ cs"] by simp qed next case Acute show ?thesis proof (cases as) case Nil note as = this show ?thesis proof (cases bs) case Nil thus ?thesis using * as Acute IH(2)[of "[]" "[]" "d # cs"] by simp next case (Cons b bs) thus ?thesis using * as Acute IH(4)[of "[]" "bs @ cs"] by simp qed next case (Cons a as) thus ?thesis using * Acute IH(3)[of "[]" "as @ bs @ cs"] by simp qed qed qed qed auto lemma map_eq_append_split: assumes "map f xs = ys1 @ ys2" obtains xs1 xs2 where "ys1 = map f xs1" "ys2 = map f xs2" "xs = xs1 @ xs2" proof (insert assms, induct ys1 arbitrary: xs thesis) case (Cons y ys) note IH = this show ?case proof (cases xs) case (Cons x xs') show ?thesis proof (rule IH(1)) fix xs1 xs2 assume "ys = map f xs1" "ys2 = map f xs2" "xs' = xs1 @ xs2" thus ?thesis using Cons IH(2)[of "x # xs1" xs2] IH(3) by simp next show "map f xs' = ys @ ys2" using Cons IH(3) by simp qed qed (insert Cons, simp) qed auto lemmas map_eq_append_splits = map_eq_append_split map_eq_append_split[OF sym] abbreviation "conversion' M \ ((\i \ M. R i) \ (\i \ M. E i) \ (\i \ M. L i)\)\<^sup>*" abbreviation "valley' M \ (\i \ M. R i)\<^sup>* O (\i \ M. E i)\<^sup>* O ((\i \ M. L i)\)\<^sup>*" lemma conversion_to_lconv: assumes "(u, v) \ conversion' M" obtains xs where "snd ` set xs \ M" and "(u, v) \ lconv xs" using assms proof (induct arbitrary: thesis rule: converse_rtrancl_induct) case base show ?case using base[of "[]"] by simp next case (step u' x) from step(1) obtain p where "snd p \ M" and "(u', x) \ lstep p" by (force split: accent.splits) moreover obtain xs where "snd ` set xs \ M" "(x, v) \ lconv xs" by (rule step(3)) ultimately show ?case using step(4)[of "p # xs"] by auto qed definition lpeak :: "'b rel \ 'b \ 'b \ 'b greek \ bool" where "lpeak r a b xs \ (\as b' cs a' bs. snd ` set as \ under r a \ b' \ {[(Grave,b)],[]} \ snd ` set cs \ under r a \ under r b \ a' \ {[(Acute,a)],[]} \ snd ` set bs \ under r b \ xs = as @ b' @ cs @ a' @ bs)" definition lcliff :: "'b rel \ 'b \ 'b \ 'b greek \ bool" where "lcliff r a b xs \ (\as b' cs a' bs. snd ` set as \ under r a \ b' = [(Macron,b)] \ snd ` set cs \ under r a \ a' \ {[(Acute,a)],[]} \ snd ` set bs \ under r a \ under r b \ xs = as @ b' @ cs @ a' @ bs) \ (\cs a' bs. snd ` set cs \ under r a \ under r b \ a' \ {[(Acute,a)],[]} \ snd ` set bs \ under r b \ xs = cs @ a' @ bs)" definition rcliff :: "'b rel \ 'b \ 'b \ 'b greek \ bool" where "rcliff r a b xs \ (\as b' cs a' bs. snd ` set as \ under r a \ under r b \ b' \ {[(Grave,b)],[]} \ snd ` set cs \ under r b \ a' = [(Macron,a)] \ snd ` set bs \ under r b \ xs = as @ b' @ cs @ a' @ bs) \ (\as b' cs. snd ` set as \ under r a \ b' \ {[(Grave,b)],[]} \ snd ` set cs \ under r a \ under r b \ xs = as @ b' @ cs)" lemma dd_commute_modulo_conv[case_names wf trans peak lcliff rcliff]: assumes "wf r" and "trans r" and pk: "\a b s t u. (s, t) \ L a \ (s, u) \ R b \ \xs. lpeak r a b xs \ (t, u) \ lconv xs" and lc: "\a b s t u. (s, t) \ L a \ (s, u) \ E b \ \xs. lcliff r a b xs \ (t, u) \ lconv xs" and rc: "\a b s t u. (s, t) \ (E a)\ \ (s, u) \ R b \ \xs. rcliff r a b xs \ (t, u) \ lconv xs" shows "conversion' UNIV \ valley' UNIV" proof (intro subrelI) fix u v assume "(u,v) \ conversion' UNIV" then obtain xs where "(u, v) \ lconv xs" by (auto intro: conversion_to_lconv[of u v]) then show "(u, v) \ valley' UNIV" proof (induct xs rule: wf_induct[of "greek_less r"]) case 1 thus ?case using wf_greek_less[OF \wf r\ \trans r\] . next case (2 xs) show ?case proof (rule conversion_join_or_peak_or_cliff[of "map fst xs"]) fix as bs cs assume *: "set as \ {Grave}" "set bs \ {Macron}" "set cs \ {Acute}" "map fst xs = as @ bs @ cs" then show "(u, v) \ valley' UNIV" proof (elim map_eq_append_splits) fix as' bs' cs' bcs' assume as: "set as \ {Grave}" "as = map fst as'" and bs: "set bs \ {Macron}" "bs = map fst bs'" and cs: "set cs \ {Acute}" "cs = map fst cs'" and xs: "xs = as' @ bcs'" "bcs' = bs' @ cs'" from as(1)[unfolded as(2)] have as': "\x y. (x,y) \ lconv as' \ (x,y) \ (\a. R a)\<^sup>*" proof (induct as') case (Cons x' xs) have "\x y z i. (x,y) \ R i \ (y,z) \ (\a. R a)\<^sup>* \ (x,z) \ (\a. R a)\<^sup>*" by (rule rtrancl_trans) auto with Cons show ?case by auto qed simp from bs(1)[unfolded bs(2)] have bs': "\x y. (x,y) \ lconv bs' \ (x,y) \ (\a. E a)\<^sup>*" proof (induct bs') case (Cons x' xs) have "\x y z i. (x,y) \ E i \ (y,z) \ (\a. E a)\<^sup>* \ (x,z) \ (\a. E a)\<^sup>*" by (rule rtrancl_trans) auto with Cons show ?case by auto qed simp from cs(1)[unfolded cs(2)] have cs': "\x y. (x,y) \ lconv cs' \ (x,y) \ ((\a. L a)\)\<^sup>*" proof (induct cs') case (Cons x' xs) have "\x y z i. (x,y) \ (L i)\ \ (y,z) \ ((\a. L a)\)\<^sup>* \ (x,z) \ ((\a. L a)\)\<^sup>*" by (rule rtrancl_trans) auto with Cons show ?case by auto qed simp from 2(2) as' bs' cs' show "(u, v) \ valley' UNIV" unfolding xs lconv_append by auto (meson relcomp.simps) qed next fix as bs assume *: "map fst xs = as @ ([Acute] @ [Grave]) @ bs" { fix p a b q t' s' u' assume xs: "xs = p @ [(Acute,a),(Grave,b)] @ q" and p: "(u,t') \ lconv p" and a: "(s',t') \ L a" and b: "(s',u') \ R b" and q: "(u',v) \ lconv q" obtain js where lp: "lpeak r a b js" and js: "(t',u') \ lconv js" using pk[OF a b] by auto from lp have "(js, [(Acute,a),(Grave,b)]) \ greek_less r" unfolding lpeak_def using peak_greek_less[of _ r a _ b] by fastforce then have "(p @ js @ q, xs) \ greek_less r" unfolding xs by (intro greek_less_app_mono1 greek_less_app_mono2 \trans r\) auto moreover have "(u, v) \ lconv (p @ js @ q)" using p q js by auto ultimately have "(u, v) \ valley' UNIV" using 2(1) by blast } with * show "(u, v) \ valley' UNIV" using 2(2) by (auto elim!: map_eq_append_splits relcompEpair simp del: append.simps) simp next fix as bs assume *: "map fst xs = as @ ([Acute] @ [Macron]) @ bs" { fix p a b q t' s' u' assume xs: "xs = p @ [(Acute,a),(Macron,b)] @ q" and p: "(u,t') \ lconv p" and a: "(s',t') \ L a" and b: "(s',u') \ E b" and q: "(u',v) \ lconv q" obtain js where lp: "lcliff r a b js" and js: "(t',u') \ lconv js" using lc[OF a b] by auto from lp have "(js, [(Acute,a),(Macron,b)]) \ greek_less r" unfolding lcliff_def using lcliff_greek_less1[OF \trans r\, of _ a _ b] lcliff_greek_less2[OF \trans r\, of _ a b] by fastforce then have "(p @ js @ q, xs) \ greek_less r" unfolding xs by (intro greek_less_app_mono1 greek_less_app_mono2 \trans r\) auto moreover have "(u, v) \ lconv (p @ js @ q)" using p q js by auto ultimately have "(u, v) \ valley' UNIV" using 2(1) by blast } with * show "(u, v) \ valley' UNIV" using 2(2) by (auto elim!: map_eq_append_splits relcompEpair simp del: append.simps) simp next fix as bs assume *: "map fst xs = as @ ([Macron] @ [Grave]) @ bs" { fix p a b q t' s' u' assume xs: "xs = p @ [(Macron,a),(Grave,b)] @ q" and p: "(u,t') \ lconv p" and a: "(s',t') \ (E a)\" and b: "(s',u') \ R b" and q: "(u',v) \ lconv q" obtain js where lp: "rcliff r a b js" and js: "(t',u') \ lconv js" using rc[OF a b] by auto from lp have "(js, [(Macron,a),(Grave,b)]) \ greek_less r" unfolding rcliff_def using rcliff_greek_less1[OF \trans r\, of _ a b] rcliff_greek_less2[OF \trans r\, of _ a _ b] by fastforce then have "(p @ js @ q, xs) \ greek_less r" unfolding xs by (intro greek_less_app_mono1 greek_less_app_mono2 \trans r\) auto moreover have "(u, v) \ lconv (p @ js @ q)" using p q js by auto ultimately have "(u, v) \ valley' UNIV" using 2(1) by blast } with * show "(u, v) \ valley' UNIV" using 2(2) by (auto elim!: map_eq_append_splits relcompEpair simp del: append.simps) simp qed qed qed section \Results\ subsection \Church-Rosser modulo\ text \Decreasing diagrams for Church-Rosser modulo, commutation version.\ lemma dd_commute_modulo[case_names wf trans peak lcliff rcliff]: assumes "wf r" and "trans r" and pk: "\a b s t u. (s, t) \ L a \ (s, u) \ R b \ (t, u) \ conversion' (under r a) O (R b)\<^sup>= O conversion' (under r a \ under r b) O ((L a)\)\<^sup>= O conversion' (under r b)" and lc: "\a b s t u. (s, t) \ L a \ (s, u) \ E b \ (t, u) \ conversion' (under r a) O E b O conversion' (under r a) O ((L a)\)\<^sup>= O conversion' (under r a \ under r b) \ (t, u) \ conversion' (under r a \ under r b) O ((L a )\)\<^sup>= O conversion' (under r b)" and rc: "\a b s t u. (s, t) \ (E a)\ \ (s, u) \ R b \ (t, u) \ conversion' (under r a \ under r b) O (R b)\<^sup>= O conversion' (under r b) O E a O conversion' (under r b) \ (t, u) \ conversion' (under r a) O (R b)\<^sup>= O conversion' (under r a \ under r b)" shows "conversion' UNIV \ valley' UNIV" proof (cases rule: dd_commute_modulo_conv[of r]) case (peak a b s t u) { fix w x y z assume "(t, w) \ conversion' (under r a)" from conversion_to_lconv[OF this] obtain as where "snd ` set as \ under r a" "(t, w) \ lconv as" by auto moreover assume "(w, x) \ (R b)\<^sup>=" then obtain b' where "b' \ {[(Grave,b)],[]}" "(w, x) \ lconv b'" by fastforce moreover assume "(x, y) \ conversion' (under r a \ under r b)" from conversion_to_lconv[OF this] obtain cs where "snd ` set cs \ under r a \ under r b" "(x, y) \ lconv cs" by auto moreover assume "(y, z) \ ((L a)\)\<^sup>=" then obtain a' where "a' \ {[(Acute,a)],[]}" "(y, z) \ lconv a'" by fastforce moreover assume "(z, u) \ conversion' (under r b)" from conversion_to_lconv[OF this] obtain bs where "snd ` set bs \ under r b" "(z, u) \ lconv bs" by auto ultimately have "\xs. lpeak r a b xs \ (t, u) \ lconv xs" by (intro exI[of _ "as @ b' @ cs @ a' @ bs"], unfold lconv_append lpeak_def) blast } then show ?case using pk[OF peak] by blast next case (lcliff a b s t u) { fix w x y z assume "(t, w) \ conversion' (under r a)" from conversion_to_lconv[OF this] obtain as where "snd ` set as \ under r a" "(t, w) \ lconv as" by auto moreover assume "(w, x) \ E b" then obtain b' where "b' = [(Macron,b)]" "(w, x) \ lconv b'" by fastforce moreover assume "(x, y) \ conversion' (under r a)" from conversion_to_lconv[OF this] obtain cs where "snd ` set cs \ under r a" "(x, y) \ lconv cs" by auto moreover assume "(y, z) \ ((L a)\)\<^sup>=" then obtain a' where "a' \ {[(Acute,a)],[]}" "(y, z) \ lconv a'" by fastforce moreover assume "(z, u) \ conversion' (under r a \ under r b)" from conversion_to_lconv[OF this] obtain bs where "snd ` set bs \ under r a \ under r b" "(z, u) \ lconv bs" by auto ultimately have "\xs. lcliff r a b xs \ (t, u) \ lconv xs" by (intro exI[of _ "as @ b' @ cs @ a' @ bs"], unfold lconv_append lcliff_def) blast } moreover { fix w x assume "(t, w) \ conversion' (under r a \ under r b)" from conversion_to_lconv[OF this] obtain cs where "snd ` set cs \ under r a \ under r b" "(t, w) \ lconv cs" by auto moreover assume "(w, x) \ ((L a)\)\<^sup>=" then obtain a' where "a' \ {[(Acute,a)],[]}" "(w, x) \ lconv a'" by fastforce moreover assume "(x, u) \ conversion' (under r b)" from conversion_to_lconv[OF this] obtain bs where "snd ` set bs \ under r b" "(x, u) \ lconv bs" by auto ultimately have "\xs. lcliff r a b xs \ (t, u) \ lconv xs" by (intro exI[of _ "cs @ a' @ bs"], unfold lconv_append lcliff_def) blast } ultimately show ?case using lc[OF lcliff] by blast next case (rcliff a b s t u) { fix w x y z assume "(t, w) \ conversion' (under r a \ under r b)" from conversion_to_lconv[OF this] obtain as where "snd ` set as \ under r a \ under r b" "(t, w) \ lconv as" by auto moreover assume "(w, x) \ (R b)\<^sup>=" then obtain b' where "b' \ {[(Grave,b)],[]}" "(w, x) \ lconv b'" by fastforce moreover assume "(x, y) \ conversion' (under r b)" from conversion_to_lconv[OF this] obtain cs where "snd ` set cs \ under r b" "(x, y) \ lconv cs" by auto moreover assume "(y, z) \ E a" then obtain a' where "a' = [(Macron,a)]" "(y, z) \ lconv a'" by fastforce moreover assume "(z, u) \ conversion' (under r b)" from conversion_to_lconv[OF this] obtain bs where "snd ` set bs \ under r b" "(z, u) \ lconv bs" by auto ultimately have "\xs. rcliff r a b xs \ (t, u) \ lconv xs" by (intro exI[of _ "as @ b' @ cs @ a' @ bs"], unfold lconv_append rcliff_def) blast } moreover { fix w x assume "(t, w) \ conversion' (under r a)" from conversion_to_lconv[OF this] obtain as where "snd ` set as \ under r a" "(t, w) \ lconv as" by auto moreover assume "(w, x) \ (R b)\<^sup>=" then obtain b' where "b' \ {[(Grave,b)],[]}" "(w, x) \ lconv b'" by fastforce moreover assume "(x, u) \ conversion' (under r a \ under r b)" from conversion_to_lconv[OF this] obtain cs where "snd ` set cs \ under r a \ under r b" "(x, u) \ lconv cs" by auto ultimately have "\xs. rcliff r a b xs \ (t, u) \ lconv xs" by (intro exI[of _ "as @ b' @ cs"], unfold lconv_append rcliff_def) blast } ultimately show ?case using rc[OF rcliff] by blast qed fact+ end (* context *) text \Decreasing diagrams for Church-Rosser modulo.\ lemma dd_cr_modulo[case_names wf trans symE peak cliff]: assumes "wf r" and "trans r" and E: "\i. sym (E i)" and pk: "\a b s t u. (s, t) \ L a \ (s, u) \ L b \ (t, u) \ conversion' L L E (under r a) O (L b)\<^sup>= O conversion' L L E (under r a \ under r b) O ((L a)\)\<^sup>= O conversion' L L E (under r b)" and cl: "\a b s t u. (s, t) \ L a \ (s, u) \ E b \ (t, u) \ conversion' L L E (under r a) O E b O conversion' L L E (under r a) O ((L a)\)\<^sup>= O conversion' L L E (under r a \ under r b) \ (t, u) \ conversion' L L E (under r a \ under r b) O ((L a )\)\<^sup>= O conversion' L L E (under r b)" shows "conversion' L L E UNIV \ valley' L L E UNIV" proof (induct rule: dd_commute_modulo[of r]) note E' = E[unfolded sym_conv_converse_eq] case (rcliff a b s t u) show ?case - using cl[OF rcliff(2) rcliff(1)[unfolded E'], unfolded converse_iff[of t u,symmetric]] - by (auto simp only: ac_simps E' converse_inward) + using cl[OF rcliff(2) rcliff(1)[unfolded E'], unfolded converse_iff[of t u,symmetric]] + by (auto simp only: E' converse_inward) (auto simp only: ac_simps) qed fact+ subsection \Commutation and confluence\ abbreviation "conversion'' L R M \ ((\i \ M. R i) \ (\i \ M. L i)\)\<^sup>*" abbreviation "valley'' L R M \ (\i \ M. R i)\<^sup>* O ((\i \ M. L i)\)\<^sup>*" text \Decreasing diagrams for commutation.\ lemma dd_commute[case_names wf trans peak]: assumes "wf r" and "trans r" and pk: "\a b s t u. (s, t) \ L a \ (s, u) \ R b \ (t, u) \ conversion'' L R (under r a) O (R b)\<^sup>= O conversion'' L R (under r a \ under r b) O ((L a)\)\<^sup>= O conversion'' L R (under r b)" shows "commute (\i. L i) (\i. R i)" proof - have "((\i. L i)\)\<^sup>* O (\i. R i)\<^sup>* \ conversion'' L R UNIV" by regexp also have "\ \ valley'' L R UNIV" using dd_commute_modulo[OF assms(1,2), of L R "\_. {}"] pk by auto finally show ?thesis by (simp only: commute_def) qed text \Decreasing diagrams for confluence.\ lemmas dd_cr[case_names wf trans peak] = dd_commute[of _ L L for L, unfolded CR_iff_self_commute[symmetric]] subsection \Extended decreasing diagrams\ context fixes r q :: "'b rel" assumes "wf r" and "trans r" and "trans q" and "refl q" and compat: "r O q \ r" begin private abbreviation (input) down :: "('b \ 'a rel) \ ('b \ 'a rel)" where "down L \ \i. \j \ under q i. L j" private lemma Union_down: "(\i. down L i) = (\i. L i)" using \refl q\ by (auto simp: refl_on_def under_def) text \Extended decreasing diagrams for commutation.\ lemma edd_commute[case_names wf transr transq reflq compat peak]: assumes pk: "\a b s t u. (s, t) \ L a \ (s, u) \ R b \ (t, u) \ conversion'' L R (under r a) O (down R b)\<^sup>= O conversion'' L R (under r a \ under r b) O ((down L a)\)\<^sup>= O conversion'' L R (under r b)" shows "commute (\i. L i) (\i. R i)" unfolding Union_down[of L, symmetric] Union_down[of R, symmetric] proof (induct rule: dd_commute[of r "down L" "down R"]) case (peak a b s t u) then obtain a' b' where a': "(a', a) \ q" "(s, t) \ L a'" and b': "(b', b) \ q" "(s, u) \ R b'" by (auto simp: under_def) have "\a' a. (a',a) \ q \ under r a' \ under r a" using compat by (auto simp: under_def) then have aux1: "\a' a L. (a',a) \ q \ (\i \ under r a'. L i) \ (\i \ under r a. L i)" by auto have aux2: "\a' a L. (a',a) \ q \ down L a' \ down L a" using \trans q\ by (auto simp: under_def trans_def) have aux3: "\a L. (\i \ under r a. L i) \ (\i \ under r a. down L i)" using \refl q\ by (auto simp: under_def refl_on_def) from aux1[OF a'(1), of L] aux1[OF a'(1), of R] aux2[OF a'(1), of L] aux1[OF b'(1), of L] aux1[OF b'(1), of R] aux2[OF b'(1), of R] aux3[of L] aux3[of R] show ?case by (intro subsetD[OF _ pk[OF \(s, t) \ L a'\ \(s, u) \ R b'\]], unfold UN_Un) (intro relcomp_mono rtrancl_mono Un_mono iffD2[OF converse_mono]; fast) qed fact+ text \Extended decreasing diagrams for confluence.\ lemmas edd_cr[case_names wf transr transq reflq compat peak] = edd_commute[of L L for L, unfolded CR_iff_self_commute[symmetric]] end (* context *) end (* Decreasing_Diagrams_II *) diff --git a/thys/JinjaThreads/Compiler/JVMJ1.thy b/thys/JinjaThreads/Compiler/JVMJ1.thy --- a/thys/JinjaThreads/Compiler/JVMJ1.thy +++ b/thys/JinjaThreads/Compiler/JVMJ1.thy @@ -1,4129 +1,4131 @@ (* Title: JinjaThreads/Compiler/JVMJ1.thy Author: Andreas Lochbihler *) section \Correctness of Stage 2: From JVM to intermediate language\ theory JVMJ1 imports J1JVMBisim begin declare split_paired_Ex[simp del] lemma rec_option_is_case_option: "rec_option = case_option" apply (rule ext)+ apply (rename_tac y) apply (case_tac y) apply auto done context J1_JVM_heap_base begin lemma assumes ha: "typeof_addr h a = \Class_type D\" and subclsObj: "P \ D \\<^sup>* Throwable" shows bisim1_xcp_\Red: "\ P,e,h \ (e', xs) \ (stk, loc, pc, \a\); match_ex_table (compP f P) (cname_of h a) pc (compxE2 e 0 0) = None; \n. n + max_vars e' \ length xs \ \ e n \ \ \red1r P t h (e', xs) (Throw a, loc) \ P,e,h \ (Throw a, loc) \ (stk, loc, pc, \a\)" and bisims1_xcp_\Reds: "\ P,es,h \ (es', xs) [\] (stk, loc, pc, \a\); match_ex_table (compP f P) (cname_of h a) pc (compxEs2 es 0 0) = None; \n. n + max_varss es' \ length xs \ \s es n \ \ \vs es''. \reds1r P t h (es', xs) (map Val vs @ Throw a # es'', loc) \ P,es,h \ (map Val vs @ Throw a # es'', loc) [\] (stk, loc, pc, \a\)" proof(induct "(e', xs)" "(stk, loc, pc, \a :: 'addr\)" and "(es', xs)" "(stk, loc, pc, \a :: 'addr\)" arbitrary: e' xs stk loc pc and es' xs stk loc pc rule: bisim1_bisims1.inducts) case bisim1NewThrow thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1NewArray thus ?case by(auto intro: rtranclp.rtrancl_into_rtrancl New1ArrayThrow bisim1_bisims1.intros dest: bisim1_ThrowD elim!: NewArray_\red1r_xt) next case bisim1NewArrayThrow thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1NewArrayFail thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1Cast thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl Cast1Throw bisim1_bisims1.intros dest: bisim1_ThrowD elim!: Cast_\red1r_xt) next case bisim1CastThrow thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1CastFail thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1InstanceOf thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl InstanceOf1Throw bisim1_bisims1.intros dest: bisim1_ThrowD elim!: InstanceOf_\red1r_xt) next case bisim1InstanceOfThrow thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1BinOp1 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl Bin1OpThrow1 bisim1_bisims1.intros simp add: match_ex_table_append dest: bisim1_ThrowD elim!: BinOp_\red1r_xt1) next case bisim1BinOp2 thus ?case by(clarsimp simp add: match_ex_table_append_not_pcs compxE2_size_convs compxE2_stack_xlift_convs match_ex_table_shift_pc_None) (fastforce intro: rtranclp.rtrancl_into_rtrancl red1_reds1.intros bisim1BinOpThrow2 elim!: BinOp_\red1r_xt2) next case bisim1BinOpThrow1 thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros) next case (bisim1BinOpThrow2 e xs stk loc pc e1 bop) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ hence "xs = loc" by(auto dest: bisim1_ThrowD) with bisim show ?case by(auto intro: bisim1_bisims1.bisim1BinOpThrow2) next case bisim1BinOpThrow thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1LAss1 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl LAss1Throw bisim1_bisims1.intros dest: bisim1_ThrowD elim!: LAss_\red1r) next case bisim1LAssThrow thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1AAcc1 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl AAcc1Throw1 bisim1_bisims1.intros simp add: match_ex_table_append dest: bisim1_ThrowD elim!: AAcc_\red1r_xt1) next case bisim1AAcc2 thus ?case by(clarsimp simp add: match_ex_table_append_not_pcs compxE2_size_convs compxE2_stack_xlift_convs match_ex_table_shift_pc_None) (fastforce intro: rtranclp.rtrancl_into_rtrancl red1_reds1.intros bisim1AAccThrow2 elim!: AAcc_\red1r_xt2) next case bisim1AAccThrow1 thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros) next case bisim1AAccThrow2 thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros dest: bisim1_ThrowD) next case bisim1AAccFail thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1AAss1 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl AAss1Throw1 bisim1_bisims1.intros simp add: match_ex_table_append dest: bisim1_ThrowD elim!: AAss_\red1r_xt1) next case bisim1AAss2 thus ?case by(clarsimp simp add: compxE2_size_convs compxE2_stack_xlift_convs) (fastforce simp add: match_ex_table_append intro: rtranclp.rtrancl_into_rtrancl red1_reds1.intros bisim1AAssThrow2 elim!: AAss_\red1r_xt2) next case bisim1AAss3 thus ?case by(clarsimp simp add: compxE2_size_convs compxE2_stack_xlift_convs) (fastforce simp add: match_ex_table_append intro: rtranclp.rtrancl_into_rtrancl red1_reds1.intros bisim1AAssThrow3 elim!: AAss_\red1r_xt3) next case bisim1AAssThrow1 thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros) next case (bisim1AAssThrow2 e xs stk loc pc i e2) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ hence "xs = loc" by(auto dest: bisim1_ThrowD) with bisim show ?case by(auto intro: bisim1_bisims1.bisim1AAssThrow2) next case (bisim1AAssThrow3 e xs stk loc pc A i) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ hence "xs = loc" by(auto dest: bisim1_ThrowD) with bisim show ?case by(auto intro: bisim1_bisims1.bisim1AAssThrow3) next case bisim1AAssFail thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1ALength thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl ALength1Throw bisim1_bisims1.intros dest: bisim1_ThrowD elim!: ALength_\red1r_xt) next case bisim1ALengthThrow thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1ALengthNull thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1FAcc thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl FAcc1Throw bisim1_bisims1.intros dest: bisim1_ThrowD elim!: FAcc_\red1r_xt) next case bisim1FAccThrow thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1FAccNull thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1FAss1 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl FAss1Throw1 bisim1_bisims1.intros simp add: match_ex_table_append dest: bisim1_ThrowD elim!: FAss_\red1r_xt1) next case bisim1FAss2 thus ?case by(clarsimp simp add: match_ex_table_append_not_pcs compxE2_size_convs compxE2_stack_xlift_convs) (fastforce intro: rtranclp.rtrancl_into_rtrancl red1_reds1.intros bisim1FAssThrow2 elim!: FAss_\red1r_xt2) next case bisim1FAssThrow1 thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros) next case (bisim1FAssThrow2 e2 xs stk loc pc e) note bisim = \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ hence "xs = loc" by(auto dest: bisim1_ThrowD) with bisim show ?case by(auto intro: bisim1_bisims1.bisim1FAssThrow2) next case bisim1FAssNull thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1CAS1 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl CAS1Throw bisim1_bisims1.intros simp add: match_ex_table_append dest: bisim1_ThrowD elim!: CAS_\red1r_xt1) next case bisim1CAS2 thus ?case by(clarsimp simp add: compxE2_size_convs compxE2_stack_xlift_convs) (fastforce simp add: match_ex_table_append intro: rtranclp.rtrancl_into_rtrancl red1_reds1.intros bisim1CASThrow2 elim!: CAS_\red1r_xt2) next case bisim1CAS3 thus ?case by(clarsimp simp add: compxE2_size_convs compxE2_stack_xlift_convs) (fastforce simp add: match_ex_table_append intro: rtranclp.rtrancl_into_rtrancl red1_reds1.intros bisim1CASThrow3 elim!: CAS_\red1r_xt3) next case bisim1CASThrow1 thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros) next case (bisim1CASThrow2 e xs stk loc pc i e2) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ hence "xs = loc" by(auto dest: bisim1_ThrowD) with bisim show ?case by(auto intro: bisim1_bisims1.bisim1CASThrow2) next case (bisim1CASThrow3 e xs stk loc pc A i) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ hence "xs = loc" by(auto dest: bisim1_ThrowD) with bisim show ?case by(auto intro: bisim1_bisims1.bisim1CASThrow3) next case bisim1CASFail thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1Call1 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl Call1ThrowObj bisim1_bisims1.intros simp add: match_ex_table_append dest: bisim1_ThrowD elim!: Call_\red1r_obj) next case bisim1CallParams thus ?case by(clarsimp simp add: match_ex_table_append_not_pcs compxE2_size_convs compxEs2_size_convs compxE2_stack_xlift_convs compxEs2_stack_xlift_convs match_ex_table_shift_pc_None) (fastforce intro: rtranclp.rtrancl_into_rtrancl Call1ThrowParams[OF refl] bisim1CallThrowParams elim!: Call_\red1r_param) next case bisim1CallThrowObj thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros) next case (bisim1CallThrowParams es vs es' xs stk loc pc obj M) note bisim = \P,es,h \ (map Val vs @ Throw a # es', xs) [\] (stk, loc, pc, \a\)\ hence "xs = loc" by(auto dest: bisims1_ThrowD) with bisim show ?case by(auto intro: bisim1_bisims1.bisim1CallThrowParams) next case bisim1CallThrow thus ?case by(auto intro: bisim1_bisims1.intros) next case (bisim1BlockSome4 e e' xs stk loc pc V Ty v) from \\n. n + max_vars {V:Ty=None; e'} \ length xs \ \ {V:Ty=\v\; e} n\ have V: "V < length xs" by simp from bisim1BlockSome4 have Red: "\red1r P t h (e', xs) (Throw a, loc)" and bisim: "P,e,h \ (Throw a, loc) \ (stk, loc, pc, \a\)" by(auto simp add: match_ex_table_append_not_pcs compxE2_size_convs compxEs2_size_convs compxE2_stack_xlift_convs compxEs2_stack_xlift_convs match_ex_table_shift_pc_None intro!: exI[where x="Suc V"] elim: meta_impE) note len = \red1r_preserves_len[OF Red] from Red have "\red1r P t h ({V:Ty=None; e'}, xs) ({V:Ty=None; Throw a}, loc)" by(auto intro: Block_None_\red1r_xt) thus ?case using V len bisim by(auto intro: \move1BlockThrow Block1Throw bisim1BlockThrowSome elim!: rtranclp.rtrancl_into_rtrancl) next case bisim1BlockThrowSome thus ?case by(auto dest: bisim1_ThrowD intro: bisim1_bisims1.bisim1BlockThrowSome) next case (bisim1BlockNone e e' xs stk loc pc V Ty) hence Red: "\red1r P t h (e', xs) (Throw a, loc)" and bisim: "P,e,h \ (Throw a, loc) \ (stk, loc, pc, \a\)" by(auto elim: meta_impE intro!: exI[where x="Suc V"]) from Red have len: "length loc = length xs" by(rule \red1r_preserves_len) from \\n. n + max_vars {V:Ty=None; e'} \ length xs \ \ {V:Ty=None; e} n\ have V: "V < length xs" by simp from Red have "\red1r P t h ({V:Ty=None; e'}, xs) ({V:Ty=None; Throw a}, loc)" by(rule Block_None_\red1r_xt) thus ?case using V len bisim by(auto intro: \move1BlockThrow Block1Throw bisim1BlockThrowNone elim!: rtranclp.rtrancl_into_rtrancl) next case bisim1BlockThrowNone thus ?case by(auto dest: bisim1_ThrowD intro: bisim1_bisims1.bisim1BlockThrowNone) next case bisim1Sync1 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl Synchronized1Throw1 bisim1_bisims1.intros simp add: match_ex_table_append dest: bisim1_ThrowD elim!: Sync_\red1r_xt) next case (bisim1Sync4 e2 e' xs stk loc pc V e1 a') from \P,e2,h \ (e', xs) \ (stk, loc, pc, \a\)\ have "pc < length (compE2 e2)" by(auto dest!: bisim1_xcp_pcD) with \match_ex_table (compP f P) (cname_of h a) (Suc (Suc (Suc (length (compE2 e1) + pc)))) (compxE2 (sync\<^bsub>V\<^esub> (e1) e2) 0 0) = None\ subclsObj ha have False by(simp add: match_ex_table_append matches_ex_entry_def split: if_split_asm) thus ?case .. next case bisim1Sync10 thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1Sync11 thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1Sync12 thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1Sync14 thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1SyncThrow thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros) next case bisim1Seq1 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl Seq1Throw bisim1_bisims1.intros dest: bisim1_ThrowD elim!: Seq_\red1r_xt simp add: match_ex_table_append) next case bisim1SeqThrow1 thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros) next case bisim1Seq2 thus ?case by(auto simp add: match_ex_table_append_not_pcs compxE2_size_convs compxE2_stack_xlift_convs match_ex_table_shift_pc_None intro: bisim1_bisims1.bisim1Seq2) next case bisim1Cond1 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl Cond1Throw bisim1_bisims1.intros dest: bisim1_ThrowD elim!: Cond_\red1r_xt simp add: match_ex_table_append) next case bisim1CondThen thus ?case by(clarsimp simp add: match_ex_table_append) (auto simp add: match_ex_table_append_not_pcs compxE2_size_convs compxE2_stack_xlift_convs match_ex_table_shift_pc_None intro: bisim1_bisims1.bisim1CondThen) next case bisim1CondElse thus ?case by(clarsimp simp add: match_ex_table_append) (auto simp add: match_ex_table_append_not_pcs compxE2_size_convs compxE2_stack_xlift_convs match_ex_table_shift_pc_None intro: bisim1_bisims1.bisim1CondElse) next case bisim1CondThrow thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros) next case bisim1While3 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl Cond1Throw bisim1_bisims1.intros dest: bisim1_ThrowD elim!: Cond_\red1r_xt simp add: match_ex_table_append) next case (bisim1While4 e e' xs stk loc pc c) hence "\red1r P t h (e', xs) (Throw a, loc) \ P,e,h \ (Throw a, loc) \ (stk, loc, pc, \a\)" by(auto simp add: match_ex_table_append_not_pcs compxE2_size_convs compxEs2_size_convs compxE2_stack_xlift_convs compxEs2_stack_xlift_convs match_ex_table_shift_pc_None) hence "\red1r P t h (e';;while (c) e, xs) (Throw a, loc)" "P,while (c) e,h \ (Throw a, loc) \ (stk, loc, Suc (length (compE2 c) + pc), \a\)" by(auto intro: rtranclp.rtrancl_into_rtrancl Seq1Throw \move1SeqThrow bisim1WhileThrow2 elim!: Seq_\red1r_xt) thus ?case .. next case bisim1WhileThrow1 thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros) next case bisim1WhileThrow2 thus ?case by(auto simp add: match_ex_table_append intro: bisim1_bisims1.intros dest: bisim1_ThrowD) next case bisim1Throw1 thus ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl Throw1Throw bisim1_bisims1.intros dest: bisim1_ThrowD elim!: Throw_\red1r_xt) next case bisim1Throw2 thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1ThrowNull thus ?case by(fastforce intro: bisim1_bisims1.intros) next case bisim1ThrowThrow thus ?case by(fastforce intro: bisim1_bisims1.intros) next case (bisim1Try e e' xs stk loc pc C V e2) hence red: "\red1r P t h (e', xs) (Throw a, loc)" and bisim: "P,e,h \ (Throw a, loc) \ (stk, loc, pc, \a\)" by(auto simp add: match_ex_table_append) from red have Red: "\red1r P t h (try e' catch(C V) e2, xs) (try Throw a catch(C V) e2, loc)" by(rule Try_\red1r_xt) from \match_ex_table (compP f P) (cname_of h a) pc (compxE2 (try e catch(C V) e2) 0 0) = None\ have "\ matches_ex_entry (compP f P) (cname_of h a) pc (0, length (compE2 e), \C\, Suc (length (compE2 e)), 0)" by(auto simp add: match_ex_table_append split: if_split_asm) moreover from \P,e,h \ (e', xs) \ (stk, loc, pc, \a\)\ have "pc < length (compE2 e)" by(auto dest: bisim1_xcp_pcD) ultimately have subcls: "\ P \ D \\<^sup>* C" using ha by(simp add: matches_ex_entry_def cname_of_def) with ha have "True,P,t \1 \try Throw a catch(C V) e2, (h, loc)\ -\\ \Throw a, (h, loc)\" by -(rule Red1TryFail, auto) moreover from bisim ha subcls have "P,try e catch(C V) e2,h \ (Throw a, loc) \ (stk, loc, pc, \a\)" by(rule bisim1TryFail) ultimately show ?case using Red by(blast intro: rtranclp.rtrancl_into_rtrancl \move1TryThrow) next case (bisim1TryCatch2 e2 e' xs stk loc pc e1 C V) hence *: "\red1r P t h (e', xs) (Throw a, loc) \ P,e2,h \ (Throw a, loc) \ (stk, loc, pc, \a\)" by(clarsimp simp add: match_ex_table_append matches_ex_entry_def split: if_split_asm) (auto simp add: match_ex_table_append compxE2_size_convs compxE2_stack_xlift_convs match_ex_table_shift_pc_None elim: meta_impE intro!: exI[where x="Suc V"]) moreover note \red1r_preserves_len[OF *[THEN conjunct1]] moreover from \\n. n + max_vars {V:Class C=None; e'} \ length xs \ \ (try e1 catch(C V) e2) n\ have "V < length xs" by simp ultimately show ?case by(fastforce intro: rtranclp.rtrancl_into_rtrancl Block1Throw \move1BlockThrow bisim1TryCatchThrow elim!: Block_None_\red1r_xt) next case (bisim1TryFail e xs stk loc pc C'' C' V e2) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ hence "xs = loc" by(auto dest: bisim1_ThrowD) with bisim \typeof_addr h a = \Class_type C''\\ \\ P \ C'' \\<^sup>* C'\ show ?case by(auto intro: bisim1_bisims1.bisim1TryFail) next case (bisim1TryCatchThrow e2 xs stk loc pc e C' V) from \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ have "xs = loc" by(auto dest: bisim1_ThrowD) with \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ show ?case by(auto intro: bisim1_bisims1.bisim1TryCatchThrow) next case (bisims1List1 e e' xs stk loc pc es) hence "\red1r P t h (e', xs) (Throw a, loc)" and bisim: "P,e,h \ (Throw a, loc) \ (stk, loc, pc, \a\)" by(auto simp add: match_ex_table_append) hence "\reds1r P t h (e' # es, xs) (map Val [] @ Throw a # es, loc)" by(auto intro: \red1r_inj_\reds1r) moreover from bisim have "P,e#es,h \ (Throw a # es, loc) [\] (stk, loc, pc, \a\)" by(rule bisim1_bisims1.bisims1List1) ultimately show ?case by fastforce next case (bisims1List2 es es' xs stk loc pc e v) hence "\vs es''. \reds1r P t h (es', xs) (map Val vs @ Throw a # es'', loc) \ P,es,h \ (map Val vs @ Throw a # es'', loc) [\] (stk, loc, pc, \a\)" by(auto simp add: match_ex_table_append_not_pcs compxEs2_size_convs compxEs2_stack_xlift_convs match_ex_table_shift_pc_None) then obtain vs es'' where red: "\reds1r P t h (es', xs) (map Val vs @ Throw a # es'', loc)" and bisim: "P,es,h \ (map Val vs @ Throw a # es'', loc) [\] (stk, loc, pc, \a\)" by blast from red have "\reds1r P t h (Val v # es', xs) (map Val (v # vs) @ Throw a # es'', loc)" by(auto intro: \reds1r_cons_\reds1r) moreover from bisim have "P,e # es,h \ (map Val (v # vs) @ Throw a # es'', loc) [\] (stk @ [v], loc, length (compE2 e) + pc, \a\)" by(auto intro: bisim1_bisims1.bisims1List2) ultimately show ?case by fastforce qed primrec conf_xcp' :: "'m prog \ 'heap \ 'addr option \ bool" where "conf_xcp' P h None = True" | "conf_xcp' P h \a\ = (\D. typeof_addr h a = \Class_type D\ \ P \ D \\<^sup>* Throwable)" lemma conf_xcp_conf_xcp': "conf_xcp P h xcp i \ conf_xcp' P h xcp" by(cases xcp) auto lemma conf_xcp'_compP [simp]: "conf_xcp' (compP f P) = conf_xcp' P" by(clarsimp simp add: fun_eq_iff conf_xcp'_def rec_option_is_case_option) end context J1_heap_base begin lemmas \red1_Val_simps [simp] = \red1r_Val \red1t_Val \reds1r_map_Val \reds1t_map_Val end context J1_JVM_conf_read begin lemma assumes wf: "wf_J1_prog P" and hconf: "hconf h" "preallocated h" and tconf: "P,h \ t \t" shows red1_simulates_exec_instr: "\ P, E, h \ (e, xs) \ (stk, loc, pc, xcp); exec_move_d P t E h (stk, loc, pc, xcp) ta h' (stk', loc', pc', xcp'); n + max_vars e \ length xs; bsok E n; P,h \ stk [:\] ST; conf_xcp' (compP2 P) h xcp \ \ \e'' xs''. P, E, h' \ (e'', xs'') \ (stk', loc', pc', xcp') \ (if \move2 (compP2 P) h stk E pc xcp then h' = h \ (if xcp' = None \ pc < pc' then \red1r else \red1t) P t h (e, xs) (e'', xs'') else \ta' e' xs'. \red1r P t h (e, xs) (e', xs') \ True,P,t \1 \e', (h, xs')\ -ta'\ \e'', (h', xs'')\ \ ta_bisim wbisim1 (extTA2J1 P ta') ta \ \ \move1 P h e' \ (call1 e = None \ no_call2 E pc \ e' = e \ xs' = xs))" (is "\ _; ?exec E stk loc pc xcp stk' loc' pc' xcp'; _; _; _; _ \ \ \e'' xs''. P, E, h' \ (e'', xs'') \ (stk', loc', pc', xcp') \ ?red e xs e'' xs'' E stk pc pc' xcp xcp'") and reds1_simulates_exec_instr: "\ P, Es, h \ (es, xs) [\] (stk, loc, pc, xcp); exec_moves_d P t Es h (stk, loc, pc, xcp) ta h' (stk', loc', pc', xcp'); n + max_varss es \ length xs; bsoks Es n; P,h \ stk [:\] ST; conf_xcp' (compP2 P) h xcp \ \ \es'' xs''. P, Es, h' \ (es'', xs'') [\] (stk', loc', pc', xcp') \ (if \moves2 (compP2 P) h stk Es pc xcp then h' = h \ (if xcp' = None \ pc < pc' then \reds1r else \reds1t) P t h (es, xs) (es'', xs'') else \ta' es' xs'. \reds1r P t h (es, xs) (es', xs') \ True,P,t \1 \es', (h, xs')\ [-ta'\] \es'', (h', xs'')\ \ ta_bisim wbisim1 (extTA2J1 P ta') ta \ \ \moves1 P h es' \ (calls1 es = None \ no_calls2 Es pc \ es' = es \ xs' = xs))" (is "\ _; ?execs Es stk loc pc xcp stk' loc' pc' xcp'; _; _; _; _ \ \ \es'' xs''. P, Es, h' \ (es'', xs'') [\] (stk', loc', pc', xcp') \ ?reds es xs es'' xs'' Es stk pc pc' xcp xcp'") proof(induction E n e xs stk loc pc xcp and Es n es xs stk loc pc xcp arbitrary: stk' loc' pc' xcp' ST and stk' loc' pc' xcp' ST rule: bisim1_bisims1_inducts_split) case (bisim1Val2 e n v xs) from \?exec e [v] xs (length (compE2 e)) None stk' loc' pc' xcp'\ have False by(auto dest: exec_meth_length_compE2D simp add: exec_move_def) thus ?case .. next case (bisim1New C' n xs) note exec = \exec_move_d P t (new C') h ([], xs, 0, None) ta h' (stk', loc', pc', xcp')\ have \: "\ \move2 (compP2 P) h [] (new C') 0 None" "\ \move1 P h (new C')" by(auto simp add: \move2_iff) show ?case proof(cases "allocate h (Class_type C') = {}") case True have "P,new C',h' \ (THROW OutOfMemory, xs) \ ([], xs, 0, \addr_of_sys_xcpt OutOfMemory\)" by(rule bisim1NewThrow) with exec \ True show ?thesis by(fastforce intro: Red1NewFail elim!: exec_meth.cases simp add: exec_move_def) next case False have "\a h'. P,new C',h' \ (addr a, xs) \ ([Addr a], xs, length (compE2 (new C')), None)" by(rule bisim1Val2) auto thus ?thesis using exec False \ apply(simp add: exec_move_def) apply(erule exec_meth.cases) apply simp_all apply clarsimp apply(auto intro!: Red1New exI simp add: ta_bisim_def) done qed next case (bisim1NewThrow C' n xs) from \?exec (new C') [] xs 0 \addr_of_sys_xcpt OutOfMemory\ stk' loc' pc' xcp'\ have False by(auto elim: exec_meth.cases simp add: exec_move_def) thus ?case .. next case (bisim1NewArray e n e' xs stk loc pc xcp U) note IH = bisim1NewArray.IH(2) note exec = \?exec (newA U\e\) stk loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (newA U\e'\) \ length xs\ note bsok = \bsok (newA U\e\) n\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have exec': "?exec e stk loc pc xcp stk' loc' pc' xcp'" by(auto simp add: exec_move_newArray) from True have "\move2 (compP2 P) h stk (newA U\e\) pc xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff) moreover have "no_call2 e pc \ no_call2 (newA U\e\) pc" by(simp add: no_call2_def) ultimately show ?thesis using IH[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bsok by(fastforce intro: bisim1_bisims1.bisim1NewArray New1ArrayRed elim!: NewArray_\red1r_xt NewArray_\red1t_xt) next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (newA U\e'\, xs) (newA U\Val v\, loc)" by(rule NewArray_\red1r_xt) moreover have \: "\ \move2 (compP2 P) h [v] (newA U\e\) pc None" by(simp add: \move2_iff) moreover have "\ \move1 P h (newA U\Val v\)" by auto moreover from exec stk xcp obtain I where [simp]: "v = Intg I" by(auto elim!: exec_meth.cases simp add: exec_move_def) have "\ta' e''. P,newA U\e\,h' \ (e'',loc) \ (stk', loc', pc', xcp') \ True,P,t \1 \newA U\Val v\,(h, loc)\ -ta'\ \e'',(h', loc)\ \ ta_bisim wbisim1 (extTA2J1 P ta') ta" proof(cases "I a h'. P,newA U\e\,h' \ (addr a, loc) \ ([Addr a], loc, length (compE2 (newA U\e\)), None)" by(rule bisim1Val2) simp with False \\ I exec stk xcp show ?thesis apply(simp add: exec_move_def) apply(erule exec_meth.cases) apply simp_all apply clarsimp apply(auto intro!: Red1NewArray exI simp add: ta_bisim_def) done qed qed moreover have "no_call2 (newA U\e\) pc" by(simp add: no_call2_def) ultimately show ?thesis using exec stk xcp by fastforce qed next case (bisim1NewArrayThrow e n a xs stk loc pc U) note exec = \?exec (newA U\e\) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by (auto elim!: exec_meth.cases simp add: exec_move_def) thus ?case .. next case (bisim1NewArrayFail e n U a xs v) note exec = \?exec (newA U\e\) [v] xs (length (compE2 e)) \a\ stk' loc' pc' xcp'\ hence False by(auto elim!: exec_meth.cases dest: match_ex_table_pc_length_compE2 simp add: exec_move_def) thus ?case .. next case (bisim1Cast e n e' xs stk loc pc xcp U) note IH = bisim1Cast.IH(2) note exec = \?exec (Cast U e) stk loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (Cast U e') \ length xs\ note ST = \P,h \ stk [:\] ST\ note bsok = \bsok (Cast U e) n\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have exec': "?exec e stk loc pc xcp stk' loc' pc' xcp'" by(auto simp add: exec_move_Cast) from True have "\move2 (compP2 P) h stk (Cast U e) pc xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff) moreover have "no_call2 e pc \ no_call2 (Cast U e) pc" by(simp add: no_call2_def) ultimately show ?thesis using IH[OF exec' _ _ ST \conf_xcp' (compP2 P) h xcp\] len bsok by(fastforce intro: bisim1_bisims1.bisim1Cast Cast1Red elim!: Cast_\red1r_xt Cast_\red1t_xt) next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (Cast U e', xs) (Cast U (Val v), loc)" by(rule Cast_\red1r_xt) also from exec have [simp]: "h' = h" "ta = \" by(auto simp add: exec_move_def elim!: exec_meth.cases split: if_split_asm) have "\e''. P,Cast U e,h \ (e'',loc) \ (stk', loc', pc', xcp') \ True,P,t \1 \Cast U (Val v),(h, loc)\ -\\ \e'',(h, loc)\" proof(cases "P \ the (typeof\<^bsub>h\<^esub> v) \ U") case False with exec stk xcp bsok ST show ?thesis by(fastforce simp add: compP2_def exec_move_def exec_meth_instr list_all2_Cons1 conf_def intro: bisim1CastFail Red1CastFail) next case True have "P,Cast U e,h \ (Val v, loc) \ ([v], loc, length (compE2 (Cast U e)), None)" by(rule bisim1Val2) simp with exec stk xcp ST True show ?thesis by(fastforce simp add: compP2_def exec_move_def exec_meth_instr list_all2_Cons1 conf_def intro: Red1Cast) qed then obtain e'' where bisim': "P,Cast U e,h \ (e'',loc) \ (stk', loc', pc', xcp')" and red: "True,P,t \1 \Cast U (Val v),(h, loc)\ -\\ \e'',(h, loc)\" by blast have "\move1 P h (Cast U (Val v))" by(rule \move1CastRed) with red have "\red1t P t h (Cast U (Val v), loc) (e'', loc)" by(auto intro: \red1t_1step) also have \: "\move2 (compP2 P) h [v] (Cast U e) pc None" by(simp add: \move2_iff) moreover have "no_call2 (Cast U e) pc" by(simp add: no_call2_def) ultimately show ?thesis using exec stk xcp bisim' by fastforce qed next case (bisim1CastThrow e n a xs stk loc pc U) note exec = \?exec (Cast U e) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by (auto elim!: exec_meth.cases simp add: exec_move_def) thus ?case .. next case (bisim1CastFail e n U xs v) note exec = \?exec (Cast U e) [v] xs (length (compE2 e)) \addr_of_sys_xcpt ClassCast\ stk' loc' pc' xcp'\ hence False by(auto elim!: exec_meth.cases dest: match_ex_table_pc_length_compE2 simp add: exec_move_def) thus ?case .. next case (bisim1InstanceOf e n e' xs stk loc pc xcp U) note IH = bisim1InstanceOf.IH(2) note exec = \?exec (e instanceof U) stk loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (e' instanceof U) \ length xs\ note ST = \P,h \ stk [:\] ST\ note bsok = \bsok (e instanceof U) n\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have exec': "?exec e stk loc pc xcp stk' loc' pc' xcp'" by(auto simp add: exec_move_InstanceOf) from True have "\move2 (compP2 P) h stk (e instanceof U) pc xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff) moreover have "no_call2 e pc \ no_call2 (e instanceof U) pc" by(simp add: no_call2_def) ultimately show ?thesis using IH[OF exec' _ _ ST \conf_xcp' (compP2 P) h xcp\] len bsok by(fastforce intro: bisim1_bisims1.bisim1InstanceOf InstanceOf1Red elim!: InstanceOf_\red1r_xt InstanceOf_\red1t_xt) next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (e' instanceof U, xs) ((Val v) instanceof U, loc)" by(rule InstanceOf_\red1r_xt) also let ?v = "Bool (v \ Null \ P \ the (typeof\<^bsub>h\<^esub> v) \ U)" from exec ST stk xcp have [simp]: "h' = h" "ta = \" "xcp' = None" "loc' = loc" "stk' = [?v]" "pc' = Suc pc" by(auto simp add: exec_move_def list_all2_Cons1 conf_def compP2_def elim!: exec_meth.cases split: if_split_asm) have bisim': "P,e instanceof U,h \ (Val ?v, loc) \ ([?v], loc, length (compE2 (e instanceof U)), None)" by(rule bisim1Val2) simp from exec stk xcp ST have red: "True,P,t \1 \(Val v) instanceof U,(h, loc)\ -\\ \Val ?v ,(h, loc)\" by(auto simp add: compP2_def exec_move_def exec_meth_instr list_all2_Cons1 conf_def intro: Red1InstanceOf) have "\move1 P h ((Val v) instanceof U)" by(rule \move1InstanceOfRed) with red have "\red1t P t h ((Val v) instanceof U, loc) (Val ?v, loc)" by(auto intro: \red1t_1step) also have \: "\move2 (compP2 P) h [v] (e instanceof U) pc None" by(simp add: \move2_iff) moreover have "no_call2 (e instanceof U) pc" by(simp add: no_call2_def) ultimately show ?thesis using exec stk xcp bisim' by(fastforce) qed next case (bisim1InstanceOfThrow e n a xs stk loc pc U) note exec = \?exec (e instanceof U) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by (auto elim!: exec_meth.cases simp add: exec_move_def) thus ?case .. next case (bisim1Val v n xs) from \?exec (Val v) [] xs 0 None stk' loc' pc' xcp'\ have "stk' = [v]" "loc' = xs" "h' = h" "pc' = length (compE2 (Val v))" "xcp' = None" by(auto elim: exec_meth.cases simp add: exec_move_def) moreover have "P,Val v,h \ (Val v, xs) \ ([v], xs, length (compE2 (Val v)), None)" by(rule bisim1Val2) simp moreover have "\move2 (compP2 P) h [] (Val v) 0 None" by(rule \move2Val) ultimately show ?case by(auto) next case (bisim1Var V n xs) note exec = \?exec (Var V) [] xs 0 None stk' loc' pc' xcp'\ moreover note len = \n + max_vars (Var V) \ length xs\ moreover have "\move2 (compP2 P) h [] (Var V) 0 None" "\move1 P h (Var V)" by(auto intro: \move1Var simp add: \move2_iff) moreover have "P,Var V,h \ (Val (xs ! V), xs) \ ([xs ! V], xs, length (compE2 (Var V)), None)" by(rule bisim1Val2) simp ultimately show ?case by(fastforce elim!: exec_meth.cases intro: Red1Var r_into_rtranclp simp add: exec_move_def) next case (bisim1BinOp1 e1 n e1' xs stk loc pc xcp e2 bop) note IH1 = bisim1BinOp1.IH(2) note IH2 = bisim1BinOp1.IH(4) note exec = \?exec (e1 \bop\ e2) stk loc pc xcp stk' loc' pc' xcp'\ note bisim1 = \P,e1,h \ (e1', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,e2,h \ (e2, loc) \ ([], loc, 0, None)\ note len = \n + max_vars (e1' \bop\ e2) \ length xs\ note ST = \P,h \ stk [:\] ST\ note bsok = \bsok (e1 \bop\ e2) n\ from bisim1 have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e1)") case True with exec have exec': "?exec e1 stk loc pc xcp stk' loc' pc' xcp'" by(auto simp add: exec_move_BinOp1) from True have \: "\move2 (compP2 P) h stk (e1 \bop\ e2) pc xcp = \move2 (compP2 P) h stk e1 pc xcp" by(simp add: \move2_iff) with IH1[OF exec' _ _ ST \conf_xcp' (compP2 P) h xcp\] bisim2 len bsok obtain e'' xs'' where bisim': "P,e1,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red: "?red e1' xs e'' xs'' e1 stk pc pc' xcp xcp'" by auto from bisim' have "P,e1 \bop\ e2,h' \ (e'' \bop\ e2, xs'') \ (stk', loc', pc', xcp')" by(rule bisim1_bisims1.bisim1BinOp1) moreover from True have "no_call2 (e1 \bop\ e2) pc = no_call2 e1 pc" by(simp add: no_call2_def) ultimately show ?thesis using red \ by(fastforce intro: Bin1OpRed1 elim!: BinOp_\red1r_xt1 BinOp_\red1t_xt1) next case False with pc have pc: "pc = length (compE2 e1)" by auto with bisim1 obtain v where e1': "is_val e1' \ e1' = Val v" and stk: "stk = [v]" and xcp: "xcp = None" and call: "call1 e1' = None" by(auto dest: bisim1_pc_length_compE2D) with bisim1 pc len bsok have rede1': "\red1r P t h (e1', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence rede1'': "\red1r P t h (e1' \bop\ e2, xs) (Val v \bop\ e2, loc)" by(rule BinOp_\red1r_xt1) moreover from pc exec stk xcp have "exec_meth_d (compP2 P) (compE2 (e1 \bop\ e2)) (compxE2 (e1 \bop\ e2) 0 0) t h ([] @ [v], loc, length (compE2 e1) + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs compxE2_stack_xlift_convs exec_move_def) hence exec': "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v]) (compxE2 e2 0 0)) t h ([] @ [v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 e1), xcp')" and pc': "pc' \ length (compE2 e1)" by(safe dest!: BinOp_exec2D)simp_all then obtain PC' where PC': "pc' = length (compE2 e1) + PC'" by -(rule that[where PC'35="pc' - length (compE2 e1)"], simp) from exec' bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_move_d P t e2 h ([], loc, 0, None) ta h' (stk'', loc', pc' - length (compE2 e1), xcp')" by(unfold exec_move_def)(drule (1) exec_meth_stk_split, auto) with pc xcp have \: "\move2 (compP2 P) h [v] (e1 \bop\ e2) (length (compE2 e1)) None = \move2 (compP2 P) h [] e2 0 None" using \instr_stk_drop_exec_move[where stk="[]" and vs = "[v]"] by(simp add: \move2_iff \instr_stk_drop_exec_move) from bisim1 have "length xs = length loc" by(rule bisim1_length_xs) with IH2[OF exec'', of "[]"] len bsok obtain e'' xs'' where bisim': "P,e2,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 e1), xcp')" and red: "?red e2 loc e'' xs'' e2 [] 0 (pc' - length (compE2 e1)) None xcp'" by auto from bisim' have "P,e1 \bop\ e2,h' \ (Val v \bop\ e'', xs'') \ (stk'' @ [v], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule bisim1_bisims1.bisim1BinOp2) moreover from red \ have "?red (Val v \bop\ e2) loc (Val v \bop\ e'') xs'' (e1 \bop\ e2) [v] (length (compE2 e1)) pc' None xcp'" by(fastforce intro: Bin1OpRed2 elim!: BinOp_\red1r_xt2 BinOp_\red1t_xt2 simp add: no_call2_def) moreover have "no_call2 (e1 \bop\ e2) (length (compE2 e1))" by(simp add: no_call2_def) ultimately show ?thesis using \ stk' pc xcp pc' PC' bisim1 bisim2 e1' stk call by(fastforce elim!: rtranclp_trans) qed next case (bisim1BinOp2 e2 n e2' xs stk loc pc xcp e1 bop v1) note IH2 = bisim1BinOp2.IH(2) note exec = \?exec (e1 \bop\ e2) (stk @ [v1]) loc (length (compE2 e1) + pc) xcp stk' loc' pc' xcp'\ note bisim1 = \P,e1,h \ (e1, xs) \ ([], xs, 0, None)\ note bisim2 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (Val v1 \bop\ e2') \ length xs\ note bsok = \bsok (e1 \bop\ e2) n\ note ST = \P,h \ stk @ [v1] [:\] ST\ then obtain ST2 T where "ST = ST2 @ [T]" "P,h \ stk [:\] ST2" "P,h \ v1 :\ T" by(auto simp add: list_all2_append1 length_Suc_conv) from bisim2 have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e2)") case True with exec have exec': "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v1]) (compxE2 e2 0 0)) t h (stk @ [v1], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 e1), xcp')" and pc': "pc' \ length (compE2 e1)" by(unfold exec_move_def)(safe dest!: BinOp_exec2D) from exec' bisim2 obtain stk'' where stk': "stk' = stk'' @ [v1]" and exec'': "exec_move_d P t e2 h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e1), xcp')" by -(drule (1) exec_meth_stk_split, auto simp add: exec_move_def) with True have \: "\move2 (compP2 P) h (stk @ [v1]) (e1 \bop\ e2) (length (compE2 e1) + pc) xcp = \move2 (compP2 P) h stk e2 pc xcp" by(auto simp add: \move2_iff \instr_stk_drop_exec_move) from IH2[OF exec'' _ _ \P,h \ stk [:\] ST2\ \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,e2,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 e1), xcp')" and red: "?red e2' xs e'' xs'' e2 stk pc (pc' - length (compE2 e1)) xcp xcp'" by auto from bisim' have "P,e1 \bop\ e2,h' \ (Val v1 \bop\ e'', xs'') \ (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule bisim1_bisims1.bisim1BinOp2) with red \ stk' pc' True show ?thesis by(fastforce intro: Bin1OpRed2 elim!: BinOp_\red1r_xt2 BinOp_\red1t_xt2 split: if_split_asm simp add: no_call2_def) next case False with pc have [simp]: "pc = length (compE2 e2)" by simp with bisim2 obtain v2 where e2': "is_val e2' \ e2' = Val v2" and stk: "stk = [v2]" and xcp: "xcp = None" and call: "call1 e2' = None" by(auto dest: bisim1_pc_length_compE2D) with bisim2 pc len bsok have red: "\red1r P t h (e2', xs) (Val v2, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence red1: "\red1r P t h (Val v1 \bop\ e2', xs) (Val v1 \bop\ Val v2, loc)" by(rule BinOp_\red1r_xt2) show ?thesis proof(cases "the (binop bop v1 v2)") case (Inl v) note red1 also from exec xcp ST stk Inl have "\red1r P t h (Val v1 \bop\ Val v2, loc) (Val v, loc)" by(force simp add: exec_move_def exec_meth_instr list_all2_Cons1 conf_def compP2_def dest: binop_progress intro: r_into_rtranclp Red1BinOp \move1BinOp) also have \: "\move2 (compP2 P) h [v2, v1] (e1 \bop\ e2) (length (compE2 e1) + length (compE2 e2)) None" by(simp add: \move2_iff) moreover have "P,e1 \bop\ e2,h \ (Val v, loc) \ ([v], loc, length (compE2 (e1 \bop\ e2)), None)" by(rule bisim1Val2) simp ultimately show ?thesis using exec xcp stk call Inl by(auto simp add: exec_move_def exec_meth_instr) next case (Inr a) note red1 also from exec xcp ST stk Inr have "\red1r P t h (Val v1 \bop\ Val v2, loc) (Throw a, loc)" by(force simp add: exec_move_def exec_meth_instr list_all2_Cons1 conf_def compP2_def dest: binop_progress intro: r_into_rtranclp Red1BinOpFail \move1BinOp) also have \: "\move2 (compP2 P) h [v2, v1] (e1 \bop\ e2) (length (compE2 e1) + length (compE2 e2)) None" by(simp add: \move2_iff) moreover have "P,e1 \bop\ e2,h \ (Throw a, loc) \ ([v2, v1], loc, length (compE2 e1) + length (compE2 e2), \a\)" by(rule bisim1BinOpThrow) ultimately show ?thesis using exec xcp stk call Inr by(auto simp add: exec_move_def exec_meth_instr) qed qed next case (bisim1BinOpThrow1 e1 n a xs stk loc pc e2 bop) note exec = \?exec (e1 \bop\ e2) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,e1,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 e1)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e1 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: match_ex_table_not_pcs_None exec_move_def) thus ?case .. next case (bisim1BinOpThrow2 e2 n a xs stk loc pc e1 bop v1) note exec = \?exec (e1 \bop\ e2) (stk @ [v1]) loc (length (compE2 e1) + pc) \a\ stk' loc' pc' xcp'\ note bisim2 = \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ hence "match_ex_table (compP2 P) (cname_of h a) (length (compE2 e1) + pc) (compxE2 e2 (length (compE2 e1)) 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec have False apply(auto elim!: exec_meth.cases simp add: match_ex_table_append_not_pcs exec_move_def) apply(auto simp only: compxE2_size_convs compxE2_stack_xlift_convs match_ex_table_stack_xlift_eq_Some_conv) done thus ?case .. next case (bisim1BinOpThrow e1 n e2 bop a xs v1 v2) note \?exec (e1 \bop\ e2) [v1, v2] xs (length (compE2 e1) + length (compE2 e2)) \a\ stk' loc' pc' xcp'\ hence False by(auto elim!: exec_meth.cases simp add: match_ex_table_append_not_pcs compxE2_size_convs exec_move_def dest!: match_ex_table_shift_pcD match_ex_table_pc_length_compE2) thus ?case .. next case (bisim1LAss1 e n e' xs stk loc pc xcp V) note IH = bisim1LAss1.IH(2) note exec = \?exec (V := e) stk loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (V := e') \ length xs\ note bsok = \bsok (V:=e) n\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have exec': "?exec e stk loc pc xcp stk' loc' pc' xcp'" by(auto simp add: exec_move_LAss) from True have "\move2 (compP2 P) h stk (V := e) pc xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff) with IH[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bsok show ?thesis by(fastforce intro: bisim1_bisims1.bisim1LAss1 LAss1Red elim!: LAss_\red1r LAss_\red1t simp add: no_call2_def) next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (V := e', xs) (V := Val v, loc)" by(rule LAss_\red1r) also have "\move1 P h (V := Val v)" by(rule \move1LAssRed) with exec stk xcp have "\red1r P t h (V := Val v, loc) (unit, loc[V := v])" by(auto intro!: r_into_rtranclp Red1LAss simp add: exec_move_def elim!: exec_meth.cases) also have \: "\move2 (compP2 P) h [v] (V := e) pc None" by(simp add: \move2_iff) moreover have "P,(V := e),h \ (unit, loc[V := v]) \ ([], loc[V := v], Suc (length (compE2 e)), None)" by(rule bisim1LAss2) ultimately show ?thesis using exec stk xcp by(fastforce elim!: exec_meth.cases simp add: exec_move_def) qed next case (bisim1LAss2 e n V xs) note bisim = \P,e,h \ (e, xs) \ ([], xs, 0, None)\ note exec = \?exec (V := e) [] xs (Suc (length (compE2 e))) None stk' loc' pc' xcp'\ hence "stk' = [Unit]" "loc' = xs" "pc' = length (compE2 (V := e))" "xcp' = None" "h' = h" by(auto elim!: exec_meth.cases simp add: exec_move_def) moreover have "\move2 (compP2 P) h [] (V := e) (Suc (length (compE2 e))) None" by(simp add: \move2_iff) moreover have "P,V:=e,h' \ (unit, xs) \ ([Unit], xs, length (compE2 (V := e)), None)" by(rule bisim1Val2) simp ultimately show ?case by(auto) next case (bisim1LAssThrow e n a xs stk loc pc V) note exec = \?exec (V := e) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by (auto elim!: exec_meth.cases simp add: exec_move_def) thus ?case .. next case (bisim1AAcc1 a n a' xs stk loc pc xcp i) note IH1 = bisim1AAcc1.IH(2) note IH2 = bisim1AAcc1.IH(4) note exec = \?exec (a\i\) stk loc pc xcp stk' loc' pc' xcp'\ note bisim1 = \P,a,h \ (a', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,i,h \ (i, loc) \ ([], loc, 0, None)\ note len = \n + max_vars (a'\i\) \ length xs\ note bsok = \bsok (a\i\) n\ from bisim1 have pc: "pc \ length (compE2 a)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 a)") case True with exec have exec': "?exec a stk loc pc xcp stk' loc' pc' xcp'" by(auto simp add: exec_move_AAcc1) from True have \: "\move2 (compP2 P) h stk (a\i\) pc xcp = \move2 (compP2 P) h stk a pc xcp" by(auto intro: \move2AAcc1) with IH1[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,a,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red: "?red a' xs e'' xs'' a stk pc pc' xcp xcp'" by auto from bisim' have "P,a\i\,h' \ (e''\i\, xs'') \ (stk', loc', pc', xcp')" by(rule bisim1_bisims1.bisim1AAcc1) moreover from True have "no_call2 (a\i\) pc = no_call2 a pc" by(simp add: no_call2_def) ultimately show ?thesis using red \ by(fastforce intro: AAcc1Red1 elim!: AAcc_\red1r_xt1 AAcc_\red1t_xt1) next case False with pc have pc: "pc = length (compE2 a)" by auto with bisim1 obtain v where a': "is_val a' \ a' = Val v" and stk: "stk = [v]" and xcp: "xcp = None" and call: "call1 a' = None" by(auto dest: bisim1_pc_length_compE2D) with bisim1 pc len bsok have rede1': "\red1r P t h (a', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (a'\i\, xs) (Val v\i\, loc)" by(rule AAcc_\red1r_xt1) moreover from pc exec stk xcp have exec': "exec_meth_d (compP2 P) (compE2 a @ compE2 i @ [ALoad]) (compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v]) (compxE2 i 0 0))) t h ([] @ [v], loc, length (compE2 a) + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs compxE2_stack_xlift_convs exec_move_def) hence "exec_meth_d (compP2 P) (compE2 i @ [ALoad]) (stack_xlift (length [v]) (compxE2 i 0 0)) t h ([] @ [v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v]) (compxE2 i 0 0)) t h ([] @ [v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_take) simp with bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_move_d P t i h ([], loc, 0, None) ta h' (stk'', loc', pc' - length (compE2 a), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with pc xcp have \: "\move2 (compP2 P) h ([] @ [v]) (a\i\) (length (compE2 a)) None = \move2 (compP2 P) h [] i 0 None" using \instr_stk_drop_exec_move[where stk="[]" and vs="[v]"] by(auto simp add: \move2_iff \instr_stk_drop_exec_move) from bisim1 have "length xs = length loc" by(rule bisim1_length_xs) with IH2[OF exec''] len bsok obtain e'' xs'' where bisim': "P,i,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 a), xcp')" and red: "?red i loc e'' xs'' i [] 0 (pc' - length (compE2 a)) None xcp'" by(fastforce) from bisim' have "P,a\i\,h' \ (Val v\e''\, xs'') \ (stk'' @ [v], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" by(rule bisim1_bisims1.bisim1AAcc2) moreover from red \ have "?red (Val v\i\) loc (Val v\e''\) xs'' (a\i\) [v] (length (compE2 a)) pc' None xcp'" by(fastforce intro: AAcc1Red2 elim!: AAcc_\red1r_xt2 AAcc_\red1t_xt2 split: if_split_asm simp add: no_call2_def) moreover from exec' have "pc' \ length (compE2 a)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 (a\i\) pc" using pc by(simp add: no_call2_def) ultimately show ?thesis using \ stk' pc xcp stk call by(fastforce elim!: rtranclp_trans)+ qed next case (bisim1AAcc2 i n i' xs stk loc pc xcp a v) note IH2 = bisim1AAcc2.IH(2) note exec = \?exec (a\i\) (stk @ [v]) loc (length (compE2 a) + pc) xcp stk' loc' pc' xcp'\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (Val v\i'\) \ length xs\ note bsok = \bsok (a\i\) n\ from bisim2 have pc: "pc \ length (compE2 i)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 i)") case True from exec have exec': "exec_meth_d (compP2 P) (compE2 a @ compE2 i @ [ALoad]) (compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v]) (compxE2 i 0 0))) t h (stk @ [v], loc, length (compE2 a) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 exec_move_def) hence "exec_meth_d (compP2 P) (compE2 i @ [ALoad]) (stack_xlift (length [v]) (compxE2 i 0 0)) t h (stk @ [v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v]) (compxE2 i 0 0)) t h (stk @ [v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a), xcp')" using True by(rule exec_meth_take) with bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_move_d P t i h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 a), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with True have \: "\move2 (compP2 P) h (stk @ [v]) (a\i\) (length (compE2 a) + pc) xcp = \move2 (compP2 P) h stk i pc xcp" by(auto simp add: \move2_iff \instr_stk_drop_exec_move) moreover from \P,h \ stk @ [v] [:\] ST\ obtain ST2 where "P,h \ stk [:\] ST2" by(auto simp add: list_all2_append1) from IH2[OF exec'' _ _ this \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,i,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 a), xcp')" and red: "?red i' xs e'' xs'' i stk pc (pc' - length (compE2 a)) xcp xcp'" by auto from bisim' have "P,a\i\,h' \ (Val v\e''\, xs'') \ (stk'' @ [v], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" by(rule bisim1_bisims1.bisim1AAcc2) moreover from exec' have "pc' \ length (compE2 a)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 i pc \ no_call2 (a\i\) (length (compE2 a) + pc)" by(simp add: no_call2_def) ultimately show ?thesis using red \ stk' True by(fastforce intro: AAcc1Red2 elim!: AAcc_\red1r_xt2 AAcc_\red1t_xt2 split: if_split_asm) next case False with pc have [simp]: "pc = length (compE2 i)" by simp with bisim2 obtain v2 where i': "is_val i' \ i' = Val v2" and stk: "stk = [v2]" and xcp: "xcp = None" and call: "call1 i' = None" by(auto dest: bisim1_pc_length_compE2D) with bisim2 pc len bsok have red: "\red1r P t h (i', xs) (Val v2, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (Val v\i'\, xs) (Val v\Val v2\, loc)" by(rule AAcc_\red1r_xt2) moreover have \: "\ \move2 (compP2 P) h [v2, v] (a\i\) (length (compE2 a) + length (compE2 i)) None" by(simp add: \move2_iff) moreover have "\ta' e''. P,a\i\,h' \ (e'',loc) \ (stk', loc', pc', xcp') \ True,P,t \1 \Val v\Val v2\, (h, loc)\ -ta'\ \e'',(h', loc)\ \ ta_bisim wbisim1 (extTA2J1 P ta') ta" proof(cases "v = Null") case True with exec stk xcp show ?thesis by(fastforce elim!: exec_meth.cases simp add: exec_move_def intro: bisim1AAccFail Red1AAccNull) next case False with exec xcp stk obtain U el A len I where [simp]: "v = Addr A" and hA: "typeof_addr h A = \Array_type U len\" and [simp]: "v2 = Intg I" by(auto simp add: exec_move_def exec_meth_instr is_Ref_def conf_def split: if_split_asm) show ?thesis proof(cases "0 <=s I \ sint I < int len") case True hence "\ I i\,h' \ (Val v3,loc) \ ([v3], loc, length (compE2 (a\i\)), None)" by(rule bisim1Val2) simp ultimately show ?thesis using exec stk xcp True hA by(fastforce elim!: exec_meth.cases intro: Red1AAcc simp add: exec_move_def ta_upd_simps ta_bisim_def split: if_split_asm) next case False with exec stk xcp hA show ?thesis by(fastforce elim!: exec_meth.cases simp add: is_Ref_def exec_move_def intro: bisim1AAccFail Red1AAccBounds split: if_split_asm) qed qed ultimately show ?thesis using exec xcp stk call by fastforce qed next case (bisim1AAccThrow1 A n a xs stk loc pc i) note exec = \?exec (A\i\) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,A,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 A)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 A 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: match_ex_table_not_pcs_None exec_move_def) thus ?case .. next case (bisim1AAccThrow2 i n a xs stk loc pc A v) note exec = \?exec (A\i\) (stk @ [v]) loc (length (compE2 A) + pc) \a\ stk' loc' pc' xcp'\ note bisim2 = \P,i,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ hence "match_ex_table (compP2 P) (cname_of h a) (length (compE2 A) + pc) (compxE2 i (length (compE2 A)) 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec have False apply(auto elim!: exec_meth.cases simp add: match_ex_table_append_not_pcs exec_move_def) apply(auto simp only: compxE2_size_convs compxE2_stack_xlift_convs match_ex_table_stack_xlift_eq_Some_conv) done thus ?case .. next case (bisim1AAccFail a n i ad xs v v') note \?exec (a\i\) [v, v'] xs (length (compE2 a) + length (compE2 i)) \ad\ stk' loc' pc' xcp'\ hence False by(auto elim!: exec_meth.cases simp add: match_ex_table_append_not_pcs compxE2_size_convs exec_move_def dest!: match_ex_table_shift_pcD match_ex_table_pc_length_compE2) thus ?case .. next case (bisim1AAss1 a n a' xs stk loc pc xcp i e) note IH1 = bisim1AAss1.IH(2) note IH2 = bisim1AAss1.IH(4) note exec = \?exec (a\i\ := e) stk loc pc xcp stk' loc' pc' xcp'\ note bisim1 = \P,a,h \ (a', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,i,h \ (i, loc) \ ([], loc, 0, None)\ note len = \n + max_vars (a'\i\ := e) \ length xs\ note bsok = \bsok (a\i\ := e) n\ from bisim1 have pc: "pc \ length (compE2 a)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 a)") case True with exec have exec': "?exec a stk loc pc xcp stk' loc' pc' xcp'" by(simp add: exec_move_AAss1) from True have \: "\move2 (compP2 P) h stk (a\i\ := e) pc xcp = \move2 (compP2 P) h stk a pc xcp" by(simp add: \move2_iff) with IH1[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,a,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red: "?red a' xs e'' xs'' a stk pc pc' xcp xcp'" by auto from bisim' have "P,a\i\ := e,h' \ (e''\i\ := e, xs'') \ (stk', loc', pc', xcp')" by(rule bisim1_bisims1.bisim1AAss1) moreover from True have "no_call2 (a\i\ := e) pc = no_call2 a pc" by(simp add: no_call2_def) ultimately show ?thesis using red \ by(fastforce intro: AAss1Red1 elim!: AAss_\red1r_xt1 AAss_\red1t_xt1) next case False with pc have pc: "pc = length (compE2 a)" by auto with bisim1 obtain v where a': "is_val a' \ a' = Val v" and stk: "stk = [v]" and xcp: "xcp = None" and call: "call1 a' = None" by(auto dest: bisim1_pc_length_compE2D) with bisim1 pc len bsok have rede1': "\red1r P t h (a', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (a'\i\ := e, xs) (Val v\i\ := e, loc)" by(rule AAss_\red1r_xt1) moreover from pc exec stk xcp have exec': "exec_meth_d (compP2 P) (compE2 a @ compE2 i @ compE2 e @ [AStore, Push Unit]) (compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v]) (compxE2 i 0 0) @ shift (length (compE2 i)) (compxE2 e 0 (Suc (Suc 0))))) t h ([] @ [v], loc, length (compE2 a) + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs compxE2_stack_xlift_convs exec_move_def) hence "exec_meth_d (compP2 P) (compE2 i @ compE2 e @ [AStore, Push Unit]) (stack_xlift (length [v]) (compxE2 i 0 0) @ shift (length (compE2 i)) (compxE2 e 0 (Suc (Suc 0)))) t h ([] @ [v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v]) (compxE2 i 0 0)) t h ([] @ [v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_take_xt) simp with bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_move_d P t i h ([], loc, 0, None) ta h' (stk'', loc', pc' - length (compE2 a), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with pc xcp have \: "\move2 (compP2 P) h [v] (a\i\:= e) (length (compE2 a)) None = \move2 (compP2 P) h [] i 0 None" using \instr_stk_drop_exec_move[where stk="[]" and vs="[v]"] by(auto simp add: \move2_iff) from bisim1 have "length xs = length loc" by(rule bisim1_length_xs) with IH2[OF exec''] len bsok obtain e'' xs'' where bisim': "P,i,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 a), xcp')" and red: "?red i loc e'' xs'' i [] 0 (pc' - length (compE2 a)) None xcp'" by fastforce from bisim' have "P,a\i\ := e,h' \ (Val v\e''\ := e, xs'') \ (stk'' @ [v], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" by(rule bisim1_bisims1.bisim1AAss2) moreover from red \ have "?red (Val v\i\ := e) loc (Val v\e''\ := e) xs'' (a\i\ := e) [v] (length (compE2 a)) pc' None xcp'" by(fastforce intro: AAss1Red2 elim!: AAss_\red1r_xt2 AAss_\red1t_xt2 split: if_split_asm simp add: no_call2_def) moreover from exec' have "pc' \ length (compE2 a)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 (a\i\ := e) pc" using pc by(simp add: no_call2_def) ultimately show ?thesis using \ stk' pc xcp stk by(fastforce elim!: rtranclp_trans) qed next case (bisim1AAss2 i n i' xs stk loc pc xcp a e v) note IH2 = bisim1AAss2.IH(2) note IH3 = bisim1AAss2.IH(6) note exec = \?exec (a\i\ := e) (stk @ [v]) loc (length (compE2 a) + pc) xcp stk' loc' pc' xcp'\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, xcp)\ note bisim3 = \P,e,h \ (e, loc) \ ([], loc, 0, None)\ note len = \n + max_vars (Val v\i'\ := e) \ length xs\ note bsok = \bsok (a\i\ := e) n\ from bisim2 have pc: "pc \ length (compE2 i)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 i)") case True from exec have exec': "exec_meth_d (compP2 P) (compE2 a @ compE2 i @ compE2 e @ [AStore, Push Unit]) (compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v]) (compxE2 i 0 0) @ shift (length (compE2 i)) (compxE2 e 0 (Suc (Suc 0))))) t h (stk @ [v], loc, length (compE2 a) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps exec_move_def) hence "exec_meth_d (compP2 P) (compE2 i @ compE2 e @ [AStore, Push Unit]) (stack_xlift (length [v]) (compxE2 i 0 0) @ shift (length (compE2 i)) (compxE2 e 0 (Suc (Suc 0)))) t h (stk @ [v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v]) (compxE2 i 0 0)) t h (stk @ [v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a), xcp')" using True by(rule exec_meth_take_xt) with bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_move_d P t i h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 a), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with True have \: "\move2 (compP2 P) h (stk @ [v]) (a\i\ := e) (length (compE2 a) + pc) xcp = \move2 (compP2 P) h stk i pc xcp" by(auto simp add: \move2_iff \instr_stk_drop_exec_move) moreover from \P,h \ stk @ [v] [:\] ST\ obtain ST2 where "P,h \ stk [:\] ST2" by(auto simp add: list_all2_append1) from IH2[OF exec'' _ _ this \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,i,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 a), xcp')" and red: "?red i' xs e'' xs'' i stk pc (pc' - length (compE2 a)) xcp xcp'" by fastforce from bisim' have "P,a\i\ := e,h' \ (Val v\e''\ := e, xs'') \ (stk'' @ [v], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" by(rule bisim1_bisims1.bisim1AAss2) moreover from exec' have "pc' \ length (compE2 a)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 i pc \ no_call2 (a\i\ := e) (length (compE2 a) + pc)" by(simp add: no_call2_def) ultimately show ?thesis using red \ stk' True by(fastforce intro: AAss1Red2 elim!: AAss_\red1r_xt2 AAss_\red1t_xt2 split: if_split_asm) next case False with pc have [simp]: "pc = length (compE2 i)" by simp with bisim2 obtain v2 where i': "is_val i' \ i' = Val v2" and stk: "stk = [v2]" and xcp: "xcp = None" and call: "call1 i' = None" by(auto dest: bisim1_pc_length_compE2D) with bisim2 pc len bsok have red: "\red1r P t h (i', xs) (Val v2, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (Val v\i'\ := e, xs) (Val v\Val v2\ := e, loc)" by(rule AAss_\red1r_xt2) moreover from pc exec stk xcp have exec': "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e @ [AStore, Push Unit]) ((compxE2 a 0 0 @ compxE2 i (length (compE2 a)) (Suc 0)) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length [v2, v]) (compxE2 e 0 0))) t h ([] @ [v2, v], loc, length (compE2 a @ compE2 i) + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs compxE2_stack_xlift_convs exec_move_def) hence "exec_meth_d (compP2 P) (compE2 e @ [AStore, Push Unit]) (stack_xlift (length [v2, v]) (compxE2 e 0 0)) t h ([] @ [v2, v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 a @ compE2 i), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 e) (stack_xlift (length [v2, v]) (compxE2 e 0 0)) t h ([] @ [v2, v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 a @ compE2 i), xcp')" by(rule exec_meth_take) simp with bisim3 obtain stk'' where stk': "stk' = stk'' @ [v2, v]" and exec'': "exec_move_d P t e h ([], loc, 0, None) ta h' (stk'', loc', pc' - length (compE2 a @ compE2 i), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with pc xcp have \: "\move2 (compP2 P) h [v2, v] (a\i\:= e) (length (compE2 a) + length (compE2 i)) None = \move2 (compP2 P) h [] e 0 None" using \instr_stk_drop_exec_move[where stk="[]" and vs="[v2, v]"] by(simp add: \move2_iff) from bisim2 have "length xs = length loc" by(rule bisim1_length_xs) with IH3[OF exec'', of "[]"] len bsok obtain e'' xs'' where bisim': "P,e,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 a) - length (compE2 i), xcp')" and red: "?red e loc e'' xs'' e [] 0 (pc' - length (compE2 a) - length (compE2 i)) None xcp'" by auto (fastforce simp only: length_append diff_diff_left) from bisim' have "P,a\i\ := e,h' \ (Val v\Val v2\ := e'', xs'') \ (stk'' @ [v2, v], loc', length (compE2 a) + length (compE2 i) + (pc' - length (compE2 a) - length (compE2 i)), xcp')" by(rule bisim1_bisims1.bisim1AAss3) moreover from red \ have "?red (Val v\Val v2\ := e) loc (Val v\Val v2\ := e'') xs'' (a\i\ := e) [v2, v] (length (compE2 a) + length (compE2 i)) pc' None xcp'" by(fastforce intro: AAss1Red3 elim!: AAss_\red1r_xt3 AAss_\red1t_xt3 split: if_split_asm simp add: no_call2_def) moreover from exec' have "pc' \ length (compE2 a @ compE2 i)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 (a\i\ := e) (length (compE2 a) + pc)" by(simp add: no_call2_def) ultimately show ?thesis using \ stk' pc xcp stk by(fastforce elim!: rtranclp_trans) qed next case (bisim1AAss3 e n e' xs stk loc pc xcp a i v v') note IH3 = bisim1AAss3.IH(2) note exec = \?exec (a\i\ := e) (stk @ [v', v]) loc (length (compE2 a) + length (compE2 i) + pc) xcp stk' loc' pc' xcp'\ note bisim3 = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (Val v\Val v'\ := e') \ length xs\ note bsok = \bsok (a\i\ := e) n\ from \P,h \ stk @ [v', v] [:\] ST\ obtain T T' ST' where [simp]: "ST = ST' @ [T', T]" and wtv: "P,h \ v :\ T" and wtv': "P,h \ v' :\ T'" and ST': "P,h \ stk [:\] ST'" by(auto simp add: list_all2_Cons1 list_all2_append1) from bisim3 have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True from exec have exec': "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e @ [AStore, Push Unit]) ((compxE2 a 0 0 @ compxE2 i (length (compE2 a)) (Suc 0)) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length [v', v]) (compxE2 e 0 0))) t h (stk @ [v', v], loc, length (compE2 a @ compE2 i) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 exec_move_def) hence "exec_meth_d (compP2 P) (compE2 e @ [AStore, Push Unit]) (stack_xlift (length [v', v]) (compxE2 e 0 0)) t h (stk @ [v', v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a @ compE2 i), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 e) (stack_xlift (length [v', v]) (compxE2 e 0 0)) t h (stk @ [v', v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a @ compE2 i), xcp')" using True by(rule exec_meth_take) with bisim3 obtain stk'' where stk': "stk' = stk'' @ [v', v]" and exec'': "exec_move_d P t e h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 a @ compE2 i), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with True have \: "\move2 (compP2 P) h (stk @ [v', v]) (a\i\ := e) (length (compE2 a) + length (compE2 i) + pc) xcp = \move2 (compP2 P) h stk e pc xcp" by(auto simp add: \move2_iff \instr_stk_drop_exec_move) moreover from IH3[OF exec'' _ _ ST' \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,e,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 a) - length (compE2 i), xcp')" and red: "?red e' xs e'' xs'' e stk pc (pc' - length (compE2 a) - length (compE2 i)) xcp xcp'" by auto(fastforce simp only: length_append diff_diff_left) from bisim' have "P,a\i\ := e,h' \ (Val v\Val v'\ := e'', xs'') \ (stk'' @ [v', v], loc', length (compE2 a) + length (compE2 i) + (pc' - length (compE2 a) - length (compE2 i)), xcp')" by(rule bisim1_bisims1.bisim1AAss3) moreover from exec' have "pc' \ length (compE2 a @ compE2 i)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 e pc \ no_call2 (a\i\ := e) (length (compE2 a) + length (compE2 i) + pc)" by(simp add: no_call2_def) ultimately show ?thesis using red \ stk' True by(fastforce intro: AAss1Red3 elim!: AAss_\red1r_xt3 AAss_\red1t_xt3 split: if_split_asm) next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim3 obtain v2 where stk: "stk = [v2]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim3 pc len bsok have red: "\red1r P t h (e', xs) (Val v2, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (Val v\Val v'\ := e', xs) (Val v\Val v'\ := Val v2, loc)" by(rule AAss_\red1r_xt3) moreover have \: "\ \move2 (compP2 P) h [v2, v', v] (a\i\ := e) (length (compE2 a) + length (compE2 i) + length (compE2 e)) None" by(simp add: \move2_iff) moreover have "\ta' e''. P,a\i\ := e,h' \ (e'',loc) \ (stk', loc', pc', xcp') \ True,P,t \1 \Val v\Val v'\ := Val v2, (h, loc)\ -ta'\ \e'',(h', loc)\ \ ta_bisim wbisim1 (extTA2J1 P ta') ta" proof(cases "v = Null") case True with exec stk xcp show ?thesis by(fastforce elim!: exec_meth.cases simp add: exec_move_def intro: bisim1AAssFail Red1AAssNull) next case False with exec stk xcp obtain U A len I where [simp]: "v = Addr A" "v' = Intg I" and hA: "typeof_addr h A = \Array_type U len\" by(fastforce simp add: exec_move_def exec_meth_instr is_Ref_def) from ST' stk obtain T3 where wt3': "typeof\<^bsub>h\<^esub> v2 = \T3\" by(auto simp add: list_all2_Cons1 conf_def) show ?thesis proof(cases "0 <=s I \ sint I < int len") case True note I = True show ?thesis proof(cases "P \ T3 \ U") case True with exec stk xcp True hA I wt3' show ?thesis by(fastforce elim!: exec_meth.cases simp add: compP2_def exec_move_def ta_bisim_def ta_upd_simps intro: Red1AAss bisim1AAss4 split: if_split_asm) next case False with exec stk xcp True hA I wt3' show ?thesis by(fastforce elim!: exec_meth.cases simp add: compP2_def exec_move_def intro: Red1AAssStore bisim1AAssFail split: if_split_asm) qed next case False with exec stk xcp hA show ?thesis by(fastforce elim!: exec_meth.cases intro: bisim1AAssFail Red1AAssBounds simp add: exec_move_def split: if_split_asm) qed qed ultimately show ?thesis using exec xcp stk by(fastforce simp add: no_call2_def) qed next case (bisim1AAssThrow1 A n a xs stk loc pc i e) note exec = \?exec (A\i\ := e) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,A,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 A)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 A 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: match_ex_table_not_pcs_None exec_move_def) thus ?case .. next case (bisim1AAssThrow2 i n a xs stk loc pc A e v) note exec = \?exec (A\i\ := e) (stk @ [v]) loc (length (compE2 A) + pc) \a\ stk' loc' pc' xcp'\ note bisim2 = \P,i,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim2 have pc: "pc < length (compE2 i)" by(auto dest: bisim1_ThrowD) from bisim2 have "match_ex_table (compP2 P) (cname_of h a) (length (compE2 A) + pc) (compxE2 i (length (compE2 A)) 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False apply(auto elim!: exec_meth.cases simp add: compxE2_stack_xlift_convs compxE2_size_convs exec_move_def) apply(auto simp add: match_ex_table_append_not_pcs) done thus ?case .. next case (bisim1AAssThrow3 e n a xs stk loc pc A i v' v) note exec = \?exec (A\i\ := e) (stk @ [v', v]) loc (length (compE2 A) + length (compE2 i) + pc) \a\ stk' loc' pc' xcp'\ note bisim2 = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim2 have "match_ex_table (compP2 P) (cname_of h a) (length (compE2 A) + length (compE2 i) + pc) (compxE2 e (length (compE2 A) + length (compE2 i)) 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec have False apply(auto elim!: exec_meth.cases simp add: compxE2_stack_xlift_convs compxE2_size_convs exec_move_def) apply(auto dest!: match_ex_table_stack_xliftD match_ex_table_shift_pcD dest: match_ex_table_pcsD simp add: match_ex_table_append match_ex_table_shift_pc_None) done thus ?case .. next case (bisim1AAssFail a n i e ad xs v' v v'') note exec = \?exec (a\i\ := e) [v', v, v''] xs (length (compE2 a) + length (compE2 i) + length (compE2 e)) \ad\ stk' loc' pc' xcp'\ hence False by(auto elim!: exec_meth.cases simp add: match_ex_table_append exec_move_def dest!: match_ex_table_shift_pcD match_ex_table_pc_length_compE2) thus ?case .. next case (bisim1AAss4 a n i e xs) have "P,a\i\ := e,h \ (unit, xs) \ ([Unit], xs, length (compE2 (a\i\ := e)), None)" by(rule bisim1Val2) simp moreover have "\move2 (compP2 P) h [] (a\i\ := e) (Suc (length (compE2 a) + length (compE2 i) + length (compE2 e))) None" by(simp add: \move2_iff) moreover note \?exec (a\i\ := e) [] xs (Suc (length (compE2 a) + length (compE2 i) + length (compE2 e))) None stk' loc' pc' xcp'\ ultimately show ?case by(fastforce elim!: exec_meth.cases simp add: ac_simps exec_move_def) next case (bisim1ALength a n a' xs stk loc pc xcp) note IH = bisim1ALength.IH(2) note exec = \?exec (a\length) stk loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,a,h \ (a', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (a'\length) \ length xs\ note bsok = \bsok (a\length) n\ from bisim have pc: "pc \ length (compE2 a)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 a)") case True with exec have exec': "?exec a stk loc pc xcp stk' loc' pc' xcp'" by(auto simp add: exec_move_ALength) from True have \: "\move2 (compP2 P) h stk (a\length) pc xcp = \move2 (compP2 P) h stk a pc xcp" by(simp add: \move2_iff) with IH[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,a,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red: "?red a' xs e'' xs'' a stk pc pc' xcp xcp'" by auto from bisim' have "P,a\length,h' \ (e''\length, xs'') \ (stk', loc', pc', xcp')" by(rule bisim1_bisims1.bisim1ALength) with red \ show ?thesis by(fastforce intro: ALength1Red elim!: ALength_\red1r_xt ALength_\red1t_xt simp add: no_call2_def) next case False with pc have pc: "pc = length (compE2 a)" by auto with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have "\red1r P t h (a', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (a'\length, xs) (Val v\length, loc)" by(rule ALength_\red1r_xt) moreover moreover have \: "\ \move2 (compP2 P) h [v] (a\length) (length (compE2 a)) None" by(simp add: \move2_iff) moreover have "\ta' e''. P,a\length,h' \ (e'',loc) \ (stk', loc', pc', xcp') \ True,P,t \1 \Val v\length, (h, loc)\ -ta'\ \e'',(h', loc)\ \ ta_bisim wbisim1 (extTA2J1 P ta') ta" proof(cases "v = Null") case True with exec stk xcp pc show ?thesis by(fastforce elim!: exec_meth.cases simp add: exec_move_def intro: bisim1ALengthNull Red1ALengthNull) next case False with exec stk xcp pc \P,h \ stk [:\] ST\ obtain U A len where [simp]: "v = Addr A" and hA: "typeof_addr h A = \Array_type U len\" by(fastforce simp add: exec_move_def exec_meth_instr is_Ref_def list_all2_Cons1) have "P,a\length,h' \ (Val (Intg (word_of_int (int len))),loc) \ ([Intg (word_of_int (int len))], loc, length (compE2 (a\length)), None)" by(rule bisim1Val2) simp thus ?thesis using exec stk xcp hA pc by(fastforce elim!: exec_meth.cases intro: Red1ALength simp add: exec_move_def) qed ultimately show ?thesis using \ pc xcp stk by(fastforce elim!: rtranclp_trans simp add: no_call2_def) qed next case (bisim1ALengthThrow A n a xs stk loc pc) note exec = \?exec (A\length) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,A,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 A)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 A 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: exec_move_def) thus ?case .. next case (bisim1ALengthNull a n xs) note exec = \?exec (a\length) [Null] xs (length (compE2 a)) \addr_of_sys_xcpt NullPointer\ stk' loc' pc' xcp'\ hence False by(auto elim!: exec_meth.cases dest!: match_ex_table_pc_length_compE2 simp add: exec_move_def) thus ?case .. next case (bisim1FAcc e n e' xs stk loc pc xcp F D) note IH = bisim1FAcc.IH(2) note exec = \?exec (e\F{D}) stk loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (e'\F{D}) \ length xs\ note bsok = \bsok (e\F{D}) n\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have exec': "?exec e stk loc pc xcp stk' loc' pc' xcp'" by(simp add: exec_move_FAcc) from True have \: "\move2 (compP2 P) h stk (e\F{D}) pc xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff) with IH[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,e,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red: "?red e' xs e'' xs'' e stk pc pc' xcp xcp'" by auto from bisim' have "P,e\F{D},h' \ (e''\F{D}, xs'') \ (stk', loc', pc', xcp')" by(rule bisim1_bisims1.bisim1FAcc) with red \ show ?thesis by(fastforce intro: FAcc1Red elim!: FAcc_\red1r_xt FAcc_\red1t_xt simp add: no_call2_def) next case False with pc have pc: "pc = length (compE2 e)" by auto with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (e'\F{D}, xs) (Val v\F{D}, loc)" by(rule FAcc_\red1r_xt) moreover have \: "\ \move2 (compP2 P) h [v] (e\F{D}) (length (compE2 e)) None" by(simp add: \move2_iff) moreover have "\ta' e''. P,e\F{D},h' \ (e'',loc) \ (stk', loc', pc', xcp') \ True,P,t \1 \Val v\F{D}, (h, loc)\ -ta'\ \e'',(h', loc)\ \ ta_bisim wbisim1 (extTA2J1 P ta') ta" proof(cases "v = Null") case True with exec stk xcp pc show ?thesis by(fastforce elim!: exec_meth.cases simp add: exec_move_def intro: bisim1FAccNull Red1FAccNull) next case False with exec stk xcp pc \P,h \ stk [:\] ST\ obtain A where [simp]: "v = Addr A" by(fastforce simp add: exec_move_def exec_meth_instr is_Ref_def compP2_def) from exec False pc stk xcp obtain v' where v': "heap_read h A (CField D F) v'" "stk' = [v']" by(auto simp add: exec_move_def exec_meth_instr) have "P,e\F{D},h' \ (Val v',loc) \ ([v'], loc, length (compE2 (e\F{D})), None)" by(rule bisim1Val2) simp thus ?thesis using exec stk xcp pc v' by(fastforce elim!: exec_meth.cases intro: Red1FAcc simp add: exec_move_def ta_upd_simps ta_bisim_def) qed ultimately show ?thesis using \ pc xcp stk by(fastforce elim!: rtranclp_trans simp add: no_call2_def) qed next case (bisim1FAccThrow e n a xs stk loc pc F D) note exec = \?exec (e\F{D}) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: exec_move_def) thus ?case .. next case (bisim1FAccNull e n F D xs) note exec = \?exec (e\F{D}) [Null] xs (length (compE2 e)) \addr_of_sys_xcpt NullPointer\ stk' loc' pc' xcp'\ hence False by(auto elim!: exec_meth.cases dest!: match_ex_table_pc_length_compE2 simp add: exec_move_def) thus ?case .. next case (bisim1FAss1 e n e' xs stk loc pc xcp e2 F D) note IH1 = bisim1FAss1.IH(2) note IH2 = bisim1FAss1.IH(4) note exec = \?exec (e\F{D} := e2) stk loc pc xcp stk' loc' pc' xcp'\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,e2,h \ (e2, loc) \ ([], loc, 0, None)\ note len = \n + max_vars (e'\F{D} := e2) \ length xs\ note bsok = \bsok (e\F{D} := e2) n\ from bisim1 have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have exec': "?exec e stk loc pc xcp stk' loc' pc' xcp'" by(simp add: exec_move_FAss1) from True have \: "\move2 (compP2 P) h stk (e\F{D} := e2) pc xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff) with IH1[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,e,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red: "?red e' xs e'' xs'' e stk pc pc' xcp xcp'" by auto from bisim' have "P,e\F{D} := e2,h' \ (e''\F{D} := e2, xs'') \ (stk', loc', pc', xcp')" by(rule bisim1_bisims1.bisim1FAss1) with red \ show ?thesis by(fastforce intro: FAss1Red1 elim!: FAss_\red1r_xt1 FAss_\red1t_xt1 simp add: no_call2_def) next case False with pc have pc: "pc = length (compE2 e)" by auto with bisim1 obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim1 pc len bsok have rede1': "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (e'\F{D} := e2, xs) (Val v\F{D} := e2, loc)" by(rule FAss_\red1r_xt1) moreover from pc exec stk xcp have exec': "exec_meth_d (compP2 P) (compE2 e @ compE2 e2 @ [Putfield F D, Push Unit]) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v]) (compxE2 e2 0 0))) t h ([] @ [v], loc, length (compE2 e) + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs compxE2_stack_xlift_convs exec_move_def) hence "exec_meth_d (compP2 P) (compE2 e2 @ [Putfield F D, Push Unit]) (stack_xlift (length [v]) (compxE2 e2 0 0)) t h ([] @ [v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v]) (compxE2 e2 0 0)) t h ([] @ [v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_take) simp with bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_move_d P t e2 h ([], loc, 0, None) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with pc xcp have \: "\move2 (compP2 P) h [v] (e\F{D} := e2) (length (compE2 e)) None = \move2 (compP2 P) h [] e2 0 None" using \instr_stk_drop_exec_move[where stk="[]" and vs = "[v]"] by(simp add: \move2_iff) from bisim1 have "length xs = length loc" by(rule bisim1_length_xs) with IH2[OF exec'', of "[]"] len bsok obtain e'' xs'' where bisim': "P,e2,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 e), xcp')" and red: "?red e2 loc e'' xs'' e2 [] 0 (pc' - length (compE2 e)) None xcp'" by auto from bisim' have "P,e\F{D} := e2,h' \ (Val v\F{D} := e'', xs'') \ (stk'' @ [v], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule bisim1_bisims1.bisim1FAss2) moreover from red \ have "?red (Val v\F{D} := e2) loc (Val v\F{D} := e'') xs'' (e\F{D} := e2) [v] (length (compE2 e)) pc' None xcp'" by(fastforce intro: FAss1Red2 elim!: FAss_\red1r_xt2 FAss_\red1t_xt2 split: if_split_asm simp add: no_call2_def) moreover from exec' have "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 (e\F{D} := e2) pc" using pc by(simp add: no_call2_def) ultimately show ?thesis using \ stk' pc xcp stk by(fastforce elim!: rtranclp_trans) qed next case (bisim1FAss2 e2 n e' xs stk loc pc xcp e F D v) note IH2 = bisim1FAss2.IH(2) note exec = \?exec (e\F{D} := e2) (stk @ [v]) loc (length (compE2 e) + pc) xcp stk' loc' pc' xcp'\ note bisim2 = \P,e2,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (Val v\F{D} := e') \ length xs\ note bsok = \bsok (e\F{D} := e2) n\ note ST = \P,h \ stk @ [v] [:\] ST\ then obtain T ST' where ST': "P,h \ stk [:\] ST'" and T: "typeof\<^bsub>h\<^esub> v = \T\" by(auto simp add: list_all2_append1 list_all2_Cons1 conf_def) from bisim2 have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e2)") case True from exec have exec': "exec_meth_d (compP2 P) (compE2 e @ compE2 e2 @ [Putfield F D, Push Unit]) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v]) (compxE2 e2 0 0))) t h (stk @ [v], loc, length (compE2 e) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 exec_move_def) hence "exec_meth_d (compP2 P) (compE2 e2 @ [Putfield F D, Push Unit]) (stack_xlift (length [v]) (compxE2 e2 0 0)) t h (stk @ [v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v]) (compxE2 e2 0 0)) t h (stk @ [v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 e), xcp')" using True by(rule exec_meth_take) with bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_move_d P t e2 h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with True have \: "\move2 (compP2 P) h (stk @ [v]) (e\F{D} := e2) (length (compE2 e) + pc) xcp = \move2 (compP2 P) h stk e2 pc xcp" by(auto simp add: \move2_iff \instr_stk_drop_exec_move) moreover from IH2[OF exec'' _ _ ST' \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,e2,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 e), xcp')" and red: "?red e' xs e'' xs'' e2 stk pc (pc' - length (compE2 e)) xcp xcp'" by auto from bisim' have "P,e\F{D} := e2,h' \ (Val v\F{D} := e'', xs'') \ (stk'' @ [v], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule bisim1_bisims1.bisim1FAss2) moreover from exec' have "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc) auto ultimately show ?thesis using red \ stk' True by(fastforce intro: FAss1Red2 elim!: FAss_\red1r_xt2 FAss_\red1t_xt2 split: if_split_asm simp add: no_call2_def) next case False with pc have [simp]: "pc = length (compE2 e2)" by simp with bisim2 obtain v2 where stk: "stk = [v2]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim2 pc len bsok have red: "\red1r P t h (e', xs) (Val v2, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (Val v\F{D} := e', xs) (Val v\F{D} := Val v2, loc)" by(rule FAss_\red1r_xt2) moreover have \: "\ \move2 (compP2 P) h [v2, v] (e\F{D} := e2) (length (compE2 e) + length (compE2 e2)) None" by(simp add: \move2_iff) moreover have "\ta' e''. P,e\F{D} := e2,h' \ (e'',loc) \ (stk', loc', pc', xcp') \ True,P,t \1 \Val v\F{D} := Val v2, (h, loc)\ -ta'\ \e'',(h', loc)\ \ ta_bisim wbisim1 (extTA2J1 P ta') ta" proof(cases "v = Null") case True with exec stk xcp show ?thesis by(fastforce elim!: exec_meth.cases simp add: exec_move_def intro: bisim1FAssNull Red1FAssNull) next case False with exec stk xcp T show ?thesis by(fastforce simp add: exec_move_def compP2_def exec_meth_instr is_Ref_def ta_upd_simps ta_bisim_def intro: bisim1FAss3 Red1FAss) qed ultimately show ?thesis using exec xcp stk by(fastforce simp add: no_call2_def) qed next case (bisim1FAssThrow1 e n a xs stk loc pc e2 F D) note exec = \?exec (e\F{D} := e2) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: exec_move_def match_ex_table_not_pcs_None) thus ?case .. next case (bisim1FAssThrow2 e2 n a xs stk loc pc e F D v) note exec = \?exec (e\F{D} := e2) (stk @ [v]) loc (length (compE2 e) + pc) \a\ stk' loc' pc' xcp'\ note bisim2 = \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ hence "match_ex_table (compP2 P) (cname_of h a) (length (compE2 e) + pc) (compxE2 e2 (length (compE2 e)) 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec have False by(auto elim!: exec_meth.cases simp add: compxE2_stack_xlift_convs exec_move_def)(auto dest!: match_ex_table_stack_xliftD simp add: match_ex_table_append_not_pcs) thus ?case .. next case (bisim1FAssNull e n e2 F D xs v) note exec = \?exec (e\F{D} := e2) [v, Null] xs (length (compE2 e) + length (compE2 e2)) \addr_of_sys_xcpt NullPointer\ stk' loc' pc' xcp'\ hence False by(auto elim!: exec_meth.cases simp add: match_ex_table_append_not_pcs compxE2_size_convs exec_move_def dest!: match_ex_table_shift_pcD match_ex_table_pc_length_compE2) thus ?case .. next case (bisim1FAss3 e n e2 F D xs) have "P,e\F{D} := e2,h \ (unit, xs) \ ([Unit], xs, length (compE2 (e\F{D} := e2)), None)" by(rule bisim1Val2) simp moreover have "\move2 (compP2 P) h [] (e\F{D} := e2) (Suc (length (compE2 e) + length (compE2 e2))) None" by(simp add: \move2_iff) moreover note \?exec (e\F{D} := e2) [] xs (Suc (length (compE2 e) + length (compE2 e2))) None stk' loc' pc' xcp'\ ultimately show ?case by(fastforce elim!: exec_meth.cases simp add: ac_simps exec_move_def) next case (bisim1CAS1 a n a' xs stk loc pc xcp i e D F) note IH1 = bisim1CAS1.IH(2) note IH2 = bisim1CAS1.IH(4) note exec = \?exec (a\compareAndSwap(D\F, i, e)) stk loc pc xcp stk' loc' pc' xcp'\ note bisim1 = \P,a,h \ (a', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,i,h \ (i, loc) \ ([], loc, 0, None)\ note len = \n + max_vars _ \ length xs\ note bsok = \bsok (a\compareAndSwap(D\F, i, e)) n\ from bisim1 have pc: "pc \ length (compE2 a)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 a)") case True with exec have exec': "?exec a stk loc pc xcp stk' loc' pc' xcp'" by(simp add: exec_move_CAS1) from True have \: "\move2 (compP2 P) h stk (a\compareAndSwap(D\F, i, e)) pc xcp = \move2 (compP2 P) h stk a pc xcp" by(simp add: \move2_iff) with IH1[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,a,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red: "?red a' xs e'' xs'' a stk pc pc' xcp xcp'" by auto from bisim' have "P,a\compareAndSwap(D\F, i, e),h' \ (e''\compareAndSwap(D\F, i, e), xs'') \ (stk', loc', pc', xcp')" by(rule bisim1_bisims1.bisim1CAS1) moreover from True have "no_call2 (a\compareAndSwap(D\F, i, e)) pc = no_call2 a pc" by(simp add: no_call2_def) ultimately show ?thesis using red \ by(fastforce intro: CAS1Red1 elim!: CAS_\red1r_xt1 CAS_\red1t_xt1) next case False with pc have pc: "pc = length (compE2 a)" by auto with bisim1 obtain v where a': "is_val a' \ a' = Val v" and stk: "stk = [v]" and xcp: "xcp = None" and call: "call1 a' = None" by(auto dest: bisim1_pc_length_compE2D) with bisim1 pc len bsok have rede1': "\red1r P t h (a', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (a'\compareAndSwap(D\F, i, e), xs) (Val v\compareAndSwap(D\F, i, e), loc)" by(rule CAS_\red1r_xt1) moreover from pc exec stk xcp have exec': "exec_meth_d (compP2 P) (compE2 a @ compE2 i @ compE2 e @ [CAS F D]) (compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v]) (compxE2 i 0 0) @ shift (length (compE2 i)) (compxE2 e 0 (Suc (Suc 0))))) t h ([] @ [v], loc, length (compE2 a) + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs compxE2_stack_xlift_convs exec_move_def) hence "exec_meth_d (compP2 P) (compE2 i @ compE2 e @ [CAS F D]) (stack_xlift (length [v]) (compxE2 i 0 0) @ shift (length (compE2 i)) (compxE2 e 0 (Suc (Suc 0)))) t h ([] @ [v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v]) (compxE2 i 0 0)) t h ([] @ [v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_take_xt) simp with bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_move_d P t i h ([], loc, 0, None) ta h' (stk'', loc', pc' - length (compE2 a), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with pc xcp have \: "\move2 (compP2 P) h [v] (a\compareAndSwap(D\F, i, e)) (length (compE2 a)) None = \move2 (compP2 P) h [] i 0 None" using \instr_stk_drop_exec_move[where stk="[]" and vs="[v]"] by(auto simp add: \move2_iff) from bisim1 have "length xs = length loc" by(rule bisim1_length_xs) with IH2[OF exec''] len bsok obtain e'' xs'' where bisim': "P,i,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 a), xcp')" and red: "?red i loc e'' xs'' i [] 0 (pc' - length (compE2 a)) None xcp'" by fastforce from bisim' have "P,a\compareAndSwap(D\F, i, e),h' \ (Val v\compareAndSwap(D\F, e'', e), xs'') \ (stk'' @ [v], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" by(rule bisim1_bisims1.bisim1CAS2) moreover from red \ have "?red (Val v\compareAndSwap(D\F, i, e)) loc (Val v\compareAndSwap(D\F, e'', e)) xs'' (a\compareAndSwap(D\F, i, e)) [v] (length (compE2 a)) pc' None xcp'" by(fastforce intro: CAS1Red2 elim!: CAS_\red1r_xt2 CAS_\red1t_xt2 split: if_split_asm simp add: no_call2_def) moreover from exec' have "pc' \ length (compE2 a)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 (a\compareAndSwap(D\F, i, e)) pc" using pc by(simp add: no_call2_def) ultimately show ?thesis using \ stk' pc xcp stk by(fastforce elim!: rtranclp_trans) qed next case (bisim1CAS2 i n i' xs stk loc pc xcp a e D F v) note IH2 = bisim1CAS2.IH(2) note IH3 = bisim1CAS2.IH(6) note exec = \?exec (a\compareAndSwap(D\F, i, e)) (stk @ [v]) loc (length (compE2 a) + pc) xcp stk' loc' pc' xcp'\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, xcp)\ note bisim3 = \P,e,h \ (e, loc) \ ([], loc, 0, None)\ note len = \n + max_vars (Val v\compareAndSwap(D\F, i', e)) \ length xs\ note bsok = \bsok (a\compareAndSwap(D\F, i, e)) n\ from bisim2 have pc: "pc \ length (compE2 i)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 i)") case True from exec have exec': "exec_meth_d (compP2 P) (compE2 a @ compE2 i @ compE2 e @ [CAS F D]) (compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v]) (compxE2 i 0 0) @ shift (length (compE2 i)) (compxE2 e 0 (Suc (Suc 0))))) t h (stk @ [v], loc, length (compE2 a) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps exec_move_def) hence "exec_meth_d (compP2 P) (compE2 i @ compE2 e @ [CAS F D]) (stack_xlift (length [v]) (compxE2 i 0 0) @ shift (length (compE2 i)) (compxE2 e 0 (Suc (Suc 0)))) t h (stk @ [v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v]) (compxE2 i 0 0)) t h (stk @ [v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a), xcp')" using True by(rule exec_meth_take_xt) with bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_move_d P t i h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 a), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with True have \: "\move2 (compP2 P) h (stk @ [v]) (a\compareAndSwap(D\F, i, e)) (length (compE2 a) + pc) xcp = \move2 (compP2 P) h stk i pc xcp" by(auto simp add: \move2_iff \instr_stk_drop_exec_move) moreover from \P,h \ stk @ [v] [:\] ST\ obtain ST2 where "P,h \ stk [:\] ST2" by(auto simp add: list_all2_append1) from IH2[OF exec'' _ _ this \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,i,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 a), xcp')" and red: "?red i' xs e'' xs'' i stk pc (pc' - length (compE2 a)) xcp xcp'" by fastforce from bisim' have "P,a\compareAndSwap(D\F, i, e),h' \ (Val v\compareAndSwap(D\F, e'', e), xs'') \ (stk'' @ [v], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" by(rule bisim1_bisims1.bisim1CAS2) moreover from exec' have "pc' \ length (compE2 a)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 i pc \ no_call2 (a\compareAndSwap(D\F, i, e)) (length (compE2 a) + pc)" by(simp add: no_call2_def) ultimately show ?thesis using red \ stk' True by(fastforce intro: CAS1Red2 elim!: CAS_\red1r_xt2 CAS_\red1t_xt2 split: if_split_asm) next case False with pc have [simp]: "pc = length (compE2 i)" by simp with bisim2 obtain v2 where i': "is_val i' \ i' = Val v2" and stk: "stk = [v2]" and xcp: "xcp = None" and call: "call1 i' = None" by(auto dest: bisim1_pc_length_compE2D) with bisim2 pc len bsok have red: "\red1r P t h (i', xs) (Val v2, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (Val v\compareAndSwap(D\F, i', e ), xs) (Val v\compareAndSwap(D\F, Val v2, e), loc)" by(rule CAS_\red1r_xt2) moreover from pc exec stk xcp have exec': "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e @ [CAS F D]) ((compxE2 a 0 0 @ compxE2 i (length (compE2 a)) (Suc 0)) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length [v2, v]) (compxE2 e 0 0))) t h ([] @ [v2, v], loc, length (compE2 a @ compE2 i) + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs compxE2_stack_xlift_convs exec_move_def) hence "exec_meth_d (compP2 P) (compE2 e @ [CAS F D]) (stack_xlift (length [v2, v]) (compxE2 e 0 0)) t h ([] @ [v2, v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 a @ compE2 i), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 e) (stack_xlift (length [v2, v]) (compxE2 e 0 0)) t h ([] @ [v2, v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 a @ compE2 i), xcp')" by(rule exec_meth_take) simp with bisim3 obtain stk'' where stk': "stk' = stk'' @ [v2, v]" and exec'': "exec_move_d P t e h ([], loc, 0, None) ta h' (stk'', loc', pc' - length (compE2 a @ compE2 i), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with pc xcp have \: "\move2 (compP2 P) h [v2, v] (a\compareAndSwap(D\F, i, e)) (length (compE2 a) + length (compE2 i)) None = \move2 (compP2 P) h [] e 0 None" using \instr_stk_drop_exec_move[where stk="[]" and vs="[v2, v]"] by(simp add: \move2_iff) from bisim2 have "length xs = length loc" by(rule bisim1_length_xs) with IH3[OF exec'', of "[]"] len bsok obtain e'' xs'' where bisim': "P,e,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 a) - length (compE2 i), xcp')" and red: "?red e loc e'' xs'' e [] 0 (pc' - length (compE2 a) - length (compE2 i)) None xcp'" by auto (fastforce simp only: length_append diff_diff_left) from bisim' have "P,a\compareAndSwap(D\F, i, e),h' \ (Val v\compareAndSwap(D\F, Val v2, e''), xs'') \ (stk'' @ [v2, v], loc', length (compE2 a) + length (compE2 i) + (pc' - length (compE2 a) - length (compE2 i)), xcp')" by(rule bisim1_bisims1.bisim1CAS3) moreover from red \ have "?red (Val v\compareAndSwap(D\F, Val v2, e)) loc (Val v\compareAndSwap(D\F, Val v2, e'')) xs'' (a\compareAndSwap(D\F, i, e)) [v2, v] (length (compE2 a) + length (compE2 i)) pc' None xcp'" by(fastforce intro: CAS1Red3 elim!: CAS_\red1r_xt3 CAS_\red1t_xt3 split: if_split_asm simp add: no_call2_def) moreover from exec' have "pc' \ length (compE2 a @ compE2 i)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 (a\compareAndSwap(D\F, i, e)) (length (compE2 a) + pc)" by(simp add: no_call2_def) ultimately show ?thesis using \ stk' pc xcp stk by(fastforce elim!: rtranclp_trans) qed next case (bisim1CAS3 e n e' xs stk loc pc xcp a i D F v v') note IH3 = bisim1CAS3.IH(2) note exec = \?exec (a\compareAndSwap(D\F, i, e)) (stk @ [v', v]) loc (length (compE2 a) + length (compE2 i) + pc) xcp stk' loc' pc' xcp'\ note bisim3 = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars _ \ length xs\ note bsok = \bsok (a\compareAndSwap(D\F, i, e)) n\ from \P,h \ stk @ [v', v] [:\] ST\ obtain T T' ST' where [simp]: "ST = ST' @ [T', T]" and wtv: "P,h \ v :\ T" and wtv': "P,h \ v' :\ T'" and ST': "P,h \ stk [:\] ST'" by(auto simp add: list_all2_Cons1 list_all2_append1) from bisim3 have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True from exec have exec': "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e @ [CAS F D]) ((compxE2 a 0 0 @ compxE2 i (length (compE2 a)) (Suc 0)) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length [v', v]) (compxE2 e 0 0))) t h (stk @ [v', v], loc, length (compE2 a @ compE2 i) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 exec_move_def) hence "exec_meth_d (compP2 P) (compE2 e @ [CAS F D]) (stack_xlift (length [v', v]) (compxE2 e 0 0)) t h (stk @ [v', v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a @ compE2 i), xcp')" by(rule exec_meth_drop_xt) auto hence "exec_meth_d (compP2 P) (compE2 e) (stack_xlift (length [v', v]) (compxE2 e 0 0)) t h (stk @ [v', v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a @ compE2 i), xcp')" using True by(rule exec_meth_take) with bisim3 obtain stk'' where stk': "stk' = stk'' @ [v', v]" and exec'': "exec_move_d P t e h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 a @ compE2 i), xcp')" unfolding exec_move_def by(blast dest: exec_meth_stk_split) with True have \: "\move2 (compP2 P) h (stk @ [v', v]) (a\compareAndSwap(D\F, i, e)) (length (compE2 a) + length (compE2 i) + pc) xcp = \move2 (compP2 P) h stk e pc xcp" by(auto simp add: \move2_iff \instr_stk_drop_exec_move) moreover from IH3[OF exec'' _ _ ST' \conf_xcp' (compP2 P) h xcp\] len bsok obtain e'' xs'' where bisim': "P,e,h' \ (e'', xs'') \ (stk'', loc', pc' - length (compE2 a) - length (compE2 i), xcp')" and red: "?red e' xs e'' xs'' e stk pc (pc' - length (compE2 a) - length (compE2 i)) xcp xcp'" by auto(fastforce simp only: length_append diff_diff_left) from bisim' have "P,a\compareAndSwap(D\F, i, e),h' \ (Val v\compareAndSwap(D\F, Val v', e''), xs'') \ (stk'' @ [v', v], loc', length (compE2 a) + length (compE2 i) + (pc' - length (compE2 a) - length (compE2 i)), xcp')" by(rule bisim1_bisims1.bisim1CAS3) moreover from exec' have "pc' \ length (compE2 a @ compE2 i)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 e pc \ no_call2 (a\compareAndSwap(D\F, i, e)) (length (compE2 a) + length (compE2 i) + pc)" by(simp add: no_call2_def) ultimately show ?thesis using red \ stk' True by(fastforce intro: CAS1Red3 elim!: CAS_\red1r_xt3 CAS_\red1t_xt3 split: if_split_asm) next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim3 obtain v2 where stk: "stk = [v2]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim3 pc len bsok have red: "\red1r P t h (e', xs) (Val v2, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (Val v\compareAndSwap(D\F, Val v', e'), xs) (Val v\compareAndSwap(D\F, Val v', Val v2), loc)" by(rule CAS_\red1r_xt3) moreover have \: "\ \move2 (compP2 P) h [v2, v', v] (a\compareAndSwap(D\F, i, e)) (length (compE2 a) + length (compE2 i) + length (compE2 e)) None" by(simp add: \move2_iff) moreover have "\ta' e''. P,a\compareAndSwap(D\F, i, e),h' \ (e'',loc) \ (stk', loc', pc', xcp') \ True,P,t \1 \Val v\compareAndSwap(D\F, Val v', Val v2), (h, loc)\ -ta'\ \e'',(h', loc)\ \ ta_bisim wbisim1 (extTA2J1 P ta') ta" proof(cases "v = Null") case True with exec stk xcp show ?thesis by(fastforce elim!: exec_meth.cases simp add: exec_move_def intro: bisim1CASFail CAS1Null) next case False have "P,a\compareAndSwap(D\F, i, e),h' \ (Val (Bool b), loc) \ ([Bool b], loc, length (compE2 (a\compareAndSwap(D\F, i, e))), None)" for b by(rule bisim1Val2) simp with False exec stk xcp show ?thesis - by(auto elim!: exec_meth.cases simp add: exec_move_def is_Ref_def ac_simps intro: Red1CASSucceed Red1CASFail) - (fastforce intro!: Red1CASSucceed Red1CASFail simp add: ta_bisim_def)+ + by (auto elim!: exec_meth.cases simp add: exec_move_def is_Ref_def intro: Red1CASSucceed Red1CASFail) + (fastforce intro!: Red1CASSucceed Red1CASFail simp add: ta_bisim_def ac_simps)+ qed ultimately show ?thesis using exec xcp stk by(fastforce simp add: no_call2_def) qed next case (bisim1CASThrow1 A n a xs stk loc pc i e D F) note exec = \?exec (A\compareAndSwap(D\F, i, e)) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,A,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 A)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 A 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: match_ex_table_not_pcs_None exec_move_def) thus ?case .. next case (bisim1CASThrow2 i n a xs stk loc pc A e D F v) note exec = \?exec (A\compareAndSwap(D\F, i, e)) (stk @ [v]) loc (length (compE2 A) + pc) \a\ stk' loc' pc' xcp'\ note bisim2 = \P,i,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim2 have pc: "pc < length (compE2 i)" by(auto dest: bisim1_ThrowD) from bisim2 have "match_ex_table (compP2 P) (cname_of h a) (length (compE2 A) + pc) (compxE2 i (length (compE2 A)) 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False apply(auto elim!: exec_meth.cases simp add: compxE2_stack_xlift_convs compxE2_size_convs exec_move_def) apply(auto simp add: match_ex_table_append_not_pcs) done thus ?case .. next case (bisim1CASThrow3 e n a xs stk loc pc A i D F v' v) note exec = \?exec (A\compareAndSwap(D\F, i, e)) (stk @ [v', v]) loc (length (compE2 A) + length (compE2 i) + pc) \a\ stk' loc' pc' xcp'\ note bisim2 = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim2 have "match_ex_table (compP2 P) (cname_of h a) (length (compE2 A) + length (compE2 i) + pc) (compxE2 e (length (compE2 A) + length (compE2 i)) 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec have False apply(auto elim!: exec_meth.cases simp add: compxE2_stack_xlift_convs compxE2_size_convs exec_move_def) apply(auto dest!: match_ex_table_stack_xliftD match_ex_table_shift_pcD dest: match_ex_table_pcsD simp add: match_ex_table_append match_ex_table_shift_pc_None) done thus ?case .. next case (bisim1CASFail a n i e D F ad xs v' v v'') note exec = \?exec (a\compareAndSwap(D\F, i, e)) [v', v, v''] xs (length (compE2 a) + length (compE2 i) + length (compE2 e)) \ad\ stk' loc' pc' xcp'\ hence False by(auto elim!: exec_meth.cases simp add: match_ex_table_append exec_move_def dest!: match_ex_table_shift_pcD match_ex_table_pc_length_compE2) thus ?case .. next case (bisim1Call1 obj n obj' xs stk loc pc xcp ps M') note IH1 = bisim1Call1.IH(2) note IH2 = bisim1Call1.IH(4) note exec = \?exec (obj\M'(ps)) stk loc pc xcp stk' loc' pc' xcp'\ note bisim1 = \P,obj,h \ (obj', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,ps,h \ (ps, loc) [\] ([], loc, 0, None)\ note len = \n + max_vars (obj'\M'(ps)) \ length xs\ note bsok = \bsok (obj\M'(ps)) n\ from bisim1 have pc: "pc \ length (compE2 obj)" by(rule bisim1_pc_length_compE2) from bisim1 have lenxs: "length xs = length loc" by(rule bisim1_length_xs) show ?case proof(cases "pc < length (compE2 obj)") case True with exec have exec': "?exec obj stk loc pc xcp stk' loc' pc' xcp'" by(simp add: exec_move_Call1) from True have "\move2 (compP2 P) h stk (obj\M'(ps)) pc xcp = \move2 (compP2 P) h stk obj pc xcp" by(simp add: \move2_iff) moreover from True have "no_call2 (obj\M'(ps)) pc = no_call2 obj pc" by(simp add: no_call2_def) ultimately show ?thesis using IH1[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bisim2 len bsok by(fastforce intro: bisim1_bisims1.bisim1Call1 Call1Obj elim!: Call_\red1r_obj Call_\red1t_obj) next case False with pc have pc: "pc = length (compE2 obj)" by auto with bisim1 obtain v where stk: "stk = [v]" and xcp: "xcp = None" and call: "call1 obj' = None" and v: "is_val obj' \ obj' = Val v \ xs = loc" by(auto dest: bisim1_pc_length_compE2D) with bisim1 pc len bsok have "\red1r P t h (obj', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence red: "\red1r P t h (obj'\M'(ps), xs) (Val v\M'(ps), loc)" by(rule Call_\red1r_obj) show ?thesis proof(cases ps) case (Cons p ps') from pc exec stk xcp have "exec_meth_d (compP2 P) (compE2 (obj\M'(ps))) (compxE2 (obj\M'(ps)) 0 0) t h ([] @ [v], loc, length (compE2 obj) + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs compxE2_stack_xlift_convs exec_move_def) hence exec': "exec_meth_d (compP2 P) (compEs2 ps) (stack_xlift (length [v]) (compxEs2 ps 0 0)) t h ([] @ [v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 obj), xcp')" and pc': "pc' \ length (compE2 obj)" using Cons by(safe dest!: Call_execParamD) simp_all from exec' bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_moves_d P t ps h ([], loc, 0, None) ta h' (stk'', loc', pc' - length (compE2 obj), xcp')" unfolding exec_moves_def by -(drule (1) exec_meth_stk_splits, auto) with pc xcp have \: "\move2 (compP2 P) h [v] (obj\M'(ps)) (length (compE2 obj)) None = \moves2 (compP2 P) h [] ps 0 None" using \instr_stk_drop_exec_moves[where stk="[]" and vs="[v]"] by(auto simp add: \move2_iff \moves2_iff) from IH2[OF exec'', of "[]"] len lenxs bsok obtain es'' xs'' where bisim': "P,ps,h' \ (es'', xs'') [\] (stk'', loc', pc' - length (compE2 obj), xcp')" and red': "?reds ps loc es'' xs'' ps [] 0 (pc' - length (compE2 obj)) None xcp'" by auto from bisim' have "P,obj\M'(ps),h' \ (Val v\M'(es''), xs'') \ (stk'' @ [v], loc', length (compE2 obj) + (pc' - length (compE2 obj)), xcp')" by(rule bisim1CallParams) moreover from pc Cons have "no_call2 (obj\M'(ps)) pc" by(simp add: no_call2_def) ultimately show ?thesis using red red' \ stk' pc xcp pc' stk call by(fastforce elim!: rtranclp_trans Call_\red1r_param Call_\red1t_param intro: Call1Params rtranclp_tranclp_tranclp split: if_split_asm) next case [simp]: Nil from exec pc stk xcp have "v = Null \ (is_Addr v \ (\T C' Ts' Tr' D'. typeof\<^bsub>h\<^esub> v = \T\ \ class_type_of' T = \C'\ \ P \ C' sees M':Ts'\Tr' = Native in D'))" (is "_ \ ?rest") by(fastforce elim!: exec_meth.cases simp add: is_Ref_def exec_move_def compP2_def has_method_def split: if_split_asm) thus ?thesis proof assume [simp]: "v = Null" have "P,obj\M'([]),h \ (THROW NullPointer, loc) \ ([] @ [Null], loc, length (compE2 obj) + length (compEs2 ([] :: 'addr expr1 list)), \addr_of_sys_xcpt NullPointer\)" by(safe intro!: bisim1CallThrow) simp_all moreover have "True,P,t \1 \null\M'(map Val []),(h, loc)\ -\\ \THROW NullPointer,(h, loc)\" by(rule Red1CallNull) moreover have "\move1 P h (Val v\M'([]))" "\move2 (compP2 P) h [Null] (obj\M'(ps)) (length (compE2 obj)) None" by(simp_all add: \move2_iff) ultimately show ?thesis using exec pc stk xcp red by(fastforce elim!: exec_meth.cases intro: rtranclp_trans simp add: exec_move_def) next assume ?rest then obtain a Ta Ts' Tr' D' where [simp]: "v = Addr a" and Ta: "typeof_addr h a = \Ta\" and iec: "P \ class_type_of Ta sees M': Ts'\Tr' = Native in D'" by auto with exec pc stk xcp obtain ta' va h'' U where redex: "(ta', va, h'') \ red_external_aggr (compP2 P) t a M' [] h" and ret: "(xcp', h', [(stk', loc', undefined, undefined, pc')]) = extRet2JVM 0 h'' [Addr a] loc undefined undefined (length (compE2 obj)) [] va" and wtext': "P,h \ a\M'([]) : U" and ta': "ta = extTA2JVM (compP2 P) ta'" by(fastforce simp add: is_Ref_def exec_move_def compP2_def external_WT'_iff exec_meth_instr) from Ta iec have \: "\move1 P h (Val v\M'([])) \ \move2 (compP2 P) h [Addr a] (obj\M'(ps)) (length (compE2 obj)) None" by(auto simp add: \move2_iff compP2_def) from ret have [simp]: "h'' = h'" by simp from wtext' have wtext'': "compP2 P,h \ a\M'([]) : U" by(simp add: external_WT'_iff compP2_def) from wf have "wf_jvm_prog (compP2 P)" by(rule wt_compP2) then obtain wf_md where wf': "wf_prog wf_md (compP2 P)" unfolding wf_jvm_prog_def by(blast dest: wt_jvm_progD) from tconf have tconf': "compP2 P,h \ t \t" by(simp add: compP2_def tconf_def) from redex have redex'': "compP2 P,t \ \a\M'([]), h\ -ta'\ext \va, h'\" by(simp add: WT_red_external_list_conv[OF wf' wtext'' tconf', symmetric]) hence redex': "P,t \ \a\M'([]), h\ -ta'\ext \va, h'\" by(simp add: compP2_def) with Ta iec have "True,P,t \1 \addr a\M'(map Val []), (h, loc)\ -ta'\ \extRet2J (addr a\M'(map Val [])) va, (h', loc)\" by -(rule Red1CallExternal, auto) moreover have "P,obj\M'(ps),h' \ (extRet2J (addr a\M'(map Val [])) va, loc) \ (stk', loc', pc', xcp')" proof(cases va) case (RetVal v) have "P,obj\M'([]),h' \ (Val v, loc) \ ([v], loc, length (compE2 (obj\M'([]))), None)" by(rule bisim1Val2) simp with ret RetVal show ?thesis by simp next case (RetExc ad) have "P,obj\M'([]),h' \ (Throw ad, loc) \ ([] @ [v], loc, length (compE2 (obj)) + length (compEs2 ([] :: 'addr expr1 list)), \ad\)" by(rule bisim1CallThrow) simp with ret RetExc show ?thesis by simp next case RetStaySame have "P,obj\M'([]),h' \ (addr a\M'([]), loc) \ ([Addr a], loc, length (compE2 obj), None)" by(rule bisim1_bisims1.bisim1Call1)(rule bisim1Val2, simp) thus ?thesis using ret RetStaySame by simp qed moreover from \preallocated h\ red_external_hext[OF redex'] have "preallocated h'" by(rule preallocated_hext) from redex'' wtext'' \hconf h\ have "hconf h'" by(rule external_call_hconf) with ta' redex' \preallocated h'\ have "ta_bisim wbisim1 (extTA2J1 P ta') ta" by(auto intro: red_external_ta_bisim21[OF wf]) moreover have "\move1 P h (Val v\M'([])) \ ta' = \ \ h' = h" using redex' Ta iec by(fastforce dest: \external'_red_external_heap_unchanged \external'_red_external_TA_empty sees_method_fun simp add: \external'_def \external_def) moreover from v call have "call1 (obj'\M'(ps)) \ None \ Val v\M'(ps) = obj'\M'(ps) \ loc = xs" by(auto split: if_split_asm) ultimately show ?thesis using red \ pc xcp stk Ta call iec apply(auto simp del: call1.simps calls1.simps) apply(rule exI conjI|assumption|erule rtranclp.rtrancl_into_rtrancl rtranclp_into_tranclp1|drule (1) sees_method_fun|clarsimp)+ done qed qed qed next case (bisim1CallParams ps n ps' xs stk loc pc xcp obj M' v) note bisim2 = \P,ps,h \ (ps', xs) [\] (stk, loc, pc, xcp)\ note bisim1 = \P,obj,h \ (obj, xs) \ ([], xs, 0, None)\ note IH2 = bisim1CallParams.IH(2) note exec = \exec_move_d P t (obj\M'(ps)) h (stk @ [v], loc, length (compE2 obj) + pc, xcp) ta h' (stk', loc', pc', xcp')\ note len = \n + max_vars (Val v\M'(ps')) \ length xs\ note bsok = \bsok (obj\M'(ps)) n\ from \P,h \ stk @ [v] [:\] ST\ obtain T ST' where ST': "P,h \ stk [:\] ST'" and T: "typeof\<^bsub>h\<^esub> v = \T\" by(auto simp add: list_all2_Cons1 list_all2_append1 conf_def) from bisim2 have pc: "pc \ length (compEs2 ps)" by(rule bisims1_pc_length_compEs2) show ?case proof(cases "pc < length (compEs2 ps)") case True from exec have "exec_meth_d (compP2 P) (compE2 (obj\M'(ps))) (compxE2 (obj\M'(ps)) 0 0) t h (stk @ [v], loc, length (compE2 obj) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: exec_move_def) with True have exec': "exec_meth_d (compP2 P) (compEs2 ps) (stack_xlift (length [v]) (compxEs2 ps 0 0)) t h (stk @ [v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 obj), xcp')" and pc': "pc' \ length (compE2 obj)" by(safe dest!: Call_execParamD) from exec' bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_moves_d P t ps h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 obj), xcp')" by(unfold exec_moves_def)(drule (1) exec_meth_stk_splits, auto) with True have \: "\move2 (compP2 P) h (stk @ [v]) (obj\M'(ps)) (length (compE2 obj) + pc) xcp = \moves2 (compP2 P) h stk ps pc xcp" by(auto simp add: \move2_iff \moves2_iff \instr_stk_drop_exec_moves) from IH2[OF exec'' _ _ ST' \conf_xcp' (compP2 P) h xcp\] len bsok obtain es'' xs'' where bisim': "P,ps,h' \ (es'', xs'') [\] (stk'', loc', pc' - length (compE2 obj), xcp')" and red': "?reds ps' xs es'' xs'' ps stk pc (pc' - length (compE2 obj)) xcp xcp'" by auto from bisim' have "P,obj\M'(ps),h' \ (Val v\M'(es''), xs'') \ (stk'' @ [v], loc', length (compE2 obj) + (pc' - length (compE2 obj)), xcp')" by(rule bisim1_bisims1.bisim1CallParams) moreover from True have "no_call2 (obj\M'(ps)) (length (compE2 obj) + pc) = no_calls2 ps pc" by(simp add: no_calls2_def no_call2_def) moreover have "calls1 ps' = None \ call1 (Val v\M'(ps')) = None \ is_vals ps'" by simp ultimately show ?thesis using red' \ stk' pc pc' by(fastforce intro: Call1Params elim!: Call_\red1r_param Call_\red1t_param split: if_split_asm simp add: is_vals_conv) next case False with pc have [simp]: "pc = length (compEs2 ps)" by simp with bisim2 obtain vs where [simp]: "stk = rev vs" "xcp = None" and lenvs: "length vs = length ps" and vs: "is_vals ps' \ ps' = map Val vs \ xs = loc" and call: "calls1 ps' = None" by(auto dest: bisims1_pc_length_compEs2D) with bisim2 len bsok have reds: "\reds1r P t h (ps', xs) (map Val vs, loc)" by(auto intro: bisims1_Val_\Reds1r simp add: bsok_def) from exec T lenvs have "v = Null \ (is_Addr v \ (\T C' Ts' Tr' D'. typeof\<^bsub>h\<^esub> v = \T\ \ class_type_of' T = \C'\ \ P \ C' sees M':Ts'\Tr' = Native in D'))" (is "_ \ ?rest") by(fastforce elim!: exec_meth.cases simp add: is_Ref_def exec_move_def compP2_def has_method_def split: if_split_asm) thus ?thesis proof assume [simp]: "v = Null" hence \: "\move1 P h (Val v\M'(map Val vs))" "\move2 (compP2 P) h (stk @ [v]) (obj\M'(ps)) (length (compE2 obj) + length (compEs2 ps)) None" using lenvs by(auto simp add: \move2_iff) from lenvs have "P,obj\M'(ps),h \ (THROW NullPointer, loc) \ (rev vs @ [Null], loc, length (compE2 obj) + length (compEs2 ps), \addr_of_sys_xcpt NullPointer\)" by(safe intro!: bisim1CallThrow) simp moreover have "True,P,t \1 \null\M'(map Val vs),(h, loc)\ -\\ \THROW NullPointer,(h, loc)\" by(rule Red1CallNull) ultimately show ?thesis using exec pc \ lenvs reds apply(auto elim!: exec_meth.cases simp add: exec_move_def) apply(rule exI conjI|assumption)+ apply(rule rtranclp.rtrancl_into_rtrancl) apply(erule Call_\red1r_param) by auto next assume ?rest then obtain a Ta C' Ts' Tr' D' where [simp]: "v = Addr a" and Ta: "typeof_addr h a = \Ta\" and iec: "P \ class_type_of Ta sees M': Ts'\Tr' = Native in D'" by auto with exec pc lenvs obtain ta' va h'' U Ts Ts' where redex: "(ta', va, h'') \ red_external_aggr (compP2 P) t a M' vs h" and ret: "(xcp', h', [(stk', loc', undefined, undefined, pc')]) = extRet2JVM (length vs) h'' (rev vs @ [Addr a]) loc undefined undefined (length (compE2 obj) + length (compEs2 ps)) [] va" and wtext': "P,h \ a\M'(vs) : U" and Ts: "map typeof\<^bsub>h\<^esub> vs = map Some Ts" and ta': "ta = extTA2JVM (compP2 P) ta'" by(fastforce simp add: is_Ref_def exec_move_def compP2_def external_WT'_iff exec_meth_instr confs_conv_map) have \: "\move1 P h (Val v\M'(map Val vs)) \ \move2 (compP2 P) h (stk @ [v]) (obj\M'(ps)) (length (compE2 obj) + length (compEs2 ps)) None" using Ta iec lenvs by(auto simp add: \move2_iff map_eq_append_conv compP2_def) from ret have [simp]: "h'' = h'" by simp from wtext' have wtext'': "compP2 P,h \ a\M'(vs) : U" by(simp add: compP2_def external_WT'_iff) from wf have "wf_jvm_prog (compP2 P)" by(rule wt_compP2) then obtain wf_md where wf': "wf_prog wf_md (compP2 P)" unfolding wf_jvm_prog_def by(blast dest: wt_jvm_progD) from tconf have tconf': "compP2 P,h \ t \t" by(simp add: compP2_def tconf_def) from redex have redex'': "compP2 P,t \ \a\M'(vs), h\ -ta'\ext \va, h'\" by(simp add: WT_red_external_list_conv[OF wf' wtext'' tconf', symmetric]) hence redex': "P,t \ \a\M'(vs), h\ -ta'\ext \va, h'\" by(simp add: compP2_def) with Ta iec have "True,P,t \1 \addr a\M'(map Val vs), (h, loc)\ -ta'\ \extRet2J (addr a\M'(map Val vs)) va, (h', loc)\" by -(rule Red1CallExternal, auto) moreover have "P,obj\M'(ps),h' \ (extRet2J (addr a\M'(map Val vs)) va, loc) \ (stk', loc', pc', xcp')" proof(cases va) case (RetVal v) have "P,obj\M'(ps),h' \ (Val v, loc) \ ([v], loc, length (compE2 (obj\M'(ps))), None)" by(rule bisim1Val2)(simp) with ret RetVal show ?thesis by simp next case (RetExc ad) from lenvs have "length ps = length (rev vs)" by simp hence "P,obj\M'(ps),h' \ (Throw ad, loc) \ (rev vs @ [v], loc, length (compE2 (obj)) + length (compEs2 ps), \ad\)" by(rule bisim1CallThrow) with ret RetExc show ?thesis by simp next case RetStaySame from lenvs have "length ps = length vs" by simp from bisims1_map_Val_append[OF bisims1Nil this, of P h' loc] have "P,ps,h' \ (map Val vs, loc) [\] (rev vs, loc, length (compEs2 ps), None)" by simp hence "P,obj\M'(ps),h' \ (addr a\M'(map Val vs), loc) \ (rev vs @ [Addr a], loc, length (compE2 obj) + length (compEs2 ps), None)" by(rule bisim1_bisims1.bisim1CallParams) thus ?thesis using ret RetStaySame by simp qed moreover from \preallocated h\ red_external_hext[OF redex'] have "preallocated h'" by(rule preallocated_hext) from redex'' wtext'' \hconf h\ have "hconf h'" by(rule external_call_hconf) with ta' redex' \preallocated h'\ have "ta_bisim wbisim1 (extTA2J1 P ta') ta" by(auto intro: red_external_ta_bisim21[OF wf]) moreover have "\move1 P h (Val v\M'(map Val vs)) \ ta' = \ \ h' = h" using Ta iec redex' by(fastforce dest: \external'_red_external_heap_unchanged \external'_red_external_TA_empty sees_method_fun simp add: \external'_def \external_def map_eq_append_conv) moreover from vs call have "call1 (Val v\M'(ps')) \ None \ ps' = map Val vs \ loc = xs" by(auto split: if_split_asm simp add: is_vals_conv) ultimately show ?thesis using reds \ pc Ta call apply(auto simp del: split_paired_Ex call1.simps calls1.simps split: if_split_asm simp add: map_eq_append_conv) apply(auto 4 4 simp del: split_paired_Ex call1.simps calls1.simps intro: rtranclp.rtrancl_into_rtrancl[OF Call_\red1r_param] rtranclp_into_tranclp1[OF Call_\red1r_param])[3] apply((assumption|rule exI conjI|erule Call_\red1r_param|simp add: map_eq_append_conv)+) done qed qed next case (bisim1CallThrowObj obj n a xs stk loc pc ps M') note exec = \?exec (obj\M'(ps)) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,obj,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 obj)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 obj 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: exec_move_def match_ex_table_not_pcs_None) thus ?case .. next case (bisim1CallThrowParams ps n vs a ps' xs stk loc pc obj M' v) note exec = \?exec (obj\M'(ps)) (stk @ [v]) loc (length (compE2 obj) + pc) \a\ stk' loc' pc' xcp'\ note bisim2 = \P,ps,h \ (map Val vs @ Throw a # ps', xs) [\] (stk, loc, pc, \a\)\ from bisim2 have pc: "pc < length (compEs2 ps)" by(auto dest: bisims1_ThrowD) from bisim2 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxEs2 ps 0 0) = None" unfolding compP2_def by(rule bisims1_xcp_Some_not_caught) with exec pc have False apply(auto elim!: exec_meth.cases simp add: compxEs2_size_convs compxEs2_stack_xlift_convs exec_move_def) apply(auto simp add: match_ex_table_append dest!: match_ex_table_shift_pcD match_ex_table_stack_xliftD match_ex_table_pc_length_compE2) done thus ?case .. next case (bisim1CallThrow ps vs obj n M' a xs v) note lenvs = \length ps = length vs\ note exec = \?exec (obj\M'(ps)) (vs @ [v]) xs (length (compE2 obj) + length (compEs2 ps)) \a\ stk' loc' pc' xcp'\ with lenvs have False by(auto elim!: exec_meth.cases simp add: match_ex_table_append_not_pcs exec_move_def dest!: match_ex_table_pc_length_compEs2) thus ?case .. next case (bisim1BlockSome1 e n V Ty v xs) have "\move2 (compP2 P) h [] {V:Ty=\v\; e} 0 None" by(simp add: \move2_iff) with \?exec {V:Ty=\v\; e} [] xs 0 None stk' loc' pc' xcp'\ \P,e,h \ (e, xs) \ ([], xs, 0, None)\ show ?case by(fastforce elim!: exec_meth.cases intro: bisim1BlockSome2 simp add: exec_move_def) next case (bisim1BlockSome2 e n V Ty v xs) have "\move2 (compP2 P) h [v] {V:Ty=\v\; e} (Suc 0) None" by(simp add: \move2_iff) with \?exec {V:Ty=\v\; e} [v] xs (Suc 0) None stk' loc' pc' xcp'\ \P,e,h \ (e, xs) \ ([], xs, 0, None)\ show ?case by(fastforce intro: r_into_rtranclp Block1Some bisim1BlockSome4[OF bisim1_refl] simp add: exec_move_def elim!: exec_meth.cases) next case (bisim1BlockSome4 e n e' xs stk loc pc xcp V Ty v) note IH = bisim1BlockSome4.IH(2) note exec = \?exec {V:Ty=\v\; e} stk loc (Suc (Suc pc)) xcp stk' loc' pc' xcp'\ note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bsok = \bsok {V:Ty=\v\; e} n\ with \n + max_vars {V:Ty=None; e'} \ length xs\ have V: "V < length xs" and len: "Suc n + max_vars e' \ length xs" and [simp]: "n = V" by simp_all let ?pre = "[Push v, Store V]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e) (shift (length ?pre) (compxE2 e 0 0)) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs exec_move_def) hence pc': "pc' \ length ?pre" by(rule exec_meth_drop_pc) auto hence pc'': "Suc (Suc (pc' - Suc (Suc 0))) = pc'" by simp moreover from exec' have "exec_move_d P t e h (stk, loc, pc, xcp) ta h' (stk', loc', pc' - length ?pre, xcp')" unfolding exec_move_def by(rule exec_meth_drop) auto from IH[OF this len _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bsok obtain e'' xs'' where bisim': "P,e,h' \ (e'', xs'') \ (stk', loc', pc' - length ?pre, xcp')" and red': "?red e' xs e'' xs'' e stk pc (pc' - length ?pre) xcp xcp'" by auto from bisim' have "P,{V:Ty=\v\; e},h' \ ({V:Ty=None; e''}, xs'') \ (stk', loc', Suc (Suc (pc' - length ?pre)), xcp')" by(rule bisim1_bisims1.bisim1BlockSome4) moreover from pc' pc'' have "\move2 (compP2 P) h stk {V:Ty=\v\; e} (Suc (Suc pc)) xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff) moreover from red' have "length xs'' = length xs" by(auto split: if_split_asm dest!: \red1r_preserves_len \red1t_preserves_len red1_preserves_len) ultimately show ?case using red' pc'' V by(fastforce elim!: Block_None_\red1r_xt Block_None_\red1t_xt intro: Block1Red split: if_split_asm simp add: no_call2_def) next case (bisim1BlockThrowSome e n a xs stk loc pc V Ty v) note exec = \?exec {V:Ty=\v\; e} stk loc (Suc (Suc pc)) \a\ stk' loc' pc' xcp'\ note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False apply(auto elim!: exec_meth.cases simp add: match_ex_table_not_pcs_None exec_move_def) apply(auto simp only: compxE2_size_convs dest!: match_ex_table_shift_pcD) apply simp done thus ?case .. next case (bisim1BlockNone e n e' xs stk loc pc xcp V Ty) note IH = bisim1BlockNone.IH(2) note exec = \?exec {V:Ty=None; e} stk loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bsok = \bsok {V:Ty=None; e} n\ with \n + max_vars {V:Ty=None; e'} \ length xs\ have V: "V < length xs" and len: "Suc n + max_vars e' \ length xs" by simp_all have "\move2 (compP2 P) h stk {V:Ty=None; e} pc xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff) moreover from exec have "?exec e stk loc pc xcp stk' loc' pc' xcp'" by(simp add: exec_move_BlockNone) from IH[OF this len _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bsok obtain e'' xs'' where bisim': "P,e,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red': "?red e' xs e'' xs'' e stk pc pc' xcp xcp'" by auto from bisim' have "P,{V:Ty=None; e},h' \ ({V:Ty=None; e''}, xs'') \ (stk', loc', pc', xcp')" by(rule bisim1_bisims1.bisim1BlockNone) ultimately show ?case using V red' by(fastforce elim!: Block1Red Block_None_\red1r_xt Block_None_\red1t_xt simp add: no_call2_def) next case (bisim1BlockThrowNone e n a xs stk loc pc V Ty) note exec = \?exec {V:Ty=None; e} stk loc pc \a\ stk' loc' pc' xcp'\ note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: exec_move_def) thus ?case .. next case (bisim1Sync1 e1 n e' xs stk loc pc xcp e2 V) note IH = bisim1Sync1.IH(2) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) stk loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e1,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,e2,h \ (e2, xs) \ ([], xs, 0, None)\ note len = \n + max_vars (sync\<^bsub>V\<^esub> (e') e2) \ length xs\ note bsok = \bsok (sync\<^bsub>V\<^esub> (e1) e2) n\ from bisim have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e1)") case True with exec have exec': "?exec e1 stk loc pc xcp stk' loc' pc' xcp'" by(simp add: exec_move_Sync1) from True have "\move2 (compP2 P) h stk (sync\<^bsub>V\<^esub> (e1) e2) pc xcp = \move2 (compP2 P) h stk e1 pc xcp" by(simp add: \move2_iff) with IH[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bisim2 bsok show ?thesis by(fastforce intro: bisim1_bisims1.bisim1Sync1 Synchronized1Red1 elim!: Sync_\red1r_xt Sync_\red1t_xt simp add: no_call2_def) next case False with pc have [simp]: "pc = length (compE2 e1)" by simp with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (sync\<^bsub>V\<^esub> (e') e2, xs) (sync\<^bsub>V\<^esub> (Val v) e2, loc)" by(rule Sync_\red1r_xt) moreover have \: "\move2 (compP2 P) h [v] (sync\<^bsub>V\<^esub> (e1) e2) pc None" by(simp add: \move2_iff) moreover have "P,(sync\<^bsub>V\<^esub> (e1) e2),h \ ((sync\<^bsub>V\<^esub> (Val v) e2), loc) \ ([v, v], loc, Suc (length (compE2 e1)), None)" by(rule bisim1Sync2) ultimately show ?thesis using exec stk xcp by(fastforce elim!: exec_meth.cases simp add: exec_move_def) qed next case (bisim1Sync2 e1 n e2 V v xs) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [v, v] xs (Suc (length (compE2 e1))) None stk' loc' pc' xcp'\ note bisim = \P,e1,h \ (e1, xs) \ ([], xs, 0, None)\ note bisim2 = \P,e2,h \ (e2, xs) \ ([], xs, 0, None)\ have "\move2 (compP2 P) h [v, v] (sync\<^bsub>V\<^esub> (e1) e2) (Suc (length (compE2 e1))) None" by(rule \move2Sync3) thus ?case using exec by(fastforce elim!: exec_meth.cases intro: bisim1Sync3 simp add: exec_move_def) next case (bisim1Sync3 e1 n e2 V v xs) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [v] (xs[V := v]) (Suc (Suc (length (compE2 e1)))) None stk' loc' pc' xcp'\ note bisim = \P,e1,h \ (e1, xs) \ ([], xs, 0, None)\ note bisim2 = \\xs. P,e2,h \ (e2, xs) \ ([], xs, 0, None)\ note len = \n + max_vars (sync\<^bsub>V\<^esub> (Val v) e2) \ length xs\ note bsok = \bsok (sync\<^bsub>V\<^esub> (e1) e2) n\ with len have V: "V < length xs" by simp have \: "\ \move2 (compP2 P) h [v] (sync\<^bsub>V\<^esub> (e1) e2) (Suc (Suc (length (compE2 e1)))) None" by(simp add: \move2_iff) from exec have "(\a. v = Addr a \ stk' = [] \ loc' = xs[V := v] \ ta = \Lock\a, SyncLock a\ \ xcp' = None \ pc' = Suc (Suc (Suc (length (compE2 e1))))) \ (v = Null \ stk' = [v] \ loc' = xs[V := v] \ ta = \ \ xcp' = \addr_of_sys_xcpt NullPointer\ \ pc' = Suc (Suc (length (compE2 e1))))" (is "?c1 \ ?c2") by(fastforce elim!: exec_meth.cases simp add: is_Ref_def expand_finfun_eq fun_eq_iff finfun_upd_apply exec_move_def) thus ?case proof assume ?c1 then obtain a where [simp]: "v = Addr a" "stk' = []" "loc' = xs[V := v]" "ta = \Lock\a, SyncLock a\" "xcp' = None" "pc' = Suc (Suc (Suc (length (compE2 e1))))" by blast have "True,P,t \1 \sync\<^bsub>V\<^esub> (addr a) e2, (h, xs)\ -\Lock\a, SyncLock a\\ \insync\<^bsub>V\<^esub> (a) e2,(h, xs[V := Addr a])\" using V by(rule Lock1Synchronized) moreover from bisim2 have "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (insync\<^bsub>V\<^esub> (a) e2, xs[V := Addr a]) \ ([], xs[V := Addr a], Suc (Suc (Suc (length (compE2 e1)))), None)" by(auto intro: bisim1Sync4[where pc = 0, simplified]) ultimately show ?case using exec \ by(fastforce elim!: exec_meth.cases split: if_split_asm simp add: is_Ref_def exec_move_def ta_bisim_def ta_upd_simps) next assume ?c2 hence [simp]: "v = Null" "stk' = [v]" "loc' = xs[V := v]" "ta = \" "xcp' = \addr_of_sys_xcpt NullPointer\" "pc' = Suc (Suc (length (compE2 e1)))" by simp_all from V have "True,P,t \1 \sync\<^bsub>V\<^esub> (null) e2, (h, xs)\ -\\ \THROW NullPointer,(h, xs[V := Null])\" by(rule Synchronized1Null) moreover have "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (THROW NullPointer, xs[V := Null]) \ ([Null], xs[V := Null], Suc (Suc (length (compE2 e1))), \addr_of_sys_xcpt NullPointer\)" by(rule bisim1Sync11) ultimately show ?case using exec \ by(fastforce elim!: exec_meth.cases split: if_split_asm simp add: is_Ref_def exec_move_def) qed next case (bisim1Sync4 e2 n e' xs stk loc pc xcp e1 V a) note IH = bisim1Sync4.IH(2) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) stk loc (Suc (Suc (Suc (length (compE2 e1) + pc)))) xcp stk' loc' pc' xcp'\ note bisim1 = \P,e1,h \ (e1, xs) \ ([], xs, 0, None)\ note bisim2 = \P,e2,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (insync\<^bsub>V\<^esub> (a) e') \ length xs\ note bsok = \bsok (sync\<^bsub>V\<^esub> (e1) e2) n\ with len have V: "V < length xs" and len': "Suc n + max_vars e' \ length xs" by simp_all from bisim2 have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) let ?pre = "compE2 e1 @ [Dup, Store V, MEnter]" let ?post = "[Load V, MExit, Goto 4, Load V, MExit, ThrowExc]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e2 @ ?post) (compxE2 e1 0 0 @ shift (length ?pre) (compxE2 e2 0 0 @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), 0)])) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: ac_simps eval_nat_numeral shift_compxE2 exec_move_def) hence pc': "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc[where n'=1]) auto from exec' have exec'': "exec_meth_d (compP2 P) (compE2 e2 @ ?post) (compxE2 e2 0 0 @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), 0)]) t h (stk, loc, pc, xcp) ta h' (stk', loc', pc' - length ?pre, xcp')" by(rule exec_meth_drop_xt[where n=1]) auto show ?case proof(cases "pc < length (compE2 e2)") case True note pc = True hence \: "\move2 (compP2 P) h stk (sync\<^bsub>V\<^esub> (e1) e2) (Suc (Suc (Suc (length (compE2 e1) + pc)))) xcp = \move2 (compP2 P) h stk e2 pc xcp" by(simp add: \move2_iff) show ?thesis proof(cases "xcp = None \ (\a'. xcp = \a'\ \ match_ex_table (compP2 P) (cname_of h a') pc (compxE2 e2 0 0) \ None)") case False then obtain a' where Some: "xcp = \a'\" and True: "match_ex_table (compP2 P) (cname_of h a') pc (compxE2 e2 0 0) = None" by(auto simp del: not_None_eq) from Some \conf_xcp' (compP2 P) h xcp\ obtain D where ha': "typeof_addr h a' = \Class_type D\" and subcls: "P \ D \\<^sup>* Throwable" by(auto simp add: compP2_def) from ha' subcls bisim2 True bsok have "\red1r P t h (e', xs) (Throw a', loc)" using len' unfolding Some compP2_def by(auto dest!: bisim1_xcp_\Red simp add: bsok_def) moreover from pc have "\move2 (compP2 P) h stk (sync\<^bsub>V\<^esub> (e1) e2) (Suc (Suc (Suc (pc + length (compE2 e1))))) \a'\" by(auto intro: \move2xcp) moreover have "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (insync\<^bsub>V\<^esub> (a) Throw a', loc) \ ([Addr a'], loc, 6 + length (compE2 e1) + length (compE2 e2), None)" by(rule bisim1Sync7) ultimately show ?thesis using exec True pc Some ha' subcls apply(auto elim!: exec_meth.cases simp add: ac_simps eval_nat_numeral match_ex_table_append matches_ex_entry_def compP2_def exec_move_def) apply(simp_all only: compxE2_size_convs, auto dest: match_ex_table_shift_pcD match_ex_table_pc_length_compE2) apply(fastforce elim!: InSync_\red1r_xt) done next case True with exec'' pc have "exec_meth_d (compP2 P) (compE2 e2 @ ?post) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk', loc', pc' - length ?pre, xcp')" by(auto elim!: exec_meth.cases intro: exec_meth.intros simp add: match_ex_table_append exec_move_def) hence "exec_move_d P t e2 h (stk, loc, pc, xcp) ta h' (stk', loc', pc' - length ?pre, xcp')" using pc unfolding exec_move_def by(rule exec_meth_take) from IH[OF this len' _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bsok obtain e'' xs'' where bisim': "P,e2,h' \ (e'', xs'') \ (stk', loc', pc' - length ?pre, xcp')" and red': "?red e' xs e'' xs'' e2 stk pc (pc' - length ?pre) xcp xcp'" by auto from bisim' have "P,sync\<^bsub>V\<^esub> (e1) e2, h' \ (insync\<^bsub>V\<^esub> (a) e'', xs'') \ (stk', loc', Suc (Suc (Suc (length (compE2 e1) + (pc' - length ?pre)))), xcp')" by(rule bisim1_bisims1.bisim1Sync4) moreover from pc' have "Suc (Suc (Suc (length (compE2 e1) + (pc' - Suc (Suc (Suc (length (compE2 e1)))))))) = pc'" "Suc (Suc (Suc (pc' - Suc (Suc (Suc 0))))) = pc'" by simp_all ultimately show ?thesis using red' \ by(fastforce intro: Synchronized1Red2 simp add: eval_nat_numeral no_call2_def elim!: InSync_\red1r_xt InSync_\red1t_xt split: if_split_asm) qed next case False with pc have [simp]: "pc = length (compE2 e2)" by simp with bisim2 obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) have "\move2 (compP2 P) h [v] (sync\<^bsub>V\<^esub> (e1) e2) (Suc (Suc (Suc (length (compE2 e1) + length (compE2 e2))))) None" by(simp add: \move2_iff) moreover from bisim2 pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro!: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (insync\<^bsub>V\<^esub> (a) e', xs) (insync\<^bsub>V\<^esub> (a) (Val v), loc)" by(rule InSync_\red1r_xt) moreover have "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (insync\<^bsub>V\<^esub> (a) (Val v), loc) \ ([loc ! V, v], loc, 4 + length (compE2 e1) + length (compE2 e2), None)" by(rule bisim1Sync5) ultimately show ?thesis using exec by(auto elim!: exec_meth.cases simp add: eval_nat_numeral exec_move_def) blast qed next case (bisim1Sync5 e1 n e2 V a v xs) note bisim1 = \P,e1,h \ (e1, xs) \ ([], xs, 0, None)\ note bisim2 = \P,e2,h \ (e2, xs) \ ([], xs, 0, None)\ note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [xs ! V, v] xs (4 + length (compE2 e1) + length (compE2 e2)) None stk' loc' pc' xcp'\ from \n + max_vars (insync\<^bsub>V\<^esub> (a) Val v) \ length xs\ \bsok (sync\<^bsub>V\<^esub> (e1) e2) n\ have V: "V < length xs" by simp have \: "\ \move2 (compP2 P) h [xs ! V, v] (sync\<^bsub>V\<^esub> (e1) e2) (4 + length (compE2 e1) + length (compE2 e2)) None" by(simp add: \move2_iff) have \': "\ \move1 P h (insync\<^bsub>V\<^esub> (a) Val v)" by auto from exec have "(\a'. xs ! V = Addr a') \ xs ! V = Null" (is "?c1 \ ?c2") by(auto elim!: exec_meth.cases simp add: split_beta is_Ref_def exec_move_def split: if_split_asm) thus ?case proof assume ?c1 then obtain a' where xsV [simp]: "xs ! V = Addr a'" .. have "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (Val v, xs) \ ([v], xs, 5 + length (compE2 e1) + length (compE2 e2), None)" "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (THROW IllegalMonitorState,xs) \ ([Addr a', v],xs,4 + length (compE2 e1) + length (compE2 e2),\addr_of_sys_xcpt IllegalMonitorState\)" by(rule bisim1Sync6, rule bisim1Sync12) moreover from xsV V have "True,P,t \1 \insync\<^bsub>V\<^esub> (a) Val v, (h, xs)\ -\Unlock\a', SyncUnlock a'\\ \Val v,(h, xs)\" by(rule Unlock1Synchronized) moreover from xsV V have "True,P,t \1 \insync\<^bsub>V\<^esub> (a) Val v, (h, xs)\ -\UnlockFail\a'\\ \THROW IllegalMonitorState,(h, xs)\" by(rule Unlock1SynchronizedFail[OF TrueI]) ultimately show ?case using \ \' exec - by(fastforce elim!: exec_meth.cases simp add: is_Ref_def ta_bisim_def ac_simps exec_move_def ta_upd_simps) + by (fastforce elim!: exec_meth.cases simp add: is_Ref_def ta_bisim_def exec_move_def ac_simps ta_upd_simps + simp del: conj.left_commute) next assume xsV: "xs ! V = Null" have "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (THROW NullPointer,xs) \ ([Null, v],xs,4 + length (compE2 e1) + length (compE2 e2),\addr_of_sys_xcpt NullPointer\)" by(rule bisim1Sync12) thus ?case using \ \' exec xsV V - by(fastforce elim!: exec_meth.cases intro: Unlock1SynchronizedNull simp add: is_Ref_def ta_bisim_def ac_simps exec_move_def) + by (fastforce elim!: exec_meth.cases intro: Unlock1SynchronizedNull simp add: is_Ref_def ta_bisim_def ac_simps exec_move_def + simp del: conj.left_commute) qed next case (bisim1Sync6 e1 n e2 V v x) note bisim1 = \P,e1,h \ (e1, xs) \ ([], xs, 0, None)\ note bisim2 = \P,e2,h \ (e2, xs) \ ([], xs, 0, None)\ note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [v] x (5 + length (compE2 e1) + length (compE2 e2)) None stk' loc' pc' xcp'\ have "\move2 (compP2 P) h [v] (sync\<^bsub>V\<^esub> (e1) e2) (5 + length (compE2 e1) + length (compE2 e2)) None" by(simp add: \move2_iff) moreover have "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (Val v, x) \ ([v], x, length (compE2 (sync\<^bsub>V\<^esub> (e1) e2)), None)" by(rule bisim1Val2) simp moreover have "nat (9 + (int (length (compE2 e1)) + int (length (compE2 e2)))) = 9 + length (compE2 e1) + length (compE2 e2)" by arith ultimately show ?case using exec by(fastforce elim!: exec_meth.cases simp add: eval_nat_numeral exec_move_def) next case (bisim1Sync7 e1 n e2 V a a' xs) note \?exec (sync\<^bsub>V\<^esub> (e1) e2) [Addr a'] xs (6 + length (compE2 e1) + length (compE2 e2)) None stk' loc' pc' xcp'\ moreover have "P,sync\<^bsub>V\<^esub> (e1) e2,h' \ (insync\<^bsub>V\<^esub> (a) Throw a', xs) \ ([xs ! V, Addr a'], xs, 7 + length (compE2 e1) + length (compE2 e2), None)" by(rule bisim1Sync8) moreover have "\move2 (compP2 P) h [Addr a'] (sync\<^bsub>V\<^esub> (e1) e2) (6 + length (compE2 e1) + length (compE2 e2)) None" by(simp add: \move2_iff) ultimately show ?case by(fastforce elim!: exec_meth.cases simp add: eval_nat_numeral exec_move_def) next case (bisim1Sync8 e1 n e2 V a a' xs) from \n + max_vars (insync\<^bsub>V\<^esub> (a) Throw a') \ length xs\ \bsok (sync\<^bsub>V\<^esub> (e1) e2) n\ have V: "V < length xs" by simp note \?exec (sync\<^bsub>V\<^esub> (e1) e2) [xs ! V, Addr a'] xs (7 + length (compE2 e1) + length (compE2 e2)) None stk' loc' pc' xcp'\ moreover have "\ \move2 (compP2 P) h [xs ! V, Addr a'] (sync\<^bsub>V\<^esub> (e1) e2) (7 + length (compE2 e1) + length (compE2 e2)) None" by(simp add: \move2_iff) moreover have "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (Throw a', xs) \ ([Addr a'], xs, 8 + length (compE2 e1) + length (compE2 e2), None)" "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (THROW IllegalMonitorState,xs) \ ([xs ! V, Addr a'],xs,7 + length (compE2 e1) + length (compE2 e2),\addr_of_sys_xcpt IllegalMonitorState\)" "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (THROW NullPointer,xs) \ ([Null, Addr a'],xs,7 + length (compE2 e1) + length (compE2 e2),\addr_of_sys_xcpt NullPointer\)" by(rule bisim1Sync9 bisim1Sync14)+ moreover { fix A assume "xs ! V = Addr A" hence "True,P,t \1 \insync\<^bsub>V\<^esub> (a) Throw a',(h, xs)\ -\Unlock\A, SyncUnlock A\\ \Throw a', (h, xs)\" "True,P,t \1 \insync\<^bsub>V\<^esub> (a) Throw a',(h, xs)\ -\UnlockFail\A\\ \THROW IllegalMonitorState, (h, xs)\" using V by(rule Synchronized1Throw2 Synchronized1Throw2Fail[OF TrueI])+ } moreover { assume "xs ! V = Null" hence "True,P,t \1 \insync\<^bsub>V\<^esub> (a) Throw a',(h, xs)\ -\\ \THROW NullPointer, (h, xs)\" using V by(rule Synchronized1Throw2Null) } moreover have "\ \move1 P h (insync\<^bsub>V\<^esub> (a) Throw a')" by fastforce ultimately show ?case by(fastforce elim!: exec_meth.cases simp add: eval_nat_numeral is_Ref_def ta_bisim_def ta_upd_simps exec_move_def split: if_split_asm) next case (bisim1Sync9 e1 n e2 V a xs) note \?exec (sync\<^bsub>V\<^esub> (e1) e2) [Addr a] xs (8 + length (compE2 e1) + length (compE2 e2)) None stk' loc' pc' xcp'\ moreover have "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (Throw a, xs) \ ([Addr a], xs, 8 + length (compE2 e1) + length (compE2 e2), \a\)" by(rule bisim1Sync10) moreover have "\move2 (compP2 P) h [Addr a] (sync\<^bsub>V\<^esub> (e1) e2) (8 + length (compE2 e1) + length (compE2 e2)) None" by(rule \move2Sync8) ultimately show ?case by(fastforce elim!: exec_meth.cases simp add: eval_nat_numeral exec_move_def split: if_split_asm) next case (bisim1Sync10 e1 n e2 V a xs) from \?exec (sync\<^bsub>V\<^esub> (e1) e2) [Addr a] xs (8 + length (compE2 e1) + length (compE2 e2)) \a\ stk' loc' pc' xcp'\ have False by(auto elim!: exec_meth.cases simp add: matches_ex_entry_def match_ex_table_append_not_pcs exec_move_def) thus ?case .. next case (bisim1Sync11 e1 n e2 V xs) from \?exec (sync\<^bsub>V\<^esub> (e1) e2) [Null] xs (Suc (Suc (length (compE2 e1)))) \addr_of_sys_xcpt NullPointer\ stk' loc' pc' xcp'\ have False by(auto elim!: exec_meth.cases simp add: matches_ex_entry_def match_ex_table_append_not_pcs exec_move_def) thus ?case .. next case (bisim1Sync12 e1 n e2 V a xs v v') from \?exec (sync\<^bsub>V\<^esub> (e1) e2) [v, v'] xs (4 + length (compE2 e1) + length (compE2 e2)) \a\ stk' loc' pc' xcp'\ have False by(auto elim!: exec_meth.cases simp add: matches_ex_entry_def match_ex_table_append_not_pcs exec_move_def) thus ?case .. next case (bisim1Sync14 e1 n e2 V a xs v a') from \?exec (sync\<^bsub>V\<^esub> (e1) e2) [v, Addr a'] xs (7 + length (compE2 e1) + length (compE2 e2)) \a\ stk' loc' pc' xcp'\ have False by(auto elim!: exec_meth.cases simp add: matches_ex_entry_def match_ex_table_append_not_pcs exec_move_def) thus ?case .. next case bisim1InSync thus ?case by simp next case (bisim1SyncThrow e1 n a xs stk loc pc e2 V) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,e1,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 e1)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e1 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False apply(auto elim!: exec_meth.cases simp add: match_ex_table_append_not_pcs exec_move_def) apply(auto dest!: match_ex_table_shift_pcD simp add: matches_ex_entry_def) done thus ?case .. next case (bisim1Seq1 e1 n e' xs stk loc pc xcp e2) note IH = bisim1Seq1.IH(2) note exec = \?exec (e1;; e2) stk loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e1,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (e';;e2) \ length xs\ note bsok = \bsok (e1;; e2) n\ from bisim have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e1)") case True let ?post = "Pop # compE2 e2" from exec have exec': "?exec e1 stk loc pc xcp stk' loc' pc' xcp'" using True by(simp add: exec_move_Seq1) from True have "\move2 (compP2 P) h stk (e1;;e2) pc xcp = \move2 (compP2 P) h stk e1 pc xcp" by(simp add: \move2_iff) with IH[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bsok show ?thesis by(fastforce intro: bisim1_bisims1.bisim1Seq1 Seq1Red elim!: Seq_\red1r_xt Seq_\red1t_xt simp add: no_call2_def) next case False with pc have [simp]: "pc = length (compE2 e1)" by simp with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (e';; e2, xs) (Val v;;e2, loc)" by(rule Seq_\red1r_xt) also have "\move1 P h (Val v;;e2)" by(rule \move1SeqRed) hence "\red1r P t h (Val v;;e2, loc) (e2, loc)" by(auto intro: Red1Seq r_into_rtranclp) also have \: "\move2 (compP2 P) h [v] (e1;;e2) pc None" by(simp add: \move2_iff) moreover from \P,e2,h \ (e2, loc) \ ([], loc, 0, None)\ have "P,e1;;e2,h \ (e2, loc) \ ([], loc, Suc (length (compE2 e1) + 0), None)" by(rule bisim1Seq2) ultimately show ?thesis using exec stk xcp by(fastforce elim!: exec_meth.cases simp add: exec_move_def) qed next case (bisim1SeqThrow1 e1 n a xs stk loc pc e2) note exec = \?exec (e1;;e2) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,e1,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 e1)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e1 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: match_ex_table_not_pcs_None exec_move_def) thus ?case .. next case (bisim1Seq2 e2 n e' xs stk loc pc xcp e1) note IH = bisim1Seq2.IH(2) note bisim1 = \\xs. P,e1,h \ (e1, xs) \ ([], xs, 0, None)\ note bisim2 = \P,e2,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars e' \ length xs\ note exec = \?exec (e1;; e2) stk loc (Suc (length (compE2 e1) + pc)) xcp stk' loc' pc' xcp'\ note bsok = \bsok (e1;; e2) n\ let ?pre = "compE2 e1 @ [Pop]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e2) (compxE2 e1 0 0 @ shift (length ?pre) (compxE2 e2 0 0)) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 exec_move_def) hence "?exec e2 stk loc pc xcp stk' loc' (pc' - length ?pre) xcp'" unfolding exec_move_def by(rule exec_meth_drop_xt, auto) from IH[OF this len _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bsok obtain e'' xs'' where bisim': "P,e2,h' \ (e'', xs'') \ (stk', loc', pc' - length ?pre, xcp')" and red: "?red e' xs e'' xs'' e2 stk pc (pc' - length ?pre) xcp xcp'" by auto from bisim' have "P,e1;;e2,h' \ (e'', xs'') \ (stk', loc', Suc (length (compE2 e1) + (pc' - length ?pre)), xcp')" by(rule bisim1_bisims1.bisim1Seq2) moreover have \: "\move2 (compP2 P) h stk (e1;;e2) (Suc (length (compE2 e1) + pc)) xcp = \move2 (compP2 P) h stk e2 pc xcp" by(simp add: \move2_iff) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc) auto ultimately show ?case using red by(fastforce split: if_split_asm simp add: no_call2_def) next case (bisim1Cond1 e n e' xs stk loc pc xcp e1 e2) note IH = bisim1Cond1.IH(2) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bisim1 = \\xs. P,e1,h \ (e1, xs) \ ([], xs, 0, None)\ note bisim2 = \\xs. P,e2,h \ (e2, xs) \ ([], xs, 0, None)\ from \n + max_vars (if (e') e1 else e2) \ length xs\ have len: "n + max_vars e' \ length xs" by simp note exec = \?exec (if (e) e1 else e2) stk loc pc xcp stk' loc' pc' xcp'\ note bsok = \bsok (if (e) e1 else e2) n\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True let ?post = "IfFalse (2 + int (length (compE2 e1))) # compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2" from exec have exec': "?exec e stk loc pc xcp stk' loc' pc' xcp'" using True by(simp add: exec_move_Cond1) from True have "\move2 (compP2 P) h stk (if (e) e1 else e2) pc xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff) with IH[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bsok show ?thesis by(fastforce intro: bisim1_bisims1.bisim1Cond1 Cond1Red elim!: Cond_\red1r_xt Cond_\red1t_xt simp add: no_call2_def) next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (if (e') e1 else e2, xs) (if (Val v) e1 else e2, loc)" by(rule Cond_\red1r_xt) moreover have "\move1 P h (if (Val v) e1 else e2)" by(rule \move1CondRed) moreover have \: "\move2 (compP2 P) h [v] (if (e) e1 else e2) pc None" by(simp add: \move2_iff) moreover from bisim1[of loc] have "P,if (e) e1 else e2,h \ (e1, loc) \ ([], loc, Suc (length (compE2 e) + 0), None)" by(rule bisim1CondThen) moreover from bisim2[of loc] have "P,if (e) e1 else e2,h \ (e2, loc) \ ([], loc, Suc (Suc (length (compE2 e) + length (compE2 e1) + 0)), None)" by(rule bisim1CondElse) moreover have "nat (int (length (compE2 e)) + (2 + int (length (compE2 e1)))) = Suc (Suc (length (compE2 e) + length (compE2 e1) + 0))" by simp moreover from exec xcp stk have "typeof\<^bsub>h\<^esub> v = \Boolean\" by(auto simp add: exec_move_def exec_meth_instr) ultimately show ?thesis using exec stk xcp by(fastforce elim!: exec_meth.cases intro: Red1CondT Red1CondF elim!: rtranclp.rtrancl_into_rtrancl simp add: eval_nat_numeral exec_move_def) qed next case (bisim1CondThen e1 n e' xs stk loc pc xcp e e2) note IH = bisim1CondThen.IH(2) note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bisim = \\xs. P,e,h \ (e, xs) \ ([], xs, 0, None)\ note bisim2 = \\xs. P,e2,h \ (e2, xs) \ ([], xs, 0, None)\ note len = \n + max_vars e' \ length xs\ note exec = \?exec (if (e) e1 else e2) stk loc (Suc (length (compE2 e) + pc)) xcp stk' loc' pc' xcp'\ note bsok = \bsok (if (e) e1 else e2) n\ from bisim1 have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e1)") case True let ?pre = "compE2 e @ [IfFalse (2 + int (length (compE2 e1)))]" let ?post = "Goto (1 + int (length (compE2 e2))) # compE2 e2" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e1 @ ?post) (compxE2 e 0 0 @ shift (length ?pre) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (compxE2 e2 (Suc 0) 0))) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 ac_simps exec_move_def) hence "exec_meth_d (compP2 P) (compE2 e1 @ ?post) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (compxE2 e2 (Suc 0) 0)) t h (stk, loc, pc, xcp) ta h' (stk', loc', pc' - length ?pre, xcp')" by(rule exec_meth_drop_xt) auto hence "exec_move_d P t e1 h (stk, loc, pc, xcp) ta h' (stk', loc', pc' - length ?pre, xcp')" using True unfolding exec_move_def by(rule exec_meth_take_xt) from IH[OF this len _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bsok obtain e'' xs'' where bisim': "P,e1,h' \ (e'', xs'') \ (stk', loc', pc' - length ?pre, xcp')" and red: "?red e' xs e'' xs'' e1 stk pc (pc' - length ?pre) xcp xcp'" by auto from bisim' have "P,if (e) e1 else e2,h' \ (e'', xs'') \ (stk', loc', Suc (length (compE2 e) + (pc' - length ?pre)), xcp')" by(rule bisim1_bisims1.bisim1CondThen) moreover from True have \: "\move2 (compP2 P) h stk (if (e) e1 else e2) (Suc (length (compE2 e) + pc)) xcp = \move2 (compP2 P) h stk e1 pc xcp" by(simp add: \move2_iff) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc) auto ultimately show ?thesis using red by(fastforce split: if_split_asm simp add: no_call2_def) next case False with pc have [simp]: "pc = length (compE2 e1)" by simp with bisim1 obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim1 pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) moreover have \: "\move2 (compP2 P) h [v] (if (e) e1 else e2) (Suc (length (compE2 e) + length (compE2 e1))) None" by(simp add: \move2_iff) moreover have "P,if (e) e1 else e2,h \ (Val v, loc) \ ([v], loc, length (compE2 (if (e) e1 else e2)), None)" by(rule bisim1Val2) simp moreover have "nat (2 + (int (length (compE2 e)) + (int (length (compE2 e1)) + int (length (compE2 e2))))) = length (compE2 (if (e) e1 else e2))" by simp ultimately show ?thesis using exec xcp stk by(fastforce elim!: exec_meth.cases simp add: exec_move_def) qed next case (bisim1CondElse e2 n e' xs stk loc pc xcp e e1) note IH = bisim1CondElse.IH(2) note bisim2 = \P,e2,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bisim = \\xs. P,e,h \ (e, xs) \ ([], xs, 0, None)\ note bisim1 = \\xs. P,e1,h \ (e1, xs) \ ([], xs, 0, None)\ note len = \n + max_vars e' \ length xs\ note exec = \?exec (if (e) e1 else e2) stk loc (Suc (Suc (length (compE2 e) + length (compE2 e1) + pc))) xcp stk' loc' pc' xcp'\ note bsok = \bsok (if (e) e1 else e2) n\ let ?pre = "compE2 e @ IfFalse (2 + int (length (compE2 e1))) # compE2 e1 @ [Goto (1 + int (length (compE2 e2)))]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e2) ((compxE2 e 0 0 @ compxE2 e1 (Suc (length (compE2 e))) 0) @ shift (length ?pre) (compxE2 e2 0 0)) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 ac_simps exec_move_def) hence "?exec e2 stk loc pc xcp stk' loc' (pc' - length ?pre) xcp'" unfolding exec_move_def by(rule exec_meth_drop_xt) auto from IH[OF this len _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bsok obtain e'' xs'' where bisim': "P,e2,h' \ (e'', xs'') \ (stk', loc', pc' - length ?pre, xcp')" and red: "?red e' xs e'' xs'' e2 stk pc (pc' - length ?pre) xcp xcp'" by auto from bisim' have "P,if (e) e1 else e2,h' \ (e'', xs'') \ (stk', loc', Suc (Suc (length (compE2 e) + length (compE2 e1) + (pc' - length ?pre))), xcp')" by(rule bisim1_bisims1.bisim1CondElse) moreover have \: "\move2 (compP2 P) h stk (if (e) e1 else e2) (Suc (Suc (length (compE2 e) + length (compE2 e1) + pc))) xcp = \move2 (compP2 P) h stk e2 pc xcp" by(simp add: \move2_iff) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc) auto moreover hence "Suc (Suc (pc' - Suc (Suc 0))) = pc'" by simp ultimately show ?case using red by(fastforce simp add: eval_nat_numeral no_call2_def split: if_split_asm) next case (bisim1CondThrow e n a xs stk loc pc e1 e2) note exec = \?exec (if (e) e1 else e2) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: match_ex_table_not_pcs_None exec_move_def) thus ?case .. next case (bisim1While1 c n e xs) note IH = bisim1While1.IH(2) note bisim = \\xs. P,c,h \ (c, xs) \ ([], xs, 0, None)\ note bisim1 = \\xs. P,e,h \ (e, xs) \ ([], xs, 0, None)\ from \n + max_vars (while (c) e) \ length xs\ have len: "n + max_vars c \ length xs" by simp note exec = \?exec (while (c) e) [] xs 0 None stk' loc' pc' xcp'\ note bsok = \bsok (while (c) e) n\ let ?post = "IfFalse (int (length (compE2 e)) + 3) # compE2 e @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]" from exec have "?exec c [] xs 0 None stk' loc' pc' xcp'" by(simp add: exec_move_While1) from IH[OF this len] bsok obtain e'' xs'' where bisim': "P,c,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red: "?red c xs e'' xs'' c [] 0 pc' None xcp'" by fastforce from bisim' have "P,while (c) e,h' \ (if (e'') (e;;while(c) e) else unit, xs'') \ (stk', loc', pc', xcp')" by(rule bisim1While3) moreover have "True,P,t \1 \while (c) e, (h, xs)\ -\\ \if (c) (e;;while (c) e) else unit, (h, xs)\" by(rule Red1While) hence "\red1r P t h (while (c) e, xs) (if (c) (e;;while (c) e) else unit, xs)" by(auto intro: r_into_rtranclp \move1WhileRed) moreover have "\move2 (compP2 P) h [] (while (c) e) 0 None = \move2 (compP2 P) h [] c 0 None" by(simp add: \move2_iff) ultimately show ?case using red by(fastforce elim!: rtranclp_trans rtranclp_tranclp_tranclp intro: Cond1Red elim!: Cond_\red1r_xt Cond_\red1t_xt simp add: no_call2_def) next case (bisim1While3 c n e' xs stk loc pc xcp e) note IH = bisim1While3.IH(2) note bisim = \P,c,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bisim1 = \\xs. P,e,h \ (e, xs) \ ([], xs, 0, None)\ from \n + max_vars (if (e') (e;;while (c) e) else unit) \ length xs\ have len: "n + max_vars e' \ length xs" by simp note bsok = \bsok (while (c) e) n\ note exec = \?exec (while (c) e) stk loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 c)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 c)") case True let ?post = "IfFalse (int (length (compE2 e)) + 3) # compE2 e @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]" from exec have "exec_meth_d (compP2 P) (compE2 c @ ?post) (compxE2 c 0 0 @ shift (length (compE2 c)) (compxE2 e (Suc 0) 0)) t h (stk, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 exec_move_def) hence "?exec c stk loc pc xcp stk' loc' pc' xcp'" using True unfolding exec_move_def by(rule exec_meth_take_xt) from IH[OF this len _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bsok obtain e'' xs'' where bisim': "P,c,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red: "?red e' xs e'' xs'' c stk pc pc' xcp xcp'" by auto from bisim' have "P,while (c) e,h' \ (if (e'') (e;;while(c) e) else unit, xs'') \ (stk', loc', pc', xcp')" by(rule bisim1_bisims1.bisim1While3) moreover have "\move2 (compP2 P) h stk (while (c) e) pc xcp = \move2 (compP2 P) h stk c pc xcp" using True by(simp add: \move2_iff) ultimately show ?thesis using red by(fastforce elim!: rtranclp_trans intro: Cond1Red elim!: Cond_\red1r_xt Cond_\red1t_xt simp add: no_call2_def) next case False with pc have [simp]: "pc = length (compE2 c)" by simp with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (if (e') (e;; while (c) e) else unit, xs) (if (Val v) (e;; while (c) e) else unit, loc)" by(rule Cond_\red1r_xt) moreover have \: "\move2 (compP2 P) h [v] (while (c) e) (length (compE2 c)) None" by(simp add: \move2_iff) moreover have "\move1 P h (if (Val v) (e;;while (c) e) else unit)" by(rule \move1CondRed) moreover from bisim1[of loc] have "P,while (c) e,h \ (e;;while(c) e, loc) \ ([], loc, Suc (length (compE2 c) + 0), None)" by(rule bisim1While4) moreover have "P,while (c) e,h \ (unit, loc) \ ([], loc, Suc (Suc (Suc (length (compE2 c) + length (compE2 e)))), None)" by(rule bisim1While7) moreover from exec stk xcp have "typeof\<^bsub>h\<^esub> v = \Boolean\" by(auto simp add: exec_meth_instr exec_move_def) moreover have "nat (int (length (compE2 c)) + (3 + int (length (compE2 e)))) = Suc (Suc (Suc (length (compE2 c) + length (compE2 e))))" by simp ultimately show ?thesis using exec stk xcp by(fastforce elim!: exec_meth.cases rtranclp_trans intro: Red1CondT Red1CondF simp add: eval_nat_numeral exec_move_def) qed next case (bisim1While4 e n e' xs stk loc pc xcp c) note IH = bisim1While4.IH(2) note bisim = \\xs. P,c,h \ (c, xs) \ ([], xs, 0, None)\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ from \n + max_vars (e';; while (c) e) \ length xs\ have len: "n + max_vars e' \ length xs" by simp note exec = \?exec (while (c) e) stk loc (Suc (length (compE2 c) + pc)) xcp stk' loc' pc' xcp'\ note bsok = \bsok (while (c) e) n\ from bisim1 have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True let ?pre = "compE2 c @ [IfFalse (int (length (compE2 e)) + 3)]" let ?post = "[Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]" from exec have "exec_meth_d (compP2 P) ((?pre @ compE2 e) @ ?post) (compxE2 c 0 0 @ shift (length ?pre) (compxE2 e 0 0)) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 exec_move_def) hence exec': "exec_meth_d (compP2 P) (?pre @ compE2 e) (compxE2 c 0 0 @ shift (length ?pre) (compxE2 e 0 0)) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "?exec e stk loc pc xcp stk' loc' (pc' - length ?pre) xcp'" unfolding exec_move_def by(rule exec_meth_drop_xt) auto from IH[OF this len _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bsok obtain e'' xs'' where bisim': "P,e,h' \ (e'', xs'') \ (stk', loc', pc' - length ?pre, xcp')" and red: "?red e' xs e'' xs'' e stk pc (pc' - length ?pre) xcp xcp'" by auto from red have "?red (e';;while (c) e) xs (e'';;while (c) e) xs'' e stk pc (pc' - length ?pre) xcp xcp'" by(fastforce intro: Seq1Red elim!: Seq_\red1r_xt Seq_\red1t_xt split: if_split_asm) moreover from bisim' have "P,while (c) e,h' \ (e'';;while(c) e, xs'') \ (stk', loc', Suc (length (compE2 c) + (pc' - length ?pre)), xcp')" by(rule bisim1_bisims1.bisim1While4) moreover have "\move2 (compP2 P) h stk (while (c) e) (Suc (length (compE2 c) + pc)) xcp = \move2 (compP2 P) h stk e pc xcp" using True by(simp add: \move2_iff) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc) auto moreover have "no_call2 e pc \ no_call2 (while (c) e) (Suc (length (compE2 c) + pc))" by(simp add: no_call2_def) ultimately show ?thesis apply(auto split: if_split_asm) apply(fastforce+)[6] apply(rule exI conjI|assumption|rule rtranclp.rtrancl_refl|simp)+ done next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim1 obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim1 pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (e';; while (c) e, xs) (Val v;; while (c) e, loc)" by(rule Seq_\red1r_xt) moreover have \: "\move2 (compP2 P) h [v] (while (c) e) (Suc (length (compE2 c) + length (compE2 e))) None" by(simp add: \move2_iff) moreover have "\move1 P h (Val v;;while (c) e)" by(rule \move1SeqRed) moreover have "P,while (c) e,h \ (while(c) e, loc) \ ([], loc, Suc (Suc (length (compE2 c) + length (compE2 e))), None)" by(rule bisim1While6) ultimately show ?thesis using exec stk xcp by(fastforce elim!: exec_meth.cases rtranclp_trans intro: Red1Seq simp add: eval_nat_numeral exec_move_def) qed next case (bisim1While6 c n e xs) note exec = \?exec (while (c) e) [] xs (Suc (Suc (length (compE2 c) + length (compE2 e)))) None stk' loc' pc' xcp'\ moreover have "\move2 (compP2 P) h [] (while (c) e) (Suc (Suc (length (compE2 c) + length (compE2 e)))) None" by(simp add: \move2_iff) moreover have "P,while (c) e,h' \ (if (c) (e;; while (c) e) else unit, xs) \ ([], xs, 0, None)" by(rule bisim1While3[OF bisim1_refl]) moreover have "\red1t P t h (while (c) e, xs) (if (c) (e;; while (c) e) else unit, xs)" by(rule tranclp.r_into_trancl)(auto intro: Red1While) ultimately show ?case by(fastforce elim!: exec_meth.cases simp add: exec_move_def) next case (bisim1While7 c n e xs) note \?exec (while (c) e) [] xs (Suc (Suc (Suc (length (compE2 c) + length (compE2 e))))) None stk' loc' pc' xcp'\ moreover have "\move2 (compP2 P) h [] (while (c) e) (Suc (Suc (Suc (length (compE2 c) + length (compE2 e))))) None" by(simp add: \move2_iff) moreover have "P,while (c) e,h' \ (unit, xs) \ ([Unit], xs, length (compE2 (while (c) e)), None)" by(rule bisim1Val2) simp ultimately show ?case by(fastforce elim!: exec_meth.cases simp add: exec_move_def) next case (bisim1WhileThrow1 c n a xs stk loc pc e) note exec = \?exec (while (c) e) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,c,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 c)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 c 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: match_ex_table_not_pcs_None exec_move_def) thus ?case .. next case (bisim1WhileThrow2 e n a xs stk loc pc c) note exec = \?exec (while (c) e) stk loc (Suc (length (compE2 c) + pc)) \a\ stk' loc' pc' xcp'\ note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False apply(auto elim!: exec_meth.cases simp add: match_ex_table_not_pcs_None exec_move_def) apply(auto dest!: match_ex_table_shift_pcD simp only: compxE2_size_convs) apply simp done thus ?case .. next case (bisim1Throw1 e n e' xs stk loc pc xcp) note IH = bisim1Throw1.IH(2) note exec = \?exec (throw e) stk loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note len = \n + max_vars (throw e') \ length xs\ note bsok = \bsok (throw e) n\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have exec': "?exec e stk loc pc xcp stk' loc' pc' xcp'" by(simp add: exec_move_Throw) from True have "\move2 (compP2 P) h stk (throw e) pc xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff) with IH[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] len bsok show ?thesis by(fastforce intro: bisim1_bisims1.bisim1Throw1 Throw1Red elim!: Throw_\red1r_xt Throw_\red1t_xt simp add: no_call2_def) next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (throw e', xs) (throw (Val v), loc)" by(rule Throw_\red1r_xt) moreover have \: "\move2 (compP2 P) h [v] (throw e) pc None" by(simp add: \move2_iff) moreover have "\a. P,throw e,h \ (Throw a, loc) \ ([Addr a], loc, length (compE2 e), \a\)" by(rule bisim1Throw2) moreover have "P,throw e,h \ (THROW NullPointer, loc) \ ([Null], loc, length (compE2 e), \addr_of_sys_xcpt NullPointer\)" by(rule bisim1ThrowNull) moreover from exec stk xcp \P,h \ stk [:\] ST\ obtain T' where T': "typeof\<^bsub>h\<^esub> v = \T'\" "P \ T' \ Class Throwable" by(auto simp add: exec_move_def exec_meth_instr list_all2_Cons1 conf_def compP2_def) moreover with T' have "v \ Null \ \C. T' = Class C" by(cases T')(auto dest: Array_widen) moreover have "\red1r P t h (throw null, loc) (THROW NullPointer, loc)" by(auto intro: r_into_rtranclp Red1ThrowNull \move1ThrowNull) ultimately show ?thesis using exec stk xcp T' unfolding exec_move_def by(cases v)(fastforce elim!: exec_meth.cases intro: rtranclp_trans)+ qed next case (bisim1Throw2 e n a xs) note exec = \?exec (throw e) [Addr a] xs (length (compE2 e)) \a\ stk' loc' pc' xcp'\ hence False by(auto elim!: exec_meth.cases dest: match_ex_table_pc_length_compE2 simp add: exec_move_def) thus ?case .. next case (bisim1ThrowNull e n xs) note exec = \?exec (throw e) [Null] xs (length (compE2 e)) \addr_of_sys_xcpt NullPointer\ stk' loc' pc' xcp'\ hence False by(auto elim!: exec_meth.cases dest: match_ex_table_pc_length_compE2 simp add: exec_move_def) thus ?case .. next case (bisim1ThrowThrow e n a xs stk loc pc) note exec = \?exec (throw e) stk loc pc \a\ stk' loc' pc' xcp'\ note bisim1 = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim1 have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim1 have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with exec pc have False by(auto elim!: exec_meth.cases simp add: exec_move_def) thus ?case .. next case (bisim1Try e n e' xs stk loc pc xcp e2 C' V) note IH = bisim1Try.IH(2) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bisim1 = \\xs. P,e2,h \ (e2, xs) \ ([], xs, 0, None)\ note exec = \?exec (try e catch(C' V) e2) stk loc pc xcp stk' loc' pc' xcp'\ note bsok = \bsok (try e catch(C' V) e2) n\ with \n + max_vars (try e' catch(C' V) e2) \ length xs\ have len: "n + max_vars e' \ length xs" and V: "V < length xs" by simp_all from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True note pc = True show ?thesis proof(cases "xcp = None \ (\a'. xcp = \a'\ \ match_ex_table (compP2 P) (cname_of h a') pc (compxE2 e 0 0) \ None)") case False then obtain a' where Some: "xcp = \a'\" and True: "match_ex_table (compP2 P) (cname_of h a') pc (compxE2 e 0 0) = None" by(auto simp del: not_None_eq) from Some bisim \conf_xcp' (compP2 P) h xcp\ have "\C. typeof\<^bsub>h\<^esub> (Addr a') = \Class C\ \ P \ C \\<^sup>* Throwable" by(auto simp add: compP2_def) then obtain C'' where ha': "typeof_addr h a' = \Class_type C''\" and subclsThrow: "P \ C'' \\<^sup>* Throwable" by(auto) with exec True Some pc have subcls: "P \ C'' \\<^sup>* C'" apply(auto elim!: exec_meth.cases simp add: match_ex_table_append compP2_def matches_ex_entry_def exec_move_def cname_of_def split: if_split_asm) apply(simp only: compxE2_size_convs, simp) done moreover from ha' subclsThrow bsok have red: "\red1r P t h (e', xs) (Throw a', loc)" and bisim': "P,e,h \ (Throw a', loc) \ (stk, loc, pc, \a'\)" using bisim True len unfolding Some compP2_def by(auto dest!: bisim1_xcp_\Red simp add: bsok_def) from red have lenloc: "length loc = length xs" by(rule \red1r_preserves_len) from red have "\red1r P t h (try e' catch(C' V) e2, xs) (try (Throw a') catch(C' V) e2, loc)" by(rule Try_\red1r_xt) hence "\red1r P t h (try e' catch(C' V) e2, xs) ({V:Class C'=None; e2}, loc[V := Addr a'])" using ha' subcls V unfolding lenloc[symmetric] by(auto intro: rtranclp.rtrancl_into_rtrancl Red1TryCatch \move1TryThrow) moreover from pc have "\move2 (compP2 P) h stk (try e catch(C' V) e2) pc \a'\" by(simp add: \move2_iff) moreover from bisim' ha' subcls have "P,try e catch(C' V) e2,h \ ({V:Class C'=None; e2}, loc[V := Addr a']) \ ([Addr a'], loc, Suc (length (compE2 e)), None)" by(rule bisim1TryCatch1) ultimately show ?thesis using exec True pc Some ha' subclsThrow apply(auto elim!: exec_meth.cases simp add: ac_simps eval_nat_numeral match_ex_table_append matches_ex_entry_def compP2_def exec_move_def cname_of_def) apply fastforce apply(simp_all only: compxE2_size_convs, auto dest: match_ex_table_shift_pcD) done next case True let ?post = "Goto (int (length (compE2 e2)) + 2) # Store V # compE2 e2" from exec True have "exec_meth_d (compP2 P) (compE2 e @ ?post) (compxE2 e 0 0 @ shift (length (compE2 e)) (compxE2 e2 (Suc (Suc 0)) 0)) t h (stk, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(auto elim!: exec_meth.cases intro: exec_meth.intros simp add: match_ex_table_append shift_compxE2 exec_move_def) hence "?exec e stk loc pc xcp stk' loc' pc' xcp'" using pc unfolding exec_move_def by(rule exec_meth_take_xt) from IH[OF this len _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bsok obtain e'' xs'' where bisim': "P,e,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red': "?red e' xs e'' xs'' e stk pc pc' xcp xcp'" by auto from bisim' have "P,try e catch(C' V) e2,h' \ (try e'' catch(C' V) e2, xs'') \ (stk', loc', pc', xcp')" by(rule bisim1_bisims1.bisim1Try) moreover from pc have "\move2 (compP2 P) h stk (try e catch(C' V) e2) pc xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff) ultimately show ?thesis using red' by(fastforce intro: Try1Red elim!: Try_\red1r_xt Try_\red1t_xt simp add: no_call2_def) qed next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where stk: "stk = [v]" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with bisim pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\red1r P t h (try e' catch(C' V) e2, xs) (try (Val v) catch(C' V) e2, loc)" by(rule Try_\red1r_xt) hence "\red1r P t h (try e' catch(C' V) e2, xs) (Val v, loc)" by(auto intro: rtranclp.rtrancl_into_rtrancl Red1Try \move1TryRed) moreover have \: "\move2 (compP2 P) h [v] (try e catch(C' V) e2) pc None" by(simp add: \move2_iff) moreover have "P,try e catch(C' V) e2,h \ (Val v, loc) \ ([v], loc, length (compE2 (try e catch(C' V) e2)), None)" by(rule bisim1Val2) simp moreover have "nat (int (length (compE2 e)) + (int (length (compE2 e2)) + 2)) = length (compE2 (try e catch(C' V) e2))" by simp ultimately show ?thesis using exec stk xcp by(fastforce elim!: exec_meth.cases simp add: exec_move_def) qed next case (bisim1TryCatch1 e n a xs stk loc pc C'' C' e2 V) note exec = \?exec (try e catch(C' V) e2) [Addr a] loc (Suc (length (compE2 e))) None stk' loc' pc' xcp'\ note bisim2 = \P,e2,h \ (e2, loc[V := Addr a]) \ ([], loc[V := Addr a], 0, None)\ note bisim1 = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ hence [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from bisim2 have "P, try e catch(C' V) e2, h \ ({V:Class C'=None; e2}, loc[V := Addr a]) \ ([], loc[V := Addr a], Suc (Suc (length (compE2 e) + 0)), None)" by(rule bisim1TryCatch2) moreover have "\move2 (compP2 P) h [Addr a] (try e catch(C' V) e2) (Suc (length (compE2 e))) None" by(simp add: \move2_iff) ultimately show ?case using exec by(fastforce elim!: exec_meth.cases simp add: exec_move_def) next case (bisim1TryCatch2 e2 n e' xs stk loc pc xcp e C' V) note IH = bisim1TryCatch2.IH(2) note bisim2 = \P,e2,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bisim = \\xs. P,e,h \ (e, xs) \ ([], xs, 0, None)\ note exec = \?exec (try e catch(C' V) e2) stk loc (Suc (Suc (length (compE2 e) + pc))) xcp stk' loc' pc' xcp'\ note bsok = \bsok (try e catch(C' V) e2) n\ with \n + max_vars {V:Class C'=None; e'} \ length xs\ have len: "Suc n + max_vars e' \ length xs" and V: "V < length xs" by simp_all let ?pre = "compE2 e @ [Goto (int (length (compE2 e2)) + 2), Store V]" from exec have "exec_meth_d (compP2 P) (?pre @ compE2 e2) (compxE2 e 0 0 @ shift (length ?pre) (compxE2 e2 0 0) @ [(0, length (compE2 e), \C'\, Suc (length (compE2 e)), 0)]) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 exec_move_def) hence exec': "exec_meth_d (compP2 P) (?pre @ compE2 e2) (compxE2 e 0 0 @ shift (length ?pre) (compxE2 e2 0 0)) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(auto elim!: exec_meth.cases intro: exec_meth.intros simp add: match_ex_table_append matches_ex_entry_def) hence "?exec e2 stk loc pc xcp stk' loc' (pc' - length ?pre) xcp'" unfolding exec_move_def by(rule exec_meth_drop_xt) auto from IH[OF this len _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bsok obtain e'' xs'' where bisim': "P,e2,h' \ (e'', xs'') \ (stk', loc', pc' - length ?pre, xcp')" and red: "?red e' xs e'' xs'' e2 stk pc (pc' - length ?pre) xcp xcp'" by auto from red have "length xs'' = length xs" by(auto dest!: \red1r_preserves_len \red1t_preserves_len red1_preserves_len split: if_split_asm) with red V have "?red {V:Class C'=None; e'} xs {V:Class C'=None; e''} xs'' e2 stk pc (pc' - length ?pre) xcp xcp'" by(fastforce elim!: Block_None_\red1r_xt Block_None_\red1t_xt intro: Block1Red split: if_split_asm) moreover from bisim' have "P,try e catch(C' V) e2,h' \ ({V:Class C'=None;e''}, xs'') \ (stk', loc', Suc (Suc (length (compE2 e) + (pc' - length ?pre))), xcp')" by(rule bisim1_bisims1.bisim1TryCatch2) moreover have "\move2 (compP2 P) h stk (try e catch(C' V) e2) (Suc (Suc (length (compE2 e) + pc))) xcp = \move2 (compP2 P) h stk e2 pc xcp" by(simp add: \move2_iff) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc) auto moreover hence "Suc (Suc (pc' - Suc (Suc 0))) = pc'" by simp moreover have "no_call2 e2 pc \ no_call2 (try e catch(C' V) e2) (Suc (Suc (length (compE2 e) + pc)))" by(simp add: no_call2_def) ultimately show ?case using red V by(fastforce simp add: eval_nat_numeral split: if_split_asm) next case (bisim1TryFail e n a xs stk loc pc C'' C' e2 V) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from bisim have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with \?exec (try e catch(C' V) e2) stk loc pc \a\ stk' loc' pc' xcp'\ pc \typeof_addr h a = \Class_type C''\\ \\ P \ C'' \\<^sup>* C'\ have False by(auto elim!: exec_meth.cases simp add: matches_ex_entry_def compP2_def match_ex_table_append_not_pcs exec_move_def cname_of_def) thus ?case .. next case (bisim1TryCatchThrow e2 n a xs stk loc pc e C' V) note bisim = \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e2)" by(auto dest: bisim1_ThrowD) from bisim have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e2 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) with \?exec (try e catch(C' V) e2) stk loc (Suc (Suc (length (compE2 e) + pc))) \a\ stk' loc' pc' xcp'\ pc have False apply(auto elim!: exec_meth.cases simp add: compxE2_size_convs match_ex_table_append_not_pcs exec_move_def) apply(auto dest!: match_ex_table_shift_pcD simp add: match_ex_table_append matches_ex_entry_def compP2_def) done thus ?case .. next case bisims1Nil hence False by(auto elim!: exec_meth.cases simp add: exec_moves_def) thus ?case .. next case (bisims1List1 e n e' xs stk loc pc xcp es) note IH1 = bisims1List1.IH(2) note IH2 = bisims1List1.IH(4) note exec = \?execs (e # es) stk loc pc xcp stk' loc' pc' xcp'\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,es,h \ (es, loc) [\] ([], loc, 0, None)\ note len = \n + max_varss (e' # es) \ length xs\ note bsok = \bsoks (e # es) n\ from bisim1 have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) from bisim1 have lenxs: "length xs = length loc" by(rule bisim1_length_xs) show ?case proof(cases "pc < length (compE2 e)") case True with exec have exec': "?exec e stk loc pc xcp stk' loc' pc' xcp'" by(auto simp add: compxEs2_size_convs exec_moves_def exec_move_def intro: exec_meth_take_xt) from True have "\moves2 (compP2 P) h stk (e # es) pc xcp = \move2 (compP2 P) h stk e pc xcp" by(simp add: \move2_iff \moves2_iff) moreover from True have "no_calls2 (e # es) pc = no_call2 e pc" by(simp add: no_call2_def no_calls2_def) ultimately show ?thesis using IH1[OF exec' _ _ \P,h \ stk [:\] ST\ \conf_xcp' (compP2 P) h xcp\] bsok len by(fastforce intro: bisim1_bisims1.bisims1List1 elim!: \red1r_inj_\reds1r \red1t_inj_\reds1t List1Red1) next case False with pc have pc [simp]: "pc = length (compE2 e)" by simp with bisim1 obtain v where stk: "stk = [v]" and xcp: "xcp = None" and v: "is_val e' \ e' = Val v \ xs = loc" and call: "call1 e' = None" by(auto dest: bisim1_pc_length_compE2D) with bisim1 pc len bsok have red: "\red1r P t h (e', xs) (Val v, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) hence "\reds1r P t h (e' # es, xs) (Val v # es, loc)" by(rule \red1r_inj_\reds1r) moreover from exec stk xcp have exec': "exec_meth_d (compP2 P) (compE2 e @ compEs2 es) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v]) (compxEs2 es 0 0))) t h ([] @ [v], loc, length (compE2 e) + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxEs2 stack_xlift_compxEs2 exec_moves_def) hence "exec_meth_d (compP2 P) (compEs2 es) (stack_xlift (length [v]) (compxEs2 es 0 0)) t h ([] @ [v], loc, 0, None) ta h' (stk', loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_drop_xt) auto with bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_moves_d P t es h ([], loc, 0, None) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" by(unfold exec_moves_def)(drule (1) exec_meth_stk_splits, auto) from IH2[OF exec''] len lenxs bsok obtain es'' xs'' where bisim': "P,es,h' \ (es'', xs'') [\] (stk'', loc', pc' - length (compE2 e), xcp')" and red': "?reds es loc es'' xs'' es [] 0 (pc' - length (compE2 e)) None xcp'" by fastforce from bisim' have "P,e # es,h' \ (Val v # es'', xs'') [\] (stk'' @ [v], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule bisims1List2) moreover from exec'' have "\moves2 (compP2 P) h [v] (e # es) (length (compE2 e)) None = \moves2 (compP2 P) h [] es 0 None" using \instr_stk_drop_exec_moves[where stk="[]" and vs="[v]"] by(simp add: \moves2_iff) moreover have \: "\es'. \moves1 P h (Val v # es') \ \moves1 P h es'" by simp from exec' have "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_calls2 es 0 \ no_calls2 (e # es) (length (compE2 e))" by(simp add: no_calls2_def) ultimately show ?thesis using red' xcp stk stk' call v apply(auto simp add: split_paired_Ex) apply(blast 25 intro: rtranclp_trans rtranclp_tranclp_tranclp \reds1r_cons_\reds1r List1Red2 \reds1t_cons_\reds1t dest: \)+ done qed next case (bisims1List2 es n es' xs stk loc pc xcp e v) note IH = bisims1List2.IH(2) note exec = \?execs (e # es) (stk @ [v]) loc (length (compE2 e) + pc) xcp stk' loc' pc' xcp'\ note bisim1 = \P,e,h \ (e, xs) \ ([], xs, 0, None)\ note bisim2 = \P,es,h \ (es', xs) [\] (stk, loc, pc, xcp)\ note len = \n + max_varss (Val v # es') \ length xs\ note bsok = \bsoks (e # es) n\ from \P,h \ stk @ [v] [:\] ST\ obtain ST' where ST': "P,h \ stk [:\] ST'" by(auto simp add: list_all2_append1) from exec have exec': "exec_meth_d (compP2 P) (compE2 e @ compEs2 es) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v]) (compxEs2 es 0 0))) t h (stk @ [v], loc, length (compE2 e) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxEs2 stack_xlift_compxEs2 exec_moves_def) hence "exec_meth_d (compP2 P) (compEs2 es) (stack_xlift (length [v]) (compxEs2 es 0 0)) t h (stk @ [v], loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_drop_xt) auto with bisim2 obtain stk'' where stk': "stk' = stk'' @ [v]" and exec'': "exec_moves_d P t es h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" by(unfold exec_moves_def)(drule (1) exec_meth_stk_splits, auto) from IH[OF exec'' _ _ ST' \conf_xcp' (compP2 P) h xcp\] len bsok obtain es'' xs'' where bisim': "P,es,h' \ (es'', xs'') [\] (stk'', loc', pc' - length (compE2 e), xcp')" and red': "?reds es' xs es'' xs'' es stk pc (pc' - length (compE2 e)) xcp xcp'" by auto from bisim' have "P,e # es,h' \ (Val v # es'', xs'') [\] (stk'' @ [v], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule bisim1_bisims1.bisims1List2) moreover from exec'' have "\moves2 (compP2 P) h (stk @ [v]) (e # es) (length (compE2 e) + pc) xcp = \moves2 (compP2 P) h stk es pc xcp" by(auto simp add: \moves2_iff \instr_stk_drop_exec_moves) moreover have \: "\es'. \moves1 P h (Val v # es') \ \moves1 P h es'" by simp from exec' have "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc) auto moreover have "no_calls2 es pc \ no_calls2 (e # es) (length (compE2 e) + pc)" by(simp add: no_calls2_def) ultimately show ?case using red' stk' by(auto split: if_split_asm simp add: split_paired_Ex)(blast intro: rtranclp_trans rtranclp_tranclp_tranclp \reds1r_cons_\reds1r List1Red2 \reds1t_cons_\reds1t dest: \)+ qed end inductive sim21_size_aux :: "nat \ (pc \ 'addr option) \ (pc \ 'addr option) \ bool" for len :: nat where "\ pc1 \ len; pc2 \ len; xcp1 \ None \ xcp2 = None \ pc1 = pc2 \ xcp1 = None \ pc1 > pc2 \ \ sim21_size_aux len (pc1, xcp1) (pc2, xcp2)" definition sim21_size :: "'addr jvm_prog \ 'addr jvm_thread_state \ 'addr jvm_thread_state \ bool" where "sim21_size P xcpfrs xcpfrs' \ (xcpfrs, xcpfrs') \ inv_image (less_than <*lex*> same_fst (\n. True) (\n. {(pcxcp, pcxcp'). sim21_size_aux n pcxcp pcxcp'})) (\(xcp, frs). (length frs, case frs of [] \ undefined | (stk, loc, C, M, pc) # frs \ (length (fst (snd (snd (the (snd (snd (snd (method P C M)))))))), pc, xcp)))" lemma wfP_sim21_size_aux: "wfP (sim21_size_aux n)" proof - let ?f = "\(pc, xcp). case xcp of None \ Suc (2 * (n - pc)) | Some _ \ 2 * (n - pc)" have "wf {(m, m'). (m :: nat) < m'}" by(rule wf_less) hence "wf (inv_image {(m, m'). m < m'} ?f)" by(rule wf_inv_image) moreover have "{(pcxcp1, pcxcp2). sim21_size_aux n pcxcp1 pcxcp2} \ inv_image {(m, m'). m < m'} ?f" by(auto elim!: sim21_size_aux.cases) ultimately show ?thesis unfolding wfP_def by(rule wf_subset) qed lemma Collect_split_mem: "{(x, y). (x, y) \ Q} = Q" by simp lemma wfP_sim21_size: "wfP (sim21_size P)" unfolding wfP_def Collect_split_mem sim21_size_def [abs_def] apply(rule wf_inv_image) apply(rule wf_lex_prod) apply(rule wf_less_than) apply(rule wf_same_fst) apply(rule wfP_sim21_size_aux[unfolded wfP_def]) done declare split_beta[simp] context J1_JVM_heap_base begin lemma bisim1_Invoke_\Red: "\ P,E,h \ (e, xs) \ (rev vs @ Addr a # stk', loc, pc, None); pc < length (compE2 E); compE2 E ! pc = Invoke M (length vs); n + max_vars e \ length xs; bsok E n \ \ \e' xs'. \red1r P t h (e, xs) (e', xs') \ P,E,h \ (e', xs') \ (rev vs @ Addr a # stk', loc, pc, None) \ call1 e' = \(a, M, vs)\" (is "\ _; _; _; _; _ \ \ ?concl e xs E n pc stk' loc") and bisims1_Invoke_\Reds: "\ P,Es,h \ (es, xs) [\] (rev vs @ Addr a # stk', loc, pc, None); pc < length (compEs2 Es); compEs2 Es ! pc = Invoke M (length vs); n + max_varss es \ length xs; bsoks Es n \ \ \es' xs'. \reds1r P t h (es, xs) (es', xs') \ P,Es,h \ (es', xs') [\] (rev vs @ Addr a # stk', loc, pc, None) \ calls1 es' = \(a, M, vs)\" (is "\ _; _; _; _; _ \ \ ?concls es xs Es n pc stk' loc") proof(induct E n e xs stk\"rev vs @ Addr a # stk'" loc pc xcp\"None::'addr option" and Es n es xs stk\"rev vs @ Addr a # stk'" loc pc xcp\"None::'addr option" arbitrary: stk' and stk' rule: bisim1_bisims1_inducts_split) case bisim1Val2 thus ?case by simp next case bisim1New thus ?case by simp next case bisim1NewArray thus ?case by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1NewArray dest: bisim1_pc_length_compE2 elim!: NewArray_\red1r_xt) next case bisim1Cast thus ?case by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1Cast dest: bisim1_pc_length_compE2 elim!: Cast_\red1r_xt) next case bisim1InstanceOf thus ?case by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1InstanceOf dest: bisim1_pc_length_compE2 elim!: InstanceOf_\red1r_xt) next case bisim1Val thus ?case by simp next case bisim1Var thus ?case by simp next case bisim1BinOp1 thus ?case apply(auto split: if_split_asm intro: bisim1_bisims1.bisim1BinOp1 dest: bisim1_pc_length_compE2 elim: BinOp_\red1r_xt1) apply(fastforce elim!: BinOp_\red1r_xt1 intro: bisim1_bisims1.bisim1BinOp1) done next case (bisim1BinOp2 e2 n e' xs stk loc pc e1 bop v1) note IH = \\stk'. \stk = rev vs @ Addr a # stk'; pc < length (compE2 e2); compE2 e2 ! pc = Invoke M (length vs); n + max_vars e' \ length xs; bsok e2 n \ \ ?concl e' xs e2 n pc stk' loc\ note inv = \compE2 (e1 \bop\ e2) ! (length (compE2 e1) + pc) = Invoke M (length vs)\ with \length (compE2 e1) + pc < length (compE2 (e1 \bop\ e2))\ have pc: "pc < length (compE2 e2)" by(auto split: bop.splits if_split_asm) moreover with inv have "compE2 e2 ! pc = Invoke M (length vs)" by simp moreover with \P,e2,h \ (e', xs) \ (stk, loc, pc, None)\ pc obtain vs'' v'' stk'' where "stk = vs'' @ v'' # stk''" and "length vs'' = length vs" by(auto dest!: bisim1_Invoke_stkD) with \stk @ [v1] = rev vs @ Addr a # stk'\ obtain stk''' where stk''': "stk = rev vs @ Addr a # stk'''" and stk: "stk' = stk''' @ [v1]" by(cases stk' rule: rev_cases) auto from \n + max_vars (Val v1 \bop\ e') \ length xs\ have "n + max_vars e' \ length xs" by simp moreover from \bsok (e1 \bop\ e2) n\ have "bsok e2 n" by simp ultimately have "?concl e' xs e2 n pc stk''' loc" using stk''' by-(rule IH) then obtain e'' xs' where IH': "\red1r P t h (e', xs) (e'', xs')" "call1 e'' = \(a, M, vs)\" and bisim: "P,e2,h \ (e'', xs') \ (rev vs @ Addr a # stk''', loc, pc, None)" by blast from bisim have "P,e1\bop\e2,h \ (Val v1 \bop\ e'', xs') \ ((rev vs @ Addr a # stk''') @ [v1], loc, length (compE2 e1) + pc, None)" by(rule bisim1_bisims1.bisim1BinOp2) with IH' stk show ?case by(fastforce elim!: BinOp_\red1r_xt2) next case bisim1LAss1 thus ?case by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1LAss1 dest: bisim1_pc_length_compE2 elim!: LAss_\red1r) next case bisim1LAss2 thus ?case by simp next case bisim1AAcc1 thus ?case apply(auto split: if_split_asm intro: bisim1_bisims1.bisim1AAcc1 dest: bisim1_pc_length_compE2 elim!: AAcc_\red1r_xt1) apply(fastforce elim!: AAcc_\red1r_xt1 intro: bisim1_bisims1.bisim1AAcc1) done next case (bisim1AAcc2 e2 n e' xs stk loc pc e1 v1) note IH = \\stk'. \stk = rev vs @ Addr a # stk'; pc < length (compE2 e2); compE2 e2 ! pc = Invoke M (length vs); n + max_vars e' \ length xs; bsok e2 n \ \ ?concl e' xs e2 n pc stk' loc\ note inv = \compE2 (e1\e2\) ! (length (compE2 e1) + pc) = Invoke M (length vs)\ with \length (compE2 e1) + pc < length (compE2 (e1\e2\))\ have pc: "pc < length (compE2 e2)" by(auto split: if_split_asm) moreover with inv have "compE2 e2 ! pc = Invoke M (length vs)" by simp moreover with \P,e2,h \ (e', xs) \ (stk, loc, pc, None)\ pc obtain vs'' v'' stk'' where "stk = vs'' @ v'' # stk''" and "length vs'' = length vs" by(auto dest!: bisim1_Invoke_stkD) with \stk @ [v1] = rev vs @ Addr a # stk'\ obtain stk''' where stk''': "stk = rev vs @ Addr a # stk'''" and stk: "stk' = stk''' @ [v1]" by(cases stk' rule: rev_cases) auto from \n + max_vars (Val v1\e'\) \ length xs\ have "n + max_vars e' \ length xs" by simp moreover from \bsok (e1\e2\) n\ have "bsok e2 n" by simp ultimately have "?concl e' xs e2 n pc stk''' loc" using stk''' by-(rule IH) then obtain e'' xs' where IH': "\red1r P t h (e', xs) (e'', xs')" "call1 e'' = \(a, M, vs)\" and bisim: "P,e2,h \ (e'', xs') \ (rev vs @ Addr a # stk''', loc, pc, None)" by blast from bisim have "P,e1\e2\,h \ (Val v1\e''\, xs') \ ((rev vs @ Addr a # stk''') @ [v1], loc, length (compE2 e1) + pc, None)" by(rule bisim1_bisims1.bisim1AAcc2) with IH' stk show ?case by(fastforce elim!: AAcc_\red1r_xt2) next case (bisim1AAss1 e n e' xs loc pc e2 e3) note IH = \\ pc < length (compE2 e); compE2 e ! pc = Invoke M (length vs); n + max_vars e' \ length xs; bsok e n \ \ ?concl e' xs e n pc stk' loc\ note bisim = \P,e,h \ (e', xs) \ (rev vs @ Addr a # stk', loc, pc, None)\ note len = \n + max_vars (e'\e2\ := e3) \ length xs\ hence len': "n + max_vars e' \ length xs" by simp note inv = \compE2 (e\e2\ := e3) ! pc = Invoke M (length vs)\ with \pc < length (compE2 (e\e2\ := e3))\ bisim have pc: "pc < length (compE2 e)" by(auto split: if_split_asm dest: bisim1_pc_length_compE2) moreover with inv have "compE2 e ! pc = Invoke M (length vs)" by simp moreover from \bsok (e\e2\ := e3) n\ have "bsok e n" by simp ultimately have "?concl e' xs e n pc stk' loc" using len' by-(rule IH) thus ?case by(fastforce intro: bisim1_bisims1.bisim1AAss1 elim!: AAss_\red1r_xt1) next case (bisim1AAss2 e2 n e' xs stk loc pc e1 e3 v1) note IH = \\stk'. \stk = rev vs @ Addr a # stk'; pc < length (compE2 e2); compE2 e2 ! pc = Invoke M (length vs); n + max_vars e' \ length xs; bsok e2 n \ \ ?concl e' xs e2 n pc stk' loc\ note inv = \compE2 (e1\e2\ := e3) ! (length (compE2 e1) + pc) = Invoke M (length vs)\ note bisim = \P,e2,h \ (e', xs) \ (stk, loc, pc, None)\ with inv \length (compE2 e1) + pc < length (compE2 (e1\e2\ := e3))\ have pc: "pc < length (compE2 e2)" by(auto split: if_split_asm dest: bisim1_pc_length_compE2) moreover with inv have "compE2 e2 ! pc = Invoke M (length vs)" by simp moreover with bisim pc obtain vs'' v'' stk'' where "stk = vs'' @ v'' # stk''" and "length vs'' = length vs" by(auto dest!: bisim1_Invoke_stkD) with \stk @ [v1] = rev vs @ Addr a # stk'\ obtain stk''' where stk''': "stk = rev vs @ Addr a # stk'''" and stk: "stk' = stk''' @ [v1]" by(cases stk' rule: rev_cases) auto from \n + max_vars (Val v1\e'\ := e3) \ length xs\ have "n + max_vars e' \ length xs" by simp moreover from \bsok (e1\e2\ := e3) n\ have "bsok e2 n" by simp ultimately have "?concl e' xs e2 n pc stk''' loc" using stk''' by-(rule IH) then obtain e'' xs' where IH': "\red1r P t h (e', xs) (e'', xs')" "call1 e'' = \(a, M, vs)\" and bisim: "P,e2,h \ (e'', xs') \ (rev vs @ Addr a # stk''', loc, pc, None)" by blast from bisim have "P,e1\e2\ := e3,h \ (Val v1\e''\ := e3, xs') \ ((rev vs @ Addr a # stk''') @ [v1], loc, length (compE2 e1) + pc, None)" by(rule bisim1_bisims1.bisim1AAss2) with IH' stk show ?case by(fastforce elim!: AAss_\red1r_xt2) next case (bisim1AAss3 e3 n e' xs stk loc pc e1 e2 v1 v2) note IH = \\stk'. \stk = rev vs @ Addr a # stk'; pc < length (compE2 e3); compE2 e3 ! pc = Invoke M (length vs); n + max_vars e' \ length xs; bsok e3 n \ \ ?concl e' xs e3 n pc stk' loc\ note inv = \compE2 (e1\e2\ := e3) ! (length (compE2 e1) + length (compE2 e2) + pc) = Invoke M (length vs)\ with \length (compE2 e1) + length (compE2 e2) + pc < length (compE2 (e1\e2\ := e3))\ have pc: "pc < length (compE2 e3)" by(simp add: nth_Cons split: nat.split_asm if_split_asm) moreover with inv have "compE2 e3 ! pc = Invoke M (length vs)" by simp moreover with \P,e3,h \ (e', xs) \ (stk, loc, pc, None)\ pc obtain vs'' v'' stk'' where "stk = vs'' @ v'' # stk''" and "length vs'' = length vs" by(auto dest!: bisim1_Invoke_stkD) with \stk @ [v2, v1] = rev vs @ Addr a # stk'\ obtain stk''' where stk''': "stk = rev vs @ Addr a # stk'''" and stk: "stk' = stk''' @ [v2, v1]" by(cases stk' rule: rev_cases) auto from \n + max_vars (Val v1\Val v2\ := e') \ length xs\ have "n + max_vars e' \ length xs" by simp moreover from \bsok (e1\e2\ := e3) n\ have "bsok e3 n" by simp ultimately have "?concl e' xs e3 n pc stk''' loc" using stk''' by-(rule IH) then obtain e'' xs' where IH': "\red1r P t h (e', xs) (e'', xs')" "call1 e'' = \(a, M, vs)\" and bisim: "P,e3,h \ (e'', xs') \ (rev vs @ Addr a # stk''', loc, pc, None)" by blast from bisim have "P,e1\e2\ := e3,h \ (Val v1\Val v2\ := e'', xs') \ ((rev vs @ Addr a # stk''') @ [v2, v1], loc, length (compE2 e1) + length (compE2 e2) + pc, None)" by -(rule bisim1_bisims1.bisim1AAss3, auto) with IH' stk show ?case by(fastforce elim!: AAss_\red1r_xt3) next case bisim1AAss4 thus ?case by simp next case bisim1ALength thus ?case by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1ALength dest: bisim1_pc_length_compE2 elim!: ALength_\red1r_xt) next case bisim1FAcc thus ?case by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1FAcc dest: bisim1_pc_length_compE2 elim!: FAcc_\red1r_xt) next case bisim1FAss1 thus ?case apply(auto split: if_split_asm intro: bisim1_bisims1.bisim1FAss1 dest: bisim1_pc_length_compE2 elim!: FAss_\red1r_xt1) by(fastforce intro: bisim1_bisims1.bisim1FAss1 elim!: FAss_\red1r_xt1) next case (bisim1FAss2 e2 n e' xs stk loc pc e1 F D v1) note IH = \\stk'. \stk = rev vs @ Addr a # stk'; pc < length (compE2 e2); compE2 e2 ! pc = Invoke M (length vs); n + max_vars e' \ length xs; bsok e2 n \ \ ?concl e' xs e2 n pc stk' loc\ note inv = \compE2 (e1\F{D} := e2) ! (length (compE2 e1) + pc) = Invoke M (length vs)\ with \length (compE2 e1) + pc < length (compE2 (e1\F{D} := e2))\ have pc: "pc < length (compE2 e2)" by(simp split: if_split_asm nat.split_asm add: nth_Cons) moreover with inv have "compE2 e2 ! pc = Invoke M (length vs)" by simp moreover with \P,e2,h \ (e', xs) \ (stk, loc, pc, None)\ pc obtain vs'' v'' stk'' where "stk = vs'' @ v'' # stk''" and "length vs'' = length vs" by(auto dest!: bisim1_Invoke_stkD) with \stk @ [v1] = rev vs @ Addr a # stk'\ obtain stk''' where stk''': "stk = rev vs @ Addr a # stk'''" and stk: "stk' = stk''' @ [v1]" by(cases stk' rule: rev_cases) auto from \n + max_vars (Val v1\F{D} := e') \ length xs\ have "n + max_vars e' \ length xs" by simp moreover from \bsok (e1\F{D} := e2) n\ have "bsok e2 n" by simp ultimately have "?concl e' xs e2 n pc stk''' loc" using stk''' by-(rule IH) then obtain e'' xs' where IH': "\red1r P t h (e', xs) (e'', xs')" "call1 e'' = \(a, M, vs)\" and bisim: "P,e2,h \ (e'', xs') \ (rev vs @ Addr a # stk''', loc, pc, None)" by blast from bisim have "P,e1\F{D} := e2,h \ (Val v1\F{D} := e'', xs') \ ((rev vs @ Addr a # stk''') @ [v1], loc, length (compE2 e1) + pc, None)" by(rule bisim1_bisims1.bisim1FAss2) with IH' stk show ?case by(fastforce elim!: FAss_\red1r_xt2) next case bisim1FAss3 thus ?case by simp next case (bisim1CAS1 e n e' xs loc pc e2 e3 D F) note IH = \\ pc < length (compE2 e); compE2 e ! pc = Invoke M (length vs); n + max_vars e' \ length xs; bsok e n \ \ ?concl e' xs e n pc stk' loc\ note bisim = \P,e,h \ (e', xs) \ (rev vs @ Addr a # stk', loc, pc, None)\ note len = \n + max_vars _ \ length xs\ hence len': "n + max_vars e' \ length xs" by simp note inv = \compE2 _ ! pc = Invoke M (length vs)\ with \pc < length _\ bisim have pc: "pc < length (compE2 e)" by(auto split: if_split_asm dest: bisim1_pc_length_compE2) moreover with inv have "compE2 e ! pc = Invoke M (length vs)" by simp moreover from \bsok _ n\ have "bsok e n" by simp ultimately have "?concl e' xs e n pc stk' loc" using len' by-(rule IH) thus ?case by(fastforce intro: bisim1_bisims1.bisim1CAS1 elim!: CAS_\red1r_xt1) next case (bisim1CAS2 e2 n e' xs stk loc pc e1 e3 D F v1) note IH = \\stk'. \stk = rev vs @ Addr a # stk'; pc < length (compE2 e2); compE2 e2 ! pc = Invoke M (length vs); n + max_vars e' \ length xs; bsok e2 n \ \ ?concl e' xs e2 n pc stk' loc\ note inv = \compE2 _ ! (length (compE2 e1) + pc) = Invoke M (length vs)\ note bisim = \P,e2,h \ (e', xs) \ (stk, loc, pc, None)\ with inv \length (compE2 e1) + pc < length (compE2 _)\ have pc: "pc < length (compE2 e2)" by(auto split: if_split_asm dest: bisim1_pc_length_compE2) moreover with inv have "compE2 e2 ! pc = Invoke M (length vs)" by simp moreover with bisim pc obtain vs'' v'' stk'' where "stk = vs'' @ v'' # stk''" and "length vs'' = length vs" by(auto dest!: bisim1_Invoke_stkD) with \stk @ [v1] = rev vs @ Addr a # stk'\ obtain stk''' where stk''': "stk = rev vs @ Addr a # stk'''" and stk: "stk' = stk''' @ [v1]" by(cases stk' rule: rev_cases) auto from \n + max_vars _ \ length xs\ have "n + max_vars e' \ length xs" by simp moreover from \bsok _ n\ have "bsok e2 n" by simp ultimately have "?concl e' xs e2 n pc stk''' loc" using stk''' by-(rule IH) then obtain e'' xs' where IH': "\red1r P t h (e', xs) (e'', xs')" "call1 e'' = \(a, M, vs)\" and bisim: "P,e2,h \ (e'', xs') \ (rev vs @ Addr a # stk''', loc, pc, None)" by blast from bisim have "P,e1\compareAndSwap(D\F, e2, e3),h \ (Val v1\compareAndSwap(D\F, e'', e3), xs') \ ((rev vs @ Addr a # stk''') @ [v1], loc, length (compE2 e1) + pc, None)" by(rule bisim1_bisims1.bisim1CAS2) with IH' stk show ?case by(fastforce elim!: CAS_\red1r_xt2) next case (bisim1CAS3 e3 n e' xs stk loc pc e1 e2 D F v1 v2) note IH = \\stk'. \stk = rev vs @ Addr a # stk'; pc < length (compE2 e3); compE2 e3 ! pc = Invoke M (length vs); n + max_vars e' \ length xs; bsok e3 n \ \ ?concl e' xs e3 n pc stk' loc\ note inv = \compE2 _ ! (length (compE2 e1) + length (compE2 e2) + pc) = Invoke M (length vs)\ with \length (compE2 e1) + length (compE2 e2) + pc < length (compE2 _)\ have pc: "pc < length (compE2 e3)" by(simp add: nth_Cons split: nat.split_asm if_split_asm) moreover with inv have "compE2 e3 ! pc = Invoke M (length vs)" by simp moreover with \P,e3,h \ (e', xs) \ (stk, loc, pc, None)\ pc obtain vs'' v'' stk'' where "stk = vs'' @ v'' # stk''" and "length vs'' = length vs" by(auto dest!: bisim1_Invoke_stkD) with \stk @ [v2, v1] = rev vs @ Addr a # stk'\ obtain stk''' where stk''': "stk = rev vs @ Addr a # stk'''" and stk: "stk' = stk''' @ [v2, v1]" by(cases stk' rule: rev_cases) auto from \n + max_vars _ \ length xs\ have "n + max_vars e' \ length xs" by simp moreover from \bsok _ n\ have "bsok e3 n" by simp ultimately have "?concl e' xs e3 n pc stk''' loc" using stk''' by-(rule IH) then obtain e'' xs' where IH': "\red1r P t h (e', xs) (e'', xs')" "call1 e'' = \(a, M, vs)\" and bisim: "P,e3,h \ (e'', xs') \ (rev vs @ Addr a # stk''', loc, pc, None)" by blast from bisim have "P,e1\compareAndSwap(D\F, e2, e3), h \ (Val v1\compareAndSwap(D\F, Val v2, e''), xs') \ ((rev vs @ Addr a # stk''') @ [v2, v1], loc, length (compE2 e1) + length (compE2 e2) + pc, None)" by -(rule bisim1_bisims1.bisim1CAS3, auto) with IH' stk show ?case by(fastforce elim!: CAS_\red1r_xt3) next case (bisim1Call1 obj n obj' xs loc pc ps M') note IH = \\pc < length (compE2 obj); compE2 obj ! pc = Invoke M (length vs); n + max_vars obj' \ length xs; bsok obj n \ \ ?concl obj' xs obj n pc stk' loc\ note bisim = \P,obj,h \ (obj', xs) \ (rev vs @ Addr a # stk', loc, pc, None)\ note len = \n + max_vars (obj'\M'(ps)) \ length xs\ hence len': "n + max_vars obj' \ length xs" by simp from \bsok (obj\M'(ps)) n\ have bsok: "bsok obj n" by simp note inv = \compE2 (obj\M'(ps)) ! pc = Invoke M (length vs)\ with \pc < length (compE2 (obj\M'(ps)))\ bisim have pc: "pc < length (compE2 obj) \ ps = [] \ pc = length (compE2 obj)" by(cases ps)(auto split: if_split_asm dest: bisim1_pc_length_compE2) thus ?case proof assume "pc < length (compE2 obj)" moreover with inv have "compE2 obj ! pc = Invoke M (length vs)" by simp ultimately have "?concl obj' xs obj n pc stk' loc" using len' bsok by(rule IH) thus ?thesis by(fastforce intro: bisim1_bisims1.bisim1Call1 elim!: Call_\red1r_obj) next assume [simp]: "ps = [] \ pc = length (compE2 obj)" with inv have [simp]: "vs = []" "M' = M" by simp_all with bisim have [simp]: "vs = []" "stk' = []" by(auto dest: bisim1_pc_length_compE2D) with bisim len' bsok have "\red1r P t h (obj', xs) (addr a, loc)" by(auto intro: bisim1_Val_\red1r simp add: bsok_def) moreover have "P,obj\M([]),h \ (addr a\M([]), loc) \ ([Addr a], loc, length (compE2 obj), None)" by(rule bisim1_bisims1.bisim1Call1[OF bisim1Val2]) simp_all ultimately show ?thesis by auto(fastforce elim!: Call_\red1r_obj) qed next case (bisim1CallParams ps n ps' xs stk loc pc obj M' v) note IH = \\stk'. \stk = rev vs @ Addr a # stk'; pc < length (compEs2 ps); compEs2 ps ! pc = Invoke M (length vs); n + max_varss ps' \ length xs; bsoks ps n \ \ ?concls ps' xs ps n pc stk' loc\ note bisim = \P,ps,h \ (ps', xs) [\] (stk, loc, pc, None)\ note len = \n + max_vars (Val v\M'(ps')) \ length xs\ hence len': "n + max_varss ps' \ length xs" by simp note stk = \stk @ [v] = rev vs @ Addr a # stk'\ note inv = \compE2 (obj\M'(ps)) ! (length (compE2 obj) + pc) = Invoke M (length vs)\ from \bsok (obj\M'(ps)) n\ have bsok: "bsoks ps n" by simp from \length (compE2 obj) + pc < length (compE2 (obj\M'(ps)))\ have "pc < length (compEs2 ps) \ pc = length (compEs2 ps)" by(auto) thus ?case proof assume pc: "pc < length (compEs2 ps)" moreover with inv have "compEs2 ps ! pc = Invoke M (length vs)" by simp moreover with bisim pc obtain vs'' v'' stk'' where "stk = vs'' @ v'' # stk''" and "length vs'' = length vs" by(auto dest!: bisims1_Invoke_stkD) with \stk @ [v] = rev vs @ Addr a # stk'\ obtain stk''' where stk''': "stk = rev vs @ Addr a # stk'''" and stk: "stk' = stk''' @ [v]" by(cases stk' rule: rev_cases) auto note len' stk''' ultimately have "?concls ps' xs ps n pc stk''' loc" using bsok by-(rule IH) then obtain es'' xs' where IH': "\reds1r P t h (ps', xs) (es'', xs')" "calls1 es'' = \(a, M, vs)\" and bisim: "P,ps,h \ (es'', xs') [\] (rev vs @ Addr a # stk''', loc, pc, None)" by blast from bisim have "P,obj\M'(ps),h \ (Val v\M'(es''), xs') \ ((rev vs @ Addr a # stk''') @ [v], loc, length (compE2 obj) + pc, None)" by(rule bisim1_bisims1.bisim1CallParams) with IH' stk show ?case by(fastforce elim!: Call_\red1r_param simp add: is_vals_conv) next assume [simp]: "pc = length (compEs2 ps)" from bisim obtain vs' where [simp]: "stk = rev vs'" and psvs': "length ps = length vs'" by(auto dest: bisims1_pc_length_compEs2D) from inv have [simp]: "M' = M" and vsps: "length vs = length ps" by simp_all with stk psvs' have [simp]: "v = Addr a" "stk' = []" "vs' = vs" by simp_all from bisim len' bsok have "\reds1r P t h (ps', xs) (map Val vs, loc)" by(auto intro: bisims1_Val_\Reds1r simp add: bsoks_def) moreover from bisims1_map_Val_append[OF bisims1Nil vsps[symmetric], simplified, of P h loc] have "P,obj\M(ps),h \ (addr a\M(map Val vs), loc) \ (rev vs @ [Addr a], loc, length (compE2 obj) + length (compEs2 ps), None)" by(rule bisim1_bisims1.bisim1CallParams) ultimately show ?thesis by(fastforce elim!: Call_\red1r_param) qed next case bisim1BlockSome1 thus ?case by simp next case bisim1BlockSome2 thus ?case by simp next case (bisim1BlockSome4 e n e' xs loc pc V T v) note IH = \\pc < length (compE2 e); compE2 e ! pc = Invoke M (length vs); Suc n + max_vars e' \ length xs; bsok e (Suc n) \ \ ?concl e' xs e (Suc V) pc stk' loc\ from \Suc (Suc pc) < length (compE2 {V:T=\v\; e})\ have "pc < length (compE2 e)" by simp moreover from \compE2 {V:T=\v\; e} ! Suc (Suc pc) = Invoke M (length vs)\ have "compE2 e ! pc = Invoke M (length vs)" by simp moreover note len = \n + max_vars {V:T=None; e'} \ length xs\ hence "Suc n + max_vars e' \ length xs" by simp moreover from \bsok {V:T=\v\; e} n\ have "bsok e (Suc n)" by simp ultimately have "?concl e' xs e (Suc V) pc stk' loc" by(rule IH) then obtain e'' xs' where red: "\red1r P t h (e', xs) (e'', xs')" and bisim': "P,e,h \ (e'', xs') \ (rev vs @ Addr a # stk', loc, pc, None)" and call: "call1 e'' = \(a, M, vs)\" by blast from red have "\red1r P t h ({V:T=None; e'}, xs) ({V:T=None; e''}, xs')" by(rule Block_None_\red1r_xt) with bisim' call show ?case by(fastforce intro: bisim1_bisims1.bisim1BlockSome4) next case (bisim1BlockNone e n e' xs loc pc V T) note IH = \\pc < length (compE2 e); compE2 e ! pc = Invoke M (length vs); Suc n + max_vars e' \ length xs; bsok e (Suc n) \ \ ?concl e' xs e (Suc V) pc stk' loc\ from \pc < length (compE2 {V:T=None; e})\ have "pc < length (compE2 e)" by simp moreover from \compE2 {V:T=None; e} ! pc = Invoke M (length vs)\ have "compE2 e ! pc = Invoke M (length vs)" by simp moreover note len = \n + max_vars {V:T=None; e'} \ length xs\ hence "Suc n + max_vars e' \ length xs" by simp moreover from \bsok {V:T=None; e} n\ have "bsok e (Suc n)" by simp ultimately have "?concl e' xs e (Suc V) pc stk' loc" by(rule IH) then obtain e'' xs' where red: "\red1r P t h (e', xs) (e'', xs')" and bisim': "P,e,h \ (e'', xs') \ (rev vs @ Addr a # stk', loc, pc, None)" and call: "call1 e'' = \(a, M, vs)\" by blast from red have "\red1r P t h ({V:T=None; e'}, xs) ({V:T=None; e''}, xs')" by(rule Block_None_\red1r_xt) with bisim' call show ?case by(fastforce intro: bisim1_bisims1.bisim1BlockNone) next case bisim1Sync1 thus ?case apply(auto split: if_split_asm intro: bisim1_bisims1.bisim1Sync1 dest: bisim1_pc_length_compE2 elim!: Sync_\red1r_xt) by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1Sync1 elim!: Sync_\red1r_xt) next case bisim1Sync2 thus ?case by simp next case bisim1Sync3 thus ?case by simp next case bisim1Sync4 thus ?case apply(auto split: if_split_asm intro: bisim1_bisims1.bisim1Sync4 dest: bisim1_pc_length_compE2 elim!: InSync_\red1r_xt) by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1Sync4 elim!: InSync_\red1r_xt) next case bisim1Sync5 thus ?case by simp next case bisim1Sync6 thus ?case by simp next case bisim1Sync7 thus ?case by simp next case bisim1Sync8 thus ?case by simp next case bisim1Sync9 thus ?case by simp next case bisim1InSync thus ?case by simp next case bisim1Seq1 thus ?case apply(auto split: if_split_asm intro: bisim1_bisims1.bisim1Seq1 dest: bisim1_pc_length_compE2 elim!: Seq_\red1r_xt) by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1Seq1 elim!: Seq_\red1r_xt) next case bisim1Seq2 thus ?case by(auto split: if_split_asm intro: bisim1_bisims1.bisim1Seq2 dest: bisim1_pc_length_compE2) next case bisim1Cond1 thus ?case apply(clarsimp split: if_split_asm) apply(fastforce intro!: exI intro: bisim1_bisims1.bisim1Cond1 elim!: Cond_\red1r_xt) by(fastforce dest: bisim1_pc_length_compE2) next case bisim1CondThen thus ?case apply(clarsimp split: if_split_asm) apply(fastforce intro!: exI intro: bisim1_bisims1.bisim1CondThen) by(fastforce dest: bisim1_pc_length_compE2) next case bisim1CondElse thus ?case by(clarsimp split: if_split_asm)(fastforce intro!: exI intro: bisim1_bisims1.bisim1CondElse) next case bisim1While1 thus ?case by simp next case bisim1While3 thus ?case apply(auto split: if_split_asm intro: bisim1_bisims1.bisim1While3 dest: bisim1_pc_length_compE2 elim!: Cond_\red1r_xt) by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1While3 elim!: Cond_\red1r_xt) next case bisim1While4 thus ?case apply(auto split: if_split_asm intro: bisim1_bisims1.bisim1While4 dest: bisim1_pc_length_compE2 elim!: Seq_\red1r_xt) by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1While4 elim!: Seq_\red1r_xt) next case bisim1While6 thus ?case by simp next case bisim1While7 thus ?case by simp next case bisim1Throw1 thus ?case by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1Throw1 dest: bisim1_pc_length_compE2 elim!: Throw_\red1r_xt) next case bisim1Try thus ?case apply(auto split: if_split_asm intro: bisim1_bisims1.bisim1Try dest: bisim1_pc_length_compE2 elim!: Try_\red1r_xt) by(fastforce split: if_split_asm intro: bisim1_bisims1.bisim1Try elim!: Try_\red1r_xt) next case bisim1TryCatch1 thus ?case by simp next case (bisim1TryCatch2 e n e' xs loc pc e1 C V) note IH = \\pc < length (compE2 e); compE2 e ! pc = Invoke M (length vs); Suc n + max_vars e' \ length xs; bsok e (Suc n) \ \ ?concl e' xs e (Suc V) pc stk' loc\ from \Suc (Suc (length (compE2 e1) + pc)) < length (compE2 (try e1 catch(C V) e))\ have "pc < length (compE2 e)" by simp moreover from \compE2 (try e1 catch(C V) e) ! Suc (Suc (length (compE2 e1) + pc)) = Invoke M (length vs)\ have "compE2 e ! pc = Invoke M (length vs)" by simp moreover note len = \n + max_vars {V:Class C=None; e'} \ length xs\ hence "Suc n + max_vars e' \ length xs" by simp moreover from \bsok (try e1 catch(C V) e) n\ have "bsok e (Suc n)" by simp ultimately have "?concl e' xs e (Suc V) pc stk' loc" by(rule IH) then obtain e'' xs' where red: "\red1r P t h (e', xs) (e'', xs')" and bisim': "P,e,h \ (e'', xs') \ (rev vs @ Addr a # stk', loc, pc, None)" and call: "call1 e'' = \(a, M, vs)\" by blast from red have "\red1r P t h ({V:Class C=None; e'}, xs) ({V:Class C=None; e''}, xs')" by(rule Block_None_\red1r_xt) with bisim' call show ?case by(fastforce intro: bisim1_bisims1.bisim1TryCatch2) next case bisims1Nil thus ?case by simp next case (bisims1List1 e n e' xs loc pc es) note IH = \\pc < length (compE2 e); compE2 e ! pc = Invoke M (length vs); n + max_vars e' \ length xs; bsok e n \ \ ?concl e' xs e n pc stk' loc\ note bisim = \P,e,h \ (e', xs) \ (rev vs @ Addr a # stk', loc, pc, None)\ note len = \n + max_varss (e' # es) \ length xs\ hence len': "n + max_vars e' \ length xs" by simp from \bsoks (e # es) n\ have bsok: "bsok e n" by simp note inv = \compEs2 (e # es) ! pc = Invoke M (length vs)\ with \pc < length (compEs2 (e # es))\ bisim have pc: "pc < length (compE2 e)" by(cases es)(auto split: if_split_asm dest: bisim1_pc_length_compE2) moreover with inv have "compE2 e ! pc = Invoke M (length vs)" by simp ultimately have "?concl e' xs e n pc stk' loc" using len' bsok by(rule IH) thus ?case by(fastforce intro: bisim1_bisims1.bisims1List1 elim!: \red1r_inj_\reds1r) next case (bisims1List2 es n es' xs stk loc pc e v) note IH = \\stk'. \stk = rev vs @ Addr a # stk'; pc < length (compEs2 es); compEs2 es ! pc = Invoke M (length vs); n + max_varss es' \ length xs; bsoks es n \ \ ?concls es' xs es n pc stk' loc\ note bisim = \P,es,h \ (es', xs) [\] (stk, loc, pc, None)\ note len = \n + max_varss (Val v # es') \ length xs\ hence len': "n + max_varss es' \ length xs" by simp from \bsoks (e # es) n\ have bsok: "bsoks es n" by simp note stk = \stk @ [v] = rev vs @ Addr a # stk'\ note inv = \compEs2 (e # es) ! (length (compE2 e) + pc) = Invoke M (length vs)\ from \length (compE2 e) + pc < length (compEs2 (e # es))\ have pc: "pc < length (compEs2 es)" by auto moreover with inv have "compEs2 es ! pc = Invoke M (length vs)" by simp moreover with bisim pc obtain vs'' v'' stk'' where "stk = vs'' @ v'' # stk''" and "length vs'' = length vs" by(auto dest!: bisims1_Invoke_stkD) with \stk @ [v] = rev vs @ Addr a # stk'\ obtain stk''' where stk''': "stk = rev vs @ Addr a # stk'''" and stk: "stk' = stk''' @ [v]" by(cases stk' rule: rev_cases) auto note len' stk''' ultimately have "?concls es' xs es n pc stk''' loc" using bsok by-(rule IH) then obtain es'' xs' where red: "\reds1r P t h (es', xs) (es'', xs')" and call: "calls1 es'' = \(a, M, vs)\" and bisim: "P,es,h \ (es'', xs') [\] (rev vs @ Addr a # stk''', loc, pc, None)" by blast from bisim have "P,e#es,h \ (Val v # es'', xs') [\] ((rev vs @ Addr a # stk''') @ [v], loc, length (compE2 e) + pc, None)" by(rule bisim1_bisims1.bisims1List2) moreover from red have "\reds1r P t h (Val v # es', xs) (Val v # es'', xs')" by(rule \reds1r_cons_\reds1r) ultimately show ?case using stk call by fastforce qed end declare split_beta [simp del] context J1_JVM_conf_read begin lemma \Red1_simulates_exec_1_\: assumes wf: "wf_J1_prog P" and exec: "exec_1_d (compP2 P) t (Normal (xcp, h, frs)) ta (Normal (xcp', h', frs'))" and bisim: "bisim1_list1 t h (e, xs) exs xcp frs" and \: "\Move2 (compP2 P) (xcp, h, frs)" shows "h = h' \ (\e' xs' exs'. (if sim21_size (compP2 P) (xcp', frs') (xcp, frs) then \Red1r else \Red1t) P t h ((e, xs), exs) ((e', xs'), exs') \ bisim1_list1 t h (e', xs') exs' xcp' frs')" using bisim proof(cases) case (bl1_Normal stk loc C M pc FRS Ts T body D) hence [simp]: "frs = (stk, loc, C, M, pc) # FRS" and conf: "compTP P \ t: (xcp, h, (stk, loc, C, M, pc) # FRS) \" and sees: "P \ C sees M: Ts\T = \body\ in D" and bisim: "P,blocks1 0 (Class D#Ts) body,h \ (e, xs) \ (stk, loc, pc, xcp)" and lenxs: "max_vars e \ length xs" and bisims: "list_all2 (bisim1_fr P h) exs FRS" by auto from sees_method_compP[OF sees, where f="\C M Ts T. compMb2"] have sees': "compP2 P \ C sees M: Ts\T = \(max_stack body, max_vars body, compE2 body @ [Return], compxE2 body 0 0)\ in D" by(simp add: compP2_def compMb2_def) from bisim have pc: "pc \ length (compE2 body)" by(auto dest: bisim1_pc_length_compE2) from conf have hconf: "hconf h" "preallocated h" by(simp_all add: correct_state_def) from sees wf have bsok: "bsok (blocks1 0 (Class D # Ts) body) 0" by(auto dest!: sees_wf_mdecl simp add: bsok_def wf_mdecl_def WT1_expr_locks) from exec obtain check: "check (compP2 P) (xcp, h, frs)" and exec: "compP2 P,t \ (xcp, h, frs) -ta-jvm\ (xcp', h', frs')" by(rule jvmd_NormalE)(auto simp add: exec_1_iff) from wt_compTP_compP2[OF wf] exec conf have conf': "compTP P \ t:(xcp', h', frs') \" by(auto intro: BV_correct_1) from conf have tconf: "P,h \ t \t" unfolding correct_state_def by(simp add: compP2_def tconf_def) show ?thesis proof(cases xcp) case [simp]: None from exec have execi: "(ta, xcp', h', frs') \ exec_instr (instrs_of (compP2 P) C M ! pc) (compP2 P) t h stk loc C M pc FRS" by(simp add: exec_1_iff) show ?thesis proof(cases "pc < length (compE2 body)") case True with execi sees' have execi: "(ta, xcp', h', frs') \ exec_instr (compE2 body ! pc) (compP2 P) t h stk loc C M pc FRS" by(simp) from \ sees' True have \i: "\move2 (compP2 P) h stk body pc None" by(simp add: \move2_iff) show ?thesis proof(cases "length frs' = Suc (length FRS)") case False with execi sees True compE2_not_Return[of body] have "(\M n. compE2 body ! pc = Invoke M n)" apply(cases "compE2 body ! pc") apply(auto split: if_split_asm sum.split_asm simp add: split_beta compP2_def compMb2_def) apply(metis in_set_conv_nth)+ done then obtain MM n where ins: "compE2 body ! pc = Invoke MM n" by blast with bisim1_Invoke_stkD[OF bisim[unfolded None], of MM n] True obtain vs' v' stk' where [simp]: "stk = vs' @ v' # stk'" "n = length vs'" by auto from check sees True ins have "is_Ref v'" by(auto split: if_split_asm simp add: split_beta compP2_def compMb2_def check_def) moreover from execi sees True ins False sees' have "v' \ Null" by auto ultimately obtain a' where [simp]: "v' = Addr a'" by(auto simp add: is_Ref_def) from bisim have Bisim': "P,blocks1 0 (Class D#Ts) body,h \ (e, xs) \ (rev (rev vs') @ Addr a' # stk', loc, pc, None)" by simp from bisim1_Invoke_\Red[OF this _ _ _ bsok, of MM t] True ins lenxs obtain e' xs' where red: "\red1r P t h (e, xs) (e', xs')" and bisim': "P,blocks1 0 (Class D#Ts) body,h \ (e', xs') \ (rev (rev vs') @ Addr a' # stk', loc, pc, None)" and call': "call1 e' = \(a', MM, rev vs')\" by auto from red have Red: "\Red1r P t h ((e, xs), exs) ((e', xs'), exs)" by(rule \red1r_into_\Red1r) from False execi True check ins sees' obtain U' Ts' T' meth D' where ha': "typeof_addr h a' = \U'\" and Sees': "compP2 P \ class_type_of U' sees MM:Ts' \ T' = \meth\ in D'" by(auto simp add: check_def has_method_def split: if_split_asm)(auto split: extCallRet.split_asm) from sees_method_compPD[OF Sees'[unfolded compP2_def]] obtain body' where Sees: "P \ class_type_of U' sees MM:Ts' \ T'=\body'\ in D'" and [simp]: "meth = (max_stack body', max_vars body', compE2 body' @ [Return], compxE2 body' 0 0)" by(auto simp add: compMb2_def) let ?e = "blocks1 0 (Class D'#Ts') body'" let ?xs = "Addr a' # rev vs' @ replicate (max_vars body') undefined_value" let ?e'xs' = "(e', xs')" let ?f = "(stk, loc, C, M, pc)" let ?f' = "([],Addr a' # rev vs' @ replicate (max_vars body') undefined_value, D', MM, 0)" from execi pc ins False ha' Sees' sees' have [simp]: "xcp' = None" "ta = \" "frs' = ?f' # ?f # FRS" "h' = h" by(auto split: if_split_asm simp add: split_beta) from bisim' have bisim'': "P,blocks1 0 (Class D#Ts) body,h \ (e', xs') \ (rev (rev vs') @ Addr a' # stk', loc, pc, None)" by simp have "n = length vs'" by simp from conf' Sees' ins sees' True have "n = length Ts'" apply(auto simp add: correct_state_def) apply(drule (1) sees_method_fun)+ apply(auto dest: sees_method_idemp sees_method_fun) done with \n = length vs'\ have vs'Ts': "length (rev vs') = length Ts'" by simp with call' ha' Sees have "True,P,t \1 \(e', xs')/exs,h\ -\\ \(?e, ?xs)/ (e', xs') # exs, h\" by(rule red1Call) hence "True,P,t \1 \(e', xs')/exs,h\ -\\ \(?e, ?xs)/ ?e'xs' # exs, h\" by(simp) moreover from call' Sees ha' have "\Move1 P h ((e', xs'), exs)" by(auto simp add: synthesized_call_def dest!: \move1_not_call1[where P=P and h=h] dest: sees_method_fun) ultimately have "\Red1t P t h ((e', xs'), exs) ((?e, ?xs), ?e'xs' # exs)" by auto moreover have "bisim1_list1 t h (?e, ?xs) (?e'xs' # exs) None (?f' # ?f # FRS)" proof from conf' show "compTP P \ t:(None, h, ?f' # ?f # FRS) \" by simp from Sees show "P \ D' sees MM: Ts'\T' = \body'\ in D'" by(rule sees_method_idemp) show "P,blocks1 0 (Class D'#Ts') body',h \ (blocks1 0 (Class D'#Ts') body', ?xs) \ ([], Addr a' # rev vs' @ replicate (max_vars body') undefined_value, 0, None)" by(rule bisim1_refl) show "max_vars (blocks1 0 (Class D'#Ts') body') \ length ?xs" using vs'Ts' by(simp add: blocks1_max_vars) from sees have "bisim1_fr P h ?e'xs' ?f" proof show "P,blocks1 0 (Class D#Ts) body,h \ (e', xs') \ (stk, loc, pc, None)" using bisim'' by simp from call' show "call1 e' = \(a', MM, rev vs')\" . from red have xs'xs: "length xs' = length xs" by(rule \red1r_preserves_len) with red lenxs show "max_vars e' \ length xs'" by(auto dest: \red1r_max_vars) qed with bisims show "list_all2 (bisim1_fr P h) (?e'xs' # exs) (?f # FRS)" by simp qed ultimately show ?thesis using Red by auto(blast intro: rtranclp_trans rtranclp_tranclp_tranclp tranclp_into_rtranclp)+ next case True note pc = \pc < length (compE2 body)\ with execi True have "\stk' loc' pc'. frs' = (stk', loc', C, M, pc') # FRS" by(cases "(compE2 body @ [Return]) ! pc")(auto split: if_split_asm sum.split_asm simp: split_beta, auto split: extCallRet.splits) then obtain stk' loc' pc' where [simp]: "frs' = (stk', loc', C, M, pc') # FRS" by blast from conf obtain ST where "compP2 P,h \ stk [:\] ST" by(auto simp add: correct_state_def conf_f_def2) hence ST: "P,h \ stk [:\] ST" by(rule List.list_all2_mono)(simp add: compP2_def) from execi sees pc check have exec': "exec_move_d P t (blocks1 0 (Class D#Ts) body) h (stk, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" apply(auto simp add: compP2_def compMb2_def exec_move_def check_def exec_meth_instr split: if_split_asm sum.split_asm) apply(cases "compE2 body ! pc") apply(auto simp add: neq_Nil_conv split_beta split: if_split_asm sum.split_asm) apply(force split: extCallRet.split_asm) apply(cases "compE2 body ! pc", auto simp add: split_beta neq_Nil_conv split: if_split_asm sum.split_asm) done from red1_simulates_exec_instr[OF wf hconf tconf bisim this _ bsok ST] lenxs \i obtain e'' xs'' where bisim': "P,blocks1 0 (Class D#Ts) body,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red: "(if xcp' = None \ pc < pc' then \red1r else \red1t) P t h (e, xs) (e'', xs'')" and [simp]: "h' = h" by(auto simp del: blocks1.simps) have Red: "(if sim21_size (compP2 P) (xcp', frs') (xcp, frs) then \Red1r else \Red1t) P t h ((e, xs), exs) ((e'', xs''), exs)" proof(cases "xcp' = None \ pc < pc'") case True from bisim bisim' have "pc \ Suc (length (compE2 body))" "pc' \ Suc (length (compE2 body))" by(auto dest: bisim1_pc_length_compE2) moreover { fix a assume "xcp' = \a\" with exec' have "pc = pc'" by(auto dest: exec_move_raise_xcp_pcD) } ultimately have "sim21_size (compP2 P) (xcp', frs') (xcp, frs)" using sees True by(auto simp add: sim21_size_def)(auto simp add: compP2_def compMb2_def intro!: sim21_size_aux.intros) with red True show ?thesis by simp(rule \red1r_into_\Red1r) next case False thus ?thesis using red by(auto intro: \red1t_into_\Red1t \red1r_into_\Red1r) qed moreover from red lenxs have "max_vars e'' \ length xs''" apply(auto dest: \red1r_max_vars \red1r_preserves_len \red1t_max_vars \red1t_preserves_len split: if_split_asm) apply(frule \red1r_max_vars \red1t_max_vars, drule \red1r_preserves_len \red1t_preserves_len, simp)+ done with conf' sees bisim' have "bisim1_list1 t h (e'', xs'') exs xcp' ((stk', loc', C, M, pc') # FRS)" unfolding \frs' = (stk', loc', C, M, pc') # FRS\ \h' = h\ using bisims by(rule bisim1_list1.bl1_Normal) ultimately show ?thesis by(auto split del: if_split) qed next case False with pc have [simp]: "pc = length (compE2 body)" by simp with execi sees have [simp]: "xcp' = None" by(cases "compE2 body ! pc")(auto split: if_split_asm simp add: compP2_def compMb2_def split_beta) from bisim have Bisim: "P,blocks1 0 (Class D#Ts) body,h \ (e, xs) \ (stk, loc, length (compE2 (blocks1 0 (Class D#Ts) body)), None)" by simp then obtain v where [simp]: "stk = [v]" by(blast dest: bisim1_pc_length_compE2D) with Bisim lenxs bsok have red: "\red1r P t h (e, xs) (Val v, loc)" by clarify (erule bisim1_Val_\red1r[where n=0], simp_all add: bsok_def) hence Red: "\Red1r P t h ((e, xs), exs) ((Val v, loc), exs)" by(rule \red1r_into_\Red1r) show ?thesis proof(cases "FRS") case [simp]: Nil with bisims have [simp]: "exs = []" by simp with exec sees' have [simp]: "ta = \" "xcp' = None" "h' = h" "frs' = []" by(auto simp add: exec_1_iff) from hconf have "bisim1_list1 t h (Val v, loc) [] None []" by(rule bl1_finalVal) then show ?thesis using Red by(auto intro: rtranclp.rtrancl_into_rtrancl rtranclp_into_tranclp1 simp del: \Red1_conv simp add: sim21_size_def) next case (Cons f' FRS') then obtain stk'' loc'' C'' M'' pc'' where [simp]: "FRS = (stk'', loc'', C'', M'', pc'') # FRS'" by(cases f') fastforce from bisims obtain e'' xs'' EXS' where [simp]: "exs = (e'', xs'') # EXS'" by(auto simp add: list_all2_Cons2) with bisims have "bisim1_fr P h (e'', xs'') (stk'', loc'', C'', M'', pc'')" by simp then obtain E'' Ts'' T'' body'' D'' a'' M''' vs'' where [simp]: "e'' = E''" and sees'': "P \ C'' sees M'':Ts''\T'' = \body''\ in D''" and bisim'': "P,blocks1 0 (Class D''#Ts'') body'',h \ (E'', xs'') \ (stk'', loc'', pc'', None)" and call'': "call1 E'' = \(a'', M''', vs'')\" and lenxs'': "max_vars E'' \ length xs''" by(cases) fastforce let ?ee' = "inline_call (Val v) E''" let ?e' = "?ee'" let ?xs' = "xs''" from bisim'' call'' have pc'': "pc'' < length (compE2 (blocks1 0 (Class D''#Ts'') body''))" by(rule bisim1_call_pcD) hence pc'': "pc'' < length (compE2 body'')" by simp with sees_method_compP[OF sees'', where f="\C M Ts T. compMb2"] sees_method_compP[OF sees, where f="\C M Ts T. compMb2"] conf obtain ST LT where \: "compTP P C'' M'' ! pc'' = \(ST, LT)\" and conff: "conf_f (compP (\C M Ts T. compMb2) P) h (ST, LT) (compE2 body'' @ [Return]) (stk'', loc'', C'', M'', pc'')" and ins: "(compE2 body'' @ [Return]) ! pc'' = Invoke M (length Ts)" unfolding correct_state_def by(fastforce simp add: compP2_def compMb2_def dest: sees_method_fun) from bisim1_callD[OF bisim'' call'', of M "length Ts"] ins pc'' have [simp]: "M''' = M" by simp from call'' have "call1 E'' = \(a'', M''', vs'')\" by simp have "True,P,t \1 \(Val v, loc)/(E'', xs'') # EXS',h\ -\\ \(inline_call (Val v) E'', xs'')/EXS', h\" by(rule red1Return) simp hence "True,P,t \1 \(Val v, loc)/(E'', xs'') # EXS',h\ -\\ \(?e', ?xs')/EXS', h\" by simp moreover have "\Move1 P h ((Val v, loc), (E'', xs'') # EXS')" by auto ultimately have "\Red1 P t h ((Val v, loc), (E'', xs'') # EXS') ((?e', ?xs'), EXS')" by auto moreover from exec sees have [simp]: "ta = \" "h' = h" and [simp]: "frs' = (v # drop (length Ts + 1) stk'', loc'', C'', M'', pc'' + 1) # FRS'" by(auto simp add: compP2_def compMb2_def exec_1_iff) have "bisim1_list1 t h (?e', ?xs') EXS' None ((v # drop (length Ts + 1) stk'', loc'', C'', M'', pc'' + 1) # FRS')" proof from conf' show "compTP P \ t: (None, h, (v # drop (length Ts + 1) stk'', loc'', C'', M'', pc'' + 1) # FRS') \" by simp from sees'' show "P \ C'' sees M'': Ts''\T'' = \body''\ in D''" . from bisim1_inline_call_Val[OF bisim'' call'', of "length Ts" v] ins pc'' show "P,blocks1 0 (Class D''#Ts'') body'',h \ (inline_call (Val v) E'', xs'') \ (v # drop (length Ts + 1) stk'', loc'', pc'' + 1, None)" by simp from lenxs'' max_vars_inline_call[of "Val v" "E''"] show "max_vars (inline_call (Val v) E'') \ length xs''" by simp from bisims show "list_all2 (bisim1_fr P h) EXS' FRS'" by simp qed ultimately show ?thesis using Red by(auto simp del: \Red1_conv intro: rtranclp_into_tranclp1 rtranclp.rtrancl_into_rtrancl) qed qed next case [simp]: (Some a') from exec have execs: "(xcp', h', frs') = exception_step (compP2 P) a' h (stk, loc, C, M, pc) FRS" and [simp]: "ta = \" by(auto simp add: exec_1_iff) from conf have confxcp': "conf_xcp' P h xcp" unfolding correct_state_def by(auto simp add: compP2_def) then obtain D' where ha': "typeof_addr h a' = \Class_type D'\" and subclsD': "P \ D' \\<^sup>* Throwable" by auto from bisim have pc: "pc < length (compE2 body)" by(auto dest: bisim1_xcp_pcD) show ?thesis proof(cases "match_ex_table (compP2 P) (cname_of h a') pc (ex_table_of (compP2 P) C M)") case None from bisim have pc: "pc < length (compE2 body)" by(auto dest: bisim1_xcp_pcD) with sees' None have match: "match_ex_table (compP2 P) (cname_of h a') pc (compxE2 body 0 0) = None" by(auto) with execs sees' have [simp]: "ta = \" "xcp' = \a'\" "h' = h" "frs' = FRS" using match sees' by auto from conf obtain CCC where ha: "typeof_addr h a' = \Class_type CCC\" and subcls: "P \ CCC \\<^sup>* Throwable" unfolding correct_state_def by(auto simp add: conf_f_def2 compP2_def) from bisim1_xcp_\Red[OF ha subcls bisim[unfolded Some], of "\C M Ts T. compMb2"] match lenxs bsok have red: "\red1r P t h (e, xs) (Throw a', loc)" and b': "P,blocks1 0 (Class D#Ts) body,h \ (Throw a', loc) \ (stk, loc, pc, \a'\)" by(auto simp add: compP2_def bsok_def) from red have Red: "\Red1r P t h ((e, xs), exs) ((Throw a', loc), exs)" by(rule \red1r_into_\Red1r) show ?thesis proof(cases "FRS") case (Cons f' FRS') then obtain stk'' loc'' C'' M'' pc'' where [simp]: "FRS = (stk'', loc'', C'', M'', pc'') # FRS'" by(cases f') fastforce from bisims obtain e'' xs'' EXS' where [simp]: "exs = (e'', xs'') # EXS'" by(auto simp add: list_all2_Cons2) with bisims have "bisim1_fr P h (e'', xs'') (stk'', loc'', C'', M'', pc'')" by simp then obtain E'' Ts'' T'' body'' D'' a'' M''' vs'' where [simp]: "e'' = E''" and sees'': "P \ C'' sees M'':Ts''\T'' = \body''\ in D''" and bisim'': "P,blocks1 0 (Class D''#Ts'') body'',h \ (E'', xs'') \ (stk'', loc'', pc'', None)" and call'': "call1 E'' = \(a'', M''', vs'')\" and lenxs'': "max_vars E'' \ length xs''" by(cases) fastforce let ?ee' = "inline_call (Throw a') E''" let ?e' = "?ee'" let ?xs' = "xs''" from bisim'' call'' have pc'': "pc'' < length (compE2 (blocks1 0 (Class D''#Ts'') body''))" by(rule bisim1_call_pcD) hence pc'': "pc'' < length (compE2 body'')" by simp with sees_method_compP[OF sees'', where f="\C M Ts T. compMb2"] sees_method_compP[OF sees, where f="\C M Ts T. compMb2"] conf obtain ST LT where \: "compTP P C'' M'' ! pc'' = \(ST, LT)\" and conff: "conf_f (compP (\C M Ts T. compMb2) P) h (ST, LT) (compE2 body'' @ [Return]) (stk'', loc'', C'', M'', pc'')" and ins: "(compE2 body'' @ [Return]) ! pc'' = Invoke M (length Ts)" unfolding correct_state_def by(fastforce simp add: compP2_def compMb2_def dest: sees_method_fun) from bisim1_callD[OF bisim'' call'', of M "length Ts"] ins pc'' have [simp]: "M''' = M" by simp have "True,P,t \1 \(Throw a', loc)/(E'', xs'') # EXS',h\ -\\ \(inline_call (Throw a') E'', xs'')/EXS', h\" by(rule red1Return) simp moreover have "\Move1 P h ((Throw a', loc), (E'', xs'') # EXS')" by fastforce ultimately have "\Red1 P t h ((Throw a', loc), (E'', xs'') # EXS') ((?e', ?xs'), EXS')" by simp moreover have "bisim1_list1 t h (?e', ?xs') EXS' \a'\ ((stk'', loc'', C'', M'', pc'') # FRS')" proof from conf' show "compTP P \ t: (\a'\, h, (stk'', loc'', C'', M'', pc'') # FRS') \" by simp from sees'' show "P \ C'' sees M'': Ts''\T'' = \body''\ in D''" . from bisim1_inline_call_Throw[OF bisim'' call'', of "length Ts" a'] ins pc'' show "P,blocks1 0 (Class D''#Ts'') body'',h \ (inline_call (Throw a') E'', xs'') \ (stk'', loc'', pc'', \a'\)" by simp from lenxs'' max_vars_inline_call[of "Throw a'" "E''"] show "max_vars (inline_call (Throw a') E'') \ length xs''" by simp from bisims show "list_all2 (bisim1_fr P h) EXS' FRS'" by simp qed ultimately show ?thesis using Red by(auto simp del: \Red1_conv intro: rtranclp.rtrancl_into_rtrancl rtranclp_into_tranclp1) next case [simp]: Nil with bisims have [simp]: "exs = []" by simp from hconf have "bisim1_list1 t h (Throw a', loc) [] \a'\ []" by(rule bl1_finalThrow) thus ?thesis using Red by(auto simp del: \Red1_conv intro: rtranclp.rtrancl_into_rtrancl rtranclp_into_tranclp1 simp add: sim21_size_def) qed next case (Some pcd) then obtain pch d where match: "match_ex_table (compP2 P) (cname_of h a') pc (ex_table_of (compP2 P) C M) = \(pch, d)\" by(cases pcd) auto with \ sees' pc have \': "\move2 (compP2 P) h stk body pc \a'\" by(simp add: compP2_def compMb2_def \move2_iff) from match execs have [simp]: "h' = h" "xcp' = None" "frs' = (Addr a' # drop (length stk - d) stk, loc, C, M, pch) # FRS" by simp_all from bisim match sees' have "d \ length stk" by(auto intro: bisim1_match_Some_stk_length simp add: compP2_def compMb2_def) with match sees' have execm: "exec_move_d P t (blocks1 0 (Class D#Ts) body) h (stk, loc, pc, \a'\) ta h' (Addr a' # drop (length stk - d) stk, loc, pch, None)" by(auto simp add: exec_move_def exec_meth_xcpt) from conf obtain ST where "compP2 P,h \ stk [:\] ST" by(auto simp add: correct_state_def conf_f_def2) hence ST: "P,h \ stk [:\] ST" by(rule List.list_all2_mono)(simp add: compP2_def) from red1_simulates_exec_instr[OF wf hconf tconf bisim[unfolded \xcp = \a'\\] execm _ bsok ST] lenxs ha' subclsD' \' obtain e'' xs'' where b': "P,blocks1 0 (Class D#Ts) body,h \ (e'', xs'') \ (Addr a' # drop (length stk - d) stk, loc, pch, None)" and red: "(if pc < pch then \red1r else \red1t) P t h (e, xs) (e'', xs'')" and [simp]: "h' = h" by(auto split: if_split_asm intro: \move2xcp simp add: compP2_def simp del: blocks1.simps) have Red: "(if sim21_size (compP2 P) (xcp', frs') (xcp, frs) then \Red1r else \Red1t) P t h ((e, xs), exs) ((e'', xs''), exs)" proof(cases "pc < pch") case True from bisim b' have "pc \ Suc (length (compE2 body))" "pch \ Suc (length (compE2 body))" by(auto dest: bisim1_pc_length_compE2) with sees True have "sim21_size (compP2 P) (xcp', frs') (xcp, frs)" by(auto simp add: sim21_size_def)(auto simp add: compP2_def compMb2_def intro: sim21_size_aux.intros) with red True show ?thesis by simp(rule \red1r_into_\Red1r) next case False thus ?thesis using red by(auto intro: \red1t_into_\Red1t \red1r_into_\Red1r) qed moreover from red lenxs have "max_vars e'' \ length xs''" apply(auto dest: \red1r_max_vars \red1r_preserves_len \red1t_max_vars \red1t_preserves_len split: if_split_asm) apply(frule \red1r_max_vars \red1t_max_vars, drule \red1r_preserves_len \red1t_preserves_len, simp)+ done with conf' sees b' have "bisim1_list1 t h (e'', xs'') exs None ((Addr a' # drop (length stk - d) stk, loc, C, M, pch) # FRS)" using bisims unfolding \h' = h\ \xcp' = None\ \frs' = (Addr a' # drop (length stk - d) stk, loc, C, M, pch) # FRS\ by rule ultimately show ?thesis by(auto split del: if_split) qed qed qed(insert exec, auto simp add: exec_1_iff elim!: jvmd_NormalE) lemma \Red1_simulates_exec_1_not_\: assumes wf: "wf_J1_prog P" and exec: "exec_1_d (compP2 P) t (Normal (xcp, h, frs)) ta (Normal (xcp', h', frs'))" and bisim: "bisim1_list1 t h (e, xs) exs xcp frs" and \: "\ \Move2 (compP2 P) (xcp, h, frs)" shows "\e' xs' exs' ta' e'' xs'' exs''. \Red1r P t h ((e, xs), exs) ((e', xs'), exs') \ True,P,t \1 \(e', xs')/exs', h\ -ta'\ \(e'', xs'')/exs'', h'\ \ \ \Move1 P h ((e', xs'), exs') \ ta_bisim wbisim1 ta' ta \ bisim1_list1 t h' (e'', xs'') exs'' xcp' frs' \ (call1 e = None \ (case frs of Nil \ False | (stk, loc, C, M, pc) # FRS \ \M' n. instrs_of (compP2 P) C M ! pc \ Invoke M' n) \ e' = e \ xs' = xs \ exs' = exs) " using bisim proof cases case (bl1_Normal stk loc C M pc FRS Ts T body D) hence [simp]: "frs = (stk, loc, C, M, pc) # FRS" and conf: "compTP P \ t: (xcp, h, (stk, loc, C, M, pc) # FRS) \" and sees: "P \ C sees M: Ts\T = \body\ in D" and bisim: "P,blocks1 0 (Class D#Ts) body,h \ (e, xs) \ (stk, loc, pc, xcp)" and lenxs: "max_vars e \ length xs" and bisims: "list_all2 (bisim1_fr P h) exs FRS" by auto from sees_method_compP[OF sees, where f="\C M Ts T. compMb2"] have sees': "compP2 P \ C sees M: Ts\T = \(max_stack body, max_vars body, compE2 body @ [Return], compxE2 body 0 0)\ in D" by(simp add: compP2_def compMb2_def) from bisim have pc: "pc \ length (compE2 body)" by(auto dest: bisim1_pc_length_compE2) from conf have hconf: "hconf h" "preallocated h" and tconf: "P,h \ t \t" unfolding correct_state_def by(simp_all add: compP2_def tconf_def) from sees wf have bsok: "bsok (blocks1 0 (Class D # Ts) body) 0" by(auto dest!: sees_wf_mdecl simp add: bsok_def wf_mdecl_def WT1_expr_locks) from exec obtain check: "check (compP2 P) (xcp, h, frs)" and exec: "compP2 P,t \ (xcp, h, frs) -ta-jvm\ (xcp', h', frs')" by(rule jvmd_NormalE)(auto simp add: exec_1_iff) from wt_compTP_compP2[OF wf] exec conf have conf': "compTP P \ t: (xcp', h', frs') \" by(auto intro: BV_correct_1) show ?thesis proof(cases xcp) case [simp]: None from exec have execi: "(ta, xcp', h', frs') \ exec_instr (instrs_of (compP2 P) C M ! pc) (compP2 P) t h stk loc C M pc FRS" by(simp add: exec_1_iff) show ?thesis proof(cases "length frs' = Suc (length FRS)") case True with pc execi sees' have pc: "pc < length (compE2 body)" by(auto split: if_split_asm simp add: split_beta) with execi sees' have execi: "(ta, xcp', h', frs') \ exec_instr (compE2 body ! pc) (compP2 P) t h stk loc C M pc FRS" by(simp) from \ sees' True pc have \i: "\ \move2 (compP2 P) h stk body pc None" by(simp add: \move2_iff) from execi True sees' pc have "\stk' loc' pc'. frs' = (stk', loc', C, M, pc') # FRS" by(cases "(compE2 body @ [Return]) ! pc")(auto split: if_split_asm sum.split_asm simp add: split_beta, auto split: extCallRet.splits) then obtain stk' loc' pc' where [simp]: "frs' = (stk', loc', C, M, pc') # FRS" by blast from conf obtain ST where "compP2 P,h \ stk [:\] ST" by(auto simp add: correct_state_def conf_f_def2) hence ST: "P,h \ stk [:\] ST" by(rule List.list_all2_mono)(simp add: compP2_def) from execi sees True check pc have exec': "exec_move_d P t (blocks1 0 (Class D#Ts) body) h (stk, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" apply(auto simp add: compP2_def compMb2_def exec_move_def check_def exec_meth_instr split: if_split_asm sum.split_asm) apply(cases "compE2 body ! pc") apply(auto simp add: neq_Nil_conv split_beta split: if_split_asm sum.split_asm) apply(force split: extCallRet.split_asm) apply(cases "compE2 body ! pc", auto simp add: split_beta neq_Nil_conv split: if_split_asm sum.split_asm) done from red1_simulates_exec_instr[OF wf hconf tconf bisim this _ bsok ST] lenxs \i obtain e'' xs'' ta' e' xs' where bisim': "P,blocks1 0 (Class D#Ts) body,h' \ (e'', xs'') \ (stk', loc', pc', xcp')" and red1: "\red1r P t h (e, xs) (e', xs')" and red2: "True,P,t \1 \e',(h, xs')\ -ta'\ \e'',(h', xs'')\" and \1: "\ \move1 P h e'" and tabisim: "ta_bisim wbisim1 (extTA2J1 P ta') ta" and call: "call1 e = None \ no_call2 (blocks1 0 (Class D # Ts) body) pc \ e' = e \ xs' = xs" by(fastforce simp del: blocks1.simps) from red1 have Red1: "\Red1r P t h ((e, xs), exs) ((e', xs'), exs)" by(rule \red1r_into_\Red1r) moreover from red2 have "True,P,t \1 \(e', xs')/exs, h\ -extTA2J1 P ta'\ \(e'', xs'')/exs, h'\" by(rule red1Red) moreover from \1 red2 have "\ \Move1 P h ((e', xs'), exs)" by auto moreover from \red1r_max_vars[OF red1] lenxs \red1r_preserves_len[OF red1] have "max_vars e' \ length xs'" by simp with red1_preserves_len[OF red2] red1_max_vars[OF red2] have "max_vars e'' \ length xs''" by simp with conf' sees bisim' have "bisim1_list1 t h' (e'', xs'') exs xcp' ((stk', loc', C, M, pc') # FRS)" unfolding \frs' = (stk', loc', C, M, pc') # FRS\ proof from red2 have "hext h h'" by(auto dest: red1_hext_incr) from bisims show "list_all2 (bisim1_fr P h') exs FRS" by(rule List.list_all2_mono)(erule bisim1_fr_hext_mono[OF _ \hext h h'\]) qed moreover from call sees' have "call1 e = None \ (\M' n. instrs_of (compP2 P) C M ! pc \ Invoke M' n) \ e' = e \ xs' = xs" by(auto simp add: no_call2_def) ultimately show ?thesis using tabisim by(fastforce simp del: split_paired_Ex) next case False with execi sees pc compE2_not_Return[of body] have "(pc = length (compE2 body) \ (\M n. compE2 body ! pc = Invoke M n)) \ xcp' = None" apply(cases "compE2 body ! pc") apply(auto split: if_split_asm sum.split_asm simp add: split_beta compP2_def compMb2_def) apply(auto split: extCallRet.splits) apply(metis in_set_conv_nth)+ done hence [simp]: "xcp' = None" and "pc = length (compE2 body) \ (\M n. compE2 body ! pc = Invoke M n)" by simp_all moreover { assume [simp]: "pc = length (compE2 body)" with sees_method_compP[OF sees, where f="\C M Ts T. compMb2"] \ have False by(auto simp add: compMb2_def compP2_def) hence ?thesis .. } moreover { assume "\M n. compE2 body ! pc = Invoke M n" and "pc \ length (compE2 body)" with pc obtain MM n where ins: "compE2 body ! pc = Invoke MM n" and pc: "pc < length (compE2 body)" by auto with bisim1_Invoke_stkD[OF bisim[unfolded None], of MM n] obtain vs' v' stk' where [simp]: "stk = vs' @ v' # stk'" "n = length vs'" by auto with False \ sees' execi pc ins have False by auto (auto split: extCallRet.split_asm) } ultimately show ?thesis by blast qed next case [simp]: (Some ad) from bisim have pc: "pc < length (compE2 body)" by(auto dest: bisim1_xcp_pcD) with \ sees' have False by auto thus ?thesis .. qed qed(insert exec, auto simp add: exec_1_iff elim!: jvmd_NormalE) end end diff --git a/thys/LinearQuantifierElim/Thys/FRE.thy b/thys/LinearQuantifierElim/Thys/FRE.thy --- a/thys/LinearQuantifierElim/Thys/FRE.thy +++ b/thys/LinearQuantifierElim/Thys/FRE.thy @@ -1,205 +1,205 @@ (* Author: Tobias Nipkow, 2007 *) theory FRE imports LinArith begin subsection\Ferrante-Rackoff \label{sec:FRE}\ text\This section formalizes a slight variant of Ferrante and Rackoff's algorithm~\cite{FerranteR-SIAM75}. We consider equalities separately, which improves performance.\ fun between :: "real * real list \ real * real list \ real * real list" where "between (r,cs) (s,ds) = ((r+s)/2, (1/2) *\<^sub>s (cs+ds))" definition FR\<^sub>1 :: "atom fm \ atom fm" where "FR\<^sub>1 \ = (let as = R.atoms\<^sub>0 \; lbs = lbounds as; ubs = ubounds as; ebs = ebounds as; intrs = [subst \ (between l u) . l \ lbs, u \ ubs] in list_disj (inf\<^sub>- \ # inf\<^sub>+ \ # intrs @ map (subst \) ebs))" lemma dense_interval: assumes "finite L" "finite U" "l \ L" "u \ U" "l < x" "x < u" "P(x::real)" and dense: "\y l u. \ \y\{l<.. L; \y\{x<.. U; l \ P y" shows "\l\L.\u\U. l (\y. l y P y)" proof - let ?L = "{l:L. l < x}" let ?U = "{u:U. x < u}" let ?ll = "Max ?L" let ?uu = "Min ?U" have "?L \ {}" using \l \ L\ \l by (blast intro:order_less_imp_le) moreover have "?U \ {}" using \u:U\ \x by (blast intro:order_less_imp_le) ultimately have "\y. ?ll y y \ L" "\y. x y y \ U" using \finite L\ \finite U\ by force+ moreover have "?ll \ L" proof show "?ll \ ?L" using \finite L\ Max_in[OF _ \?L \ {}\] by simp show "?L \ L" by blast qed moreover have "?uu \ U" proof show "?uu \ ?U" using \finite U\ Min_in[OF _ \?U \ {}\] by simp show "?U \ U" by blast qed moreover have "?ll < x" using \finite L\ \?L \ {}\ by simp moreover have "x < ?uu" using \finite U\ \?U \ {}\ by simp moreover have "?ll < ?uu" using \?ll \x by arith ultimately show ?thesis using \l < x\ \x < u\ \?L \ {}\ \?U \ {}\ by(blast intro!:dense greaterThanLessThan_iff[THEN iffD1]) qed lemma dense: "\ nqfree f; \y\{l<.. LB f xs; \y\{x<.. UB f xs; l < x; x < u; x \ EQ f xs; R.I f (x#xs); l < y; y < u\ \ R.I f (y#xs)" proof(induct f) case (Atom a) show ?case proof (cases a) case (Less r cs) show ?thesis proof (cases cs) case Nil thus ?thesis using Atom Less by (simp add:depends\<^sub>R_def) next case (Cons c cs) hence "r < c*x + \cs,xs\" using Atom Less by simp { assume "c=0" hence ?thesis using Atom Less Cons by simp } moreover { assume "c<0" hence "x < (r - \cs,xs\)/c" (is "_ < ?u") using \r < c*x + \cs,xs\\ by (simp add: field_simps) have ?thesis proof (rule ccontr) assume "\ R.I (Atom a) (y#xs)" hence "?u \ y" using Atom Less Cons \c<0\ by (auto simp add: field_simps) hence "?u < u" using \y by simp with \x show False using Atom Less Cons \c<0\ by(auto simp:depends\<^sub>R_def) qed } moreover { assume "c>0" hence "x > (r - \cs,xs\)/c" (is "_ > ?l") using \r < c*x + \cs,xs\\ by (simp add: field_simps) have ?thesis proof (rule ccontr) assume "\ R.I (Atom a) (y#xs)" hence "?l \ y" using Atom Less Cons \c>0\ by (auto simp add: field_simps) hence "?l > l" using \y>l\ by simp with \?l show False using Atom Less Cons \c>0\ by (auto simp:depends\<^sub>R_def) qed } ultimately show ?thesis by force qed next case (Eq r cs) show ?thesis proof (cases cs) case Nil thus ?thesis using Atom Eq by (simp add:depends\<^sub>R_def) next case (Cons c cs) hence "r = c*x + \cs,xs\" using Atom Eq by simp { assume "c=0" hence ?thesis using Atom Eq Cons by simp } moreover { assume "c\0" hence ?thesis using \r = c*x + \cs,xs\\ Atom Eq Cons \l \y - by(auto simp: ac_simps depends\<^sub>R_def split:if_splits) } + by (auto simp: depends\<^sub>R_def split: if_splits) } ultimately show ?thesis by force qed qed next case (And f1 f2) thus ?case by auto (metis (no_types, hide_lams))+ next case (Or f1 f2) thus ?case by auto (metis (no_types, hide_lams))+ qed fastforce+ theorem I_FR\<^sub>1: assumes "nqfree \" shows "R.I (FR\<^sub>1 \) xs = (\x. R.I \ (x#xs))" (is "?FR = ?EX") proof assume ?FR { assume "R.I (inf\<^sub>- \) xs" hence ?EX using \?FR\ min_inf[OF \nqfree \\, where xs=xs] by(auto simp add:FR\<^sub>1_def) } moreover { assume "R.I (inf\<^sub>+ \) xs" hence ?EX using \?FR\ plus_inf[OF \nqfree \\, where xs=xs] by(auto simp add:FR\<^sub>1_def) } moreover { assume "\x \ EQ \ xs. R.I \ (x#xs)" hence ?EX using \?FR\ by(auto simp add:FR\<^sub>1_def) } moreover { assume "\R.I (inf\<^sub>- \) xs \ \R.I (inf\<^sub>+ \) xs \ (\x\EQ \ xs. \R.I \ (x#xs))" with \?FR\ obtain r cs s ds where "R.I (subst \ (between (r,cs) (s,ds))) xs" by(auto simp: FR\<^sub>1_def eval_def diff_divide_distrib set_ebounds I_subst \nqfree \\) blast hence "R.I \ (eval (between (r,cs) (s,ds)) xs # xs)" by(simp add:I_subst \nqfree \\) hence ?EX .. } ultimately show ?EX by blast next assume ?EX then obtain x where x: "R.I \ (x#xs)" .. { assume "R.I (inf\<^sub>- \) xs \ R.I (inf\<^sub>+ \) xs" hence ?FR by(auto simp:FR\<^sub>1_def) } moreover { assume "x \ EQ \ xs" then obtain r cs where "(r,cs) \ set(ebounds(R.atoms\<^sub>0 \)) \ x = r + \cs,xs\" by(force simp:set_ebounds field_simps) moreover hence "R.I (subst \ (r,cs)) xs" using x by(auto simp: I_subst \nqfree \\ eval_def) ultimately have ?FR by(force simp:FR\<^sub>1_def) } moreover { assume "\ R.I (inf\<^sub>- \) xs" and "\ R.I (inf\<^sub>+ \) xs" and "x \ EQ \ xs" obtain l where "l \ LB \ xs" "l < x" using LBex[OF \nqfree \\ x \\ R.I (inf\<^sub>- \) xs\ \x \ EQ \ xs\] .. obtain u where "u \ UB \ xs" "x < u" using UBex[OF \nqfree \\ x \\ R.I (inf\<^sub>+ \) xs\ \x \ EQ \ xs\] .. have "\l\LB \ xs. \u\UB \ xs. l (\y. l < y \ y < u \ R.I \ (y#xs))" using dense_interval[where P = "\x. R.I \ (x#xs)", OF finite_LB finite_UB \l:LB \ xs\ \u:UB \ xs\ \l \x x] x dense[OF \nqfree \\ _ _ _ _ \x \ EQ \ xs\] by simp then obtain r c cs s d ds where "Less r (c # cs) \ set (R.atoms\<^sub>0 \)" "Less s (d # ds) \ set (R.atoms\<^sub>0 \)" "\y. (r - \cs,xs\) / c < y \ y < (s - \ds,xs\) / d \ R.I \ (y # xs)" and *: "c > 0" "d < 0" "(r - \cs,xs\) / c < (s - \ds,xs\) / d" by blast moreover have "(r - \cs,xs\) / c < eval (between (r / c, (-1 / c) *\<^sub>s cs) (s / d, (-1 / d) *\<^sub>s ds)) xs" (is ?P) and "eval (between (r / c, (-1 / c) *\<^sub>s cs) (s / d, (-1 / d) *\<^sub>s ds)) xs < (s - \ds,xs\) / d" (is ?Q) proof - from * have [simp]: "c * (c * (d * (d * 4))) > 0" by (simp add: algebra_split_simps) from * have "c * s + d * \cs,xs\ < d * r + c * \ds,xs\" by (simp add: field_simps) with * have "(2 * c * c * d) * (d * r + c * \ds,xs\) < (2 * c * c * d) * (c * s + d * \cs,xs\)" and "(2 * c * d * d) * (c * s + d * \cs,xs\) < (2 * c * d * d) * (d * r + c * \ds,xs\)" by simp_all with * show ?P and ?Q by (auto simp add: field_simps eval_def iprod_left_add_distrib) qed ultimately have ?FR by (fastforce simp: FR\<^sub>1_def bex_Un set_lbounds set_ubounds set_ebounds I_subst \nqfree \\) } ultimately show ?FR by blast qed definition "FR = R.lift_nnf_qe FR\<^sub>1" lemma qfree_FR\<^sub>1: "nqfree \ \ qfree (FR\<^sub>1 \)" apply(simp add:FR\<^sub>1_def) apply(rule qfree_list_disj) apply(auto simp:qfree_min_inf qfree_plus_inf set_ubounds set_lbounds set_ebounds image_def qfree_map_fm) done theorem I_FR: "R.I (FR \) xs = R.I \ xs" by(simp add:I_FR\<^sub>1 FR_def R.I_lift_nnf_qe qfree_FR\<^sub>1) theorem qfree_FR: "qfree (FR \)" by(simp add:FR_def R.qfree_lift_nnf_qe qfree_FR\<^sub>1) end diff --git a/thys/LinearQuantifierElim/Thys/QElin_inf.thy b/thys/LinearQuantifierElim/Thys/QElin_inf.thy --- a/thys/LinearQuantifierElim/Thys/QElin_inf.thy +++ b/thys/LinearQuantifierElim/Thys/QElin_inf.thy @@ -1,302 +1,302 @@ (* Author: Tobias Nipkow, 2007 *) theory QElin_inf imports LinArith begin subsection \Quantifier elimination with infinitesimals \label{sec:lin-inf}\ text\This section formalizes Loos and Weispfenning's quantifier elimination procedure based on (the simulation of) infinitesimals~\cite{LoosW93}.\ fun asubst_peps :: "real * real list \ atom \ atom fm" ("asubst\<^sub>+") where "asubst_peps (r,cs) (Less s (d#ds)) = (if d=0 then Atom(Less s ds) else let u = s - d*r; v = d *\<^sub>s cs + ds; less = Atom(Less u v) in if d<0 then less else Or less (Atom(Eq u v)))" | "asubst_peps rcs (Eq r (d#ds)) = (if d=0 then Atom(Eq r ds) else FalseF)" | "asubst_peps rcs a = Atom a" abbreviation subst_peps :: "atom fm \ real * real list \ atom fm" ("subst\<^sub>+") where "subst\<^sub>+ \ rcs \ amap\<^sub>f\<^sub>m (asubst\<^sub>+ rcs) \" definition "nolb f xs l x = (\y\{l<.. LB f xs)" lemma nolb_And[simp]: "nolb (And f g) xs l x = (nolb f xs l x \ nolb g xs l x)" apply(clarsimp simp:nolb_def) apply blast done lemma nolb_Or[simp]: "nolb (Or f g) xs l x = (nolb f xs l x \ nolb g xs l x)" apply(clarsimp simp:nolb_def) apply blast done context notes [[simp_depth_limit=4]] begin lemma innermost_intvl: "\ nqfree f; nolb f xs l x; l < x; x \ EQ f xs; R.I f (x#xs); l < y; y \ x\ \ R.I f (y#xs)" proof(induct f) case (Atom a) show ?case proof (cases a) case [simp]: (Less r cs) show ?thesis proof (cases cs) case Nil thus ?thesis using Atom by (simp add:depends\<^sub>R_def) next case [simp]: (Cons c cs) hence "r < c*x + \cs,xs\" using Atom by simp { assume "c=0" hence ?thesis using Atom by simp } moreover { assume "c<0" hence "x < (r - \cs,xs\)/c" (is "_ < ?u") using \r < c*x + \cs,xs\\ by (simp add: field_simps) have ?thesis proof (rule ccontr) assume "\ R.I (Atom a) (y#xs)" hence "?u \ y" using Atom \c<0\ by (auto simp add: field_simps) with \x show False using Atom \c<0\ by(auto simp:depends\<^sub>R_def) qed } moreover { assume "c>0" hence "x > (r - \cs,xs\)/c" (is "_ > ?l") using \r < c*x + \cs,xs\\ by (simp add: field_simps) then have "?l < y" using Atom \c>0\ by (auto simp:depends\<^sub>R_def Ball_def nolb_def) (metis linorder_not_le antisym order_less_trans) hence ?thesis using \c>0\ by (simp add: field_simps) } ultimately show ?thesis by force qed next case [simp]: (Eq r cs) show ?thesis proof (cases cs) case Nil thus ?thesis using Atom by (simp add:depends\<^sub>R_def) next case [simp]: (Cons c cs) hence "r = c*x + \cs,xs\" using Atom by simp { assume "c=0" hence ?thesis using Atom by simp } moreover { assume "c\0" hence ?thesis using \r = c*x + \cs,xs\\ Atom - by(auto simp: ac_simps depends\<^sub>R_def split:if_splits) } + by (auto simp: depends\<^sub>R_def split: if_splits) } ultimately show ?thesis by force qed qed next case (And f1 f2) thus ?case by (fastforce) next case (Or f1 f2) thus ?case by (fastforce) qed simp+ definition "EQ2 = EQ" lemma EQ2_Or[simp]: "EQ2 (Or f g) xs = (EQ2 f xs \ EQ2 g xs)" by(auto simp:EQ2_def) lemma EQ2_And[simp]: "EQ2 (And f g) xs = (EQ2 f xs \ EQ2 g xs)" by(auto simp:EQ2_def) lemma innermost_intvl2: "\ nqfree f; nolb f xs l x; l < x; x \ EQ2 f xs; R.I f (x#xs); l < y; y \ x\ \ R.I f (y#xs)" unfolding EQ2_def by(blast intro:innermost_intvl) lemma I_subst_peps2: "nqfree f \ r+\cs,xs\ < x \ nolb f xs (r+\cs,xs\) x \ \y \ {r+\cs,xs\ <.. x}. R.I f (y#xs) \ y \ EQ2 f xs \ R.I (subst\<^sub>+ f (r,cs)) xs" proof(induct f) case FalseF thus ?case by simp (metis antisym_conv1 linorder_neq_iff) next case (Atom a) show ?case proof(cases "((r,cs),a)" rule:asubst_peps.cases) case (1 r cs s d ds) { assume "d=0" hence ?thesis using Atom 1 by auto } moreover { assume "d<0" have "s < d*x + \ds,xs\" using Atom 1 by simp moreover have "d*x < d*(r + \cs,xs\)" using \d<0\ Atom 1 by (simp add: mult_strict_left_mono_neg) ultimately have "s < d * (r + \cs,xs\) + \ds,xs\" by(simp add:algebra_simps) hence ?thesis using 1 by (auto simp add: iprod_left_add_distrib algebra_simps) } moreover { let ?L = "(s - \ds,xs\) / d" let ?U = "r + \cs,xs\" assume "d>0" hence "?U < x" and "\y. ?U < y \ y < x \ y \ ?L" and "\y. ?U < y \ y \ x \ ?L < y" using Atom 1 by(simp_all add:nolb_def depends\<^sub>R_def Ball_def field_simps) hence "?L < ?U \ ?L = ?U" by (metis linorder_neqE_linordered_idom order_refl) hence ?thesis using Atom 1 \d>0\ by (simp add: iprod_left_add_distrib field_simps) } ultimately show ?thesis by force next case 2 thus ?thesis using Atom by (fastforce simp: nolb_def EQ2_def depends\<^sub>R_def field_simps split: if_split_asm) qed (insert Atom, auto) next case Or thus ?case by(simp add:Ball_def)(metis order_refl innermost_intvl2) qed simp_all end lemma I_subst_peps: "nqfree f \ R.I (subst\<^sub>+ f (r,cs)) xs \ (\leps>r+\cs,xs\. \x. r+\cs,xs\ < x \ x \ leps \ R.I f (x#xs))" proof(induct f) case TrueF thus ?case by simp (metis less_add_one) next case (Atom a) show ?case proof (cases "((r,cs),a)" rule: asubst_peps.cases) case (1 r cs s d ds) { assume "d=0" hence ?thesis using Atom 1 by auto (metis less_add_one) } moreover { assume "d<0" with Atom 1 have "r + \cs,xs\ < (s - \ds,xs\)/d" (is "?a < ?b") by(simp add:field_simps iprod_left_add_distrib) then obtain x where "?a < x" "x < ?b" by(metis dense) hence " \y. ?a < y \ y \ x \ s < d*y + \ds,xs\" using \d<0\ by (simp add:field_simps) (metis add_le_cancel_right mult_le_cancel_left order_antisym linear mult.commute xt1(8)) hence ?thesis using 1 \?a by auto } moreover { let ?a = "s - d * r" let ?b = "\d *\<^sub>s cs + ds,xs\" assume "d>0" with Atom 1 have "?a < ?b \ ?a = ?b" by auto hence ?thesis proof assume "?a = ?b" thus ?thesis using \d>0\ Atom 1 by(simp add:field_simps iprod_left_add_distrib) (metis add_0_left add_less_cancel_right distrib_left mult.commute mult_strict_left_mono) next assume "?a < ?b" { fix x assume "r+\cs,xs\ < x \ x \ r+\cs,xs\ + 1" hence "d*(r + \cs,xs\) < d*x" using \d>0\ by(metis mult_strict_left_mono) hence "s < d*x + \ds,xs\" using \d>0\ \?a < ?b\ by (simp add:algebra_simps iprod_left_add_distrib) } thus ?thesis using 1 \d>0\ by(force simp: iprod_left_add_distrib) qed } ultimately show ?thesis by (metis less_linear) qed (insert Atom, auto split:if_split_asm intro: less_add_one) next case And thus ?case apply clarsimp apply(rule_tac x="min leps lepsa" in exI) apply simp done next case Or thus ?case by force qed simp_all lemma dense_interval: assumes "finite L" "l \ L" "l < x" "P(x::real)" and dense: "\y l. \ \y\{l<.. L; lx \ \ P y" shows "\l\L. l (\y\{l<.. L) \ (\y. l y\x \ P y)" proof - let ?L = "{l\L. l < x}" let ?ll = "Max ?L" have "?L \ {}" using \l \ L\ \l by (blast intro:order_less_imp_le) hence "\y. ?ll y y \ L" using \finite L\ by force moreover have "?ll \ L" proof show "?ll \ ?L" using \finite L\ Max_in[OF _ \?L \ {}\] by simp show "?L \ L" by blast qed moreover have "?ll < x" using \finite L\ \?L \ {}\ by simp ultimately show ?thesis using \l < x\ \?L \ {}\ by(blast intro!:dense greaterThanLessThan_iff[THEN iffD1]) qed definition "qe_eps\<^sub>1(f) = (let as = R.atoms\<^sub>0 f; lbs = lbounds as; ebs = ebounds as in list_disj (inf\<^sub>- f # map (subst\<^sub>+ f) lbs @ map (subst f) ebs))" theorem I_eps1: assumes "nqfree f" shows "R.I (qe_eps\<^sub>1 f) xs = (\x. R.I f (x#xs))" (is "?QE = ?EX") proof let ?as = "R.atoms\<^sub>0 f" let ?ebs = "ebounds ?as" assume ?QE { assume "R.I (inf\<^sub>- f) xs" hence ?EX using \?QE\ min_inf[of f xs] \nqfree f\ by(auto simp add:qe_eps\<^sub>1_def amap_fm_list_disj) } moreover { assume "\x \ EQ f xs. \R.I f (x#xs)" "\ R.I (inf\<^sub>- f) xs" with \?QE\ \nqfree f\ obtain r cs where "R.I (subst\<^sub>+ f (r,cs)) xs" by (fastforce simp: qe_eps\<^sub>1_def set_ebounds diff_divide_distrib eval_def I_subst \nqfree f\) then obtain leps where "R.I f (leps#xs)" using I_subst_peps[OF \nqfree f\] by fastforce hence ?EX .. } ultimately show ?EX by blast next let ?as = "R.atoms\<^sub>0 f" let ?ebs = "ebounds ?as" assume ?EX then obtain x where x: "R.I f (x#xs)" .. { assume "R.I (inf\<^sub>- f) xs" hence ?QE using \nqfree f\ by(auto simp:qe_eps\<^sub>1_def) } moreover { assume "\rcs \ set ?ebs. R.I (subst f rcs) xs" hence ?QE by(auto simp:qe_eps\<^sub>1_def) } moreover { assume "\ R.I (inf\<^sub>- f) xs" and "\rcs \ set ?ebs. \ R.I (subst f rcs) xs" hence noE: "\e \ EQ f xs. \ R.I f (e#xs)" using \nqfree f\ by (force simp:set_ebounds I_subst diff_divide_distrib eval_def split:if_split_asm) hence "x \ EQ f xs" using x by fastforce obtain l where "l \ LB f xs" "l < x" using LBex[OF \nqfree f\ x \\ R.I(inf\<^sub>- f) xs\ \x \ EQ f xs\] .. have "\l\LB f xs. l nolb f xs l x \ (\y. l < y \ y \ x \ R.I f (y#xs))" using dense_interval[where P = "\x. R.I f (x#xs)", OF finite_LB \l\LB f xs\ \l x] x innermost_intvl[OF \nqfree f\ _ _ \x \ EQ f xs\] by (simp add:nolb_def) then obtain r c cs where *: "Less r (c#cs) \ set(R.atoms\<^sub>0 f) \ c>0 \ (r - \cs,xs\)/c < x \ nolb f xs ((r - \cs,xs\)/c) x \ (\y. (r - \cs,xs\)/c < y \ y \ x \ R.I f (y#xs))" by blast then have "R.I (subst\<^sub>+ f (r/c, (-1/c) *\<^sub>s cs)) xs" using noE by(auto intro!: I_subst_peps2[OF \nqfree f\] simp:EQ2_def diff_divide_distrib algebra_simps) with * have ?QE by(simp add:qe_eps\<^sub>1_def bex_Un set_lbounds) metis } ultimately show ?QE by blast qed lemma qfree_asubst_peps: "qfree (asubst\<^sub>+ rcs a)" by(cases "(rcs,a)" rule:asubst_peps.cases) simp_all lemma qfree_subst_peps: "nqfree \ \ qfree (subst\<^sub>+ \ rcs)" by(induct \) (simp_all add:qfree_asubst_peps) lemma qfree_qe_eps\<^sub>1: "nqfree \ \ qfree(qe_eps\<^sub>1 \)" apply(simp add:qe_eps\<^sub>1_def) apply(rule qfree_list_disj) apply (auto simp:qfree_min_inf qfree_subst_peps qfree_map_fm) done definition "qe_eps = R.lift_nnf_qe qe_eps\<^sub>1" lemma qfree_qe_eps: "qfree(qe_eps \)" by(simp add: qe_eps_def R.qfree_lift_nnf_qe qfree_qe_eps\<^sub>1) lemma I_qe_eps: "R.I (qe_eps \) xs = R.I \ xs" by(simp add:qe_eps_def R.I_lift_nnf_qe qfree_qe_eps\<^sub>1 I_eps1) end