diff --git a/thys/JinjaThreads/Basic/Auxiliary.thy b/thys/JinjaThreads/Basic/Auxiliary.thy --- a/thys/JinjaThreads/Basic/Auxiliary.thy +++ b/thys/JinjaThreads/Basic/Auxiliary.thy @@ -1,864 +1,803 @@ (* Title: JinjaThreads/Basic/Aux.thy Author: Andreas Lochbihler, David von Oheimb, Tobias Nipkow Based on the Jinja theory Common/Aux.thy by David von Oheimb and Tobias Nipkow *) section \Auxiliary Definitions and Lemmata\ theory Auxiliary imports Complex_Main "HOL-Library.Transitive_Closure_Table" "HOL-Library.Predicate_Compile_Alternative_Defs" "HOL-Library.Monad_Syntax" "HOL-Library.Infinite_Set" FinFun.FinFun begin unbundle finfun_syntax (* FIXME move and possibly turn into a general simproc *) lemma nat_add_max_le[simp]: "((n::nat) + max i j \ m) = (n + i \ m \ n + j \ m)" (*<*)by arith(*>*) lemma Suc_add_max_le[simp]: "(Suc(n + max i j) \ m) = (Suc(n + i) \ m \ Suc(n + j) \ m)" (*<*)by arith(*>*) lemma less_min_eq1: "(a :: 'a :: order) < b \ min a b = a" by(auto simp add: min_def order_less_imp_le) lemma less_min_eq2: "(a :: 'a :: order) > b \ min a b = b" by(auto simp add: min_def order_less_imp_le) no_notation floor ("\_\") notation Some ("(\_\)") (*<*) declare option.splits[split] Let_def[simp] subset_insertI2 [simp] (*>*) declare not_Cons_self [no_atp] lemma Option_bind_eq_None_conv: "x \ y = None \ x = None \ (\x'. x = Some x' \ y x' = None)" by(cases x) simp_all -lemma Option_bind_eq_Some_conv: - "x \ y = Some z \ (\x'. x = Some x' \ y x' = Some z)" -by(cases x) simp_all - lemma map_upds_xchg_snd: "\ length xs \ length ys; length xs \ length zs; \i. i < length xs \ ys ! i = zs ! i \ \ f(xs [\] ys) = f(xs [\] zs)" proof(induct xs arbitrary: ys zs f) case Nil thus ?case by simp next case (Cons x xs) note IH = \\f ys zs. \ length xs \ length ys; length xs \ length zs; \i \ f(xs [\] ys) = f(xs [\] zs)\ note leny = \length (x # xs) \ length ys\ note lenz = \length (x # xs) \ length zs\ note nth = \\i from lenz obtain z zs' where zs [simp]: "zs = z # zs'" by(cases zs, auto) from leny obtain y ys' where ys [simp]: "ys = y # ys'" by(cases ys, auto) from lenz leny nth have "(f(x \ y))(xs [\] ys') = (f(x \ y))(xs [\] zs')" by-(rule IH, auto) moreover from nth have "y = z" by auto ultimately show ?case by(simp add: map_upds_def) qed subsection \\distinct_fst\\ definition distinct_fst :: "('a \ 'b) list \ bool" where "distinct_fst \ distinct \ map fst" lemma distinct_fst_Nil [simp]: "distinct_fst []" (*<*) apply (unfold distinct_fst_def) apply (simp (no_asm)) done (*>*) lemma distinct_fst_Cons [simp]: "distinct_fst ((k,x)#kxs) = (distinct_fst kxs \ (\y. (k,y) \ set kxs))" (*<*) apply (unfold distinct_fst_def) apply (auto simp:image_def) done (*>*) lemma distinct_fstD: "\ distinct_fst xs; (x, y) \ set xs; (x, z) \ set xs \ \ y = z" by(induct xs) auto lemma map_of_SomeI: "\ distinct_fst kxs; (k,x) \ set kxs \ \ map_of kxs k = Some x" -(*<*)by (induct kxs) (auto simp:fun_upd_apply)(*>*) - -lemma rel_option_Some1: - "rel_option R (Some x) y \ (\y'. y = Some y' \ R x y')" -by(cases y) simp_all - -lemma rel_option_Some2: - "rel_option R x (Some y) \ (\x'. x = Some x' \ R x' y)" -by(cases x) simp_all +(*<*)by (induct kxs) (auto)(*>*) subsection \Using @{term list_all2} for relations\ definition fun_of :: "('a \ 'b) set \ 'a \ 'b \ bool" where "fun_of S \ \x y. (x,y) \ S" text \Convenience lemmas\ (*<*) declare fun_of_def [simp] (*>*) lemma rel_list_all2_Cons [iff]: "list_all2 (fun_of S) (x#xs) (y#ys) = ((x,y) \ S \ list_all2 (fun_of S) xs ys)" (*<*)by simp(*>*) lemma rel_list_all2_Cons1: "list_all2 (fun_of S) (x#xs) ys = (\z zs. ys = z#zs \ (x,z) \ S \ list_all2 (fun_of S) xs zs)" (*<*)by (cases ys) auto(*>*) lemma rel_list_all2_Cons2: "list_all2 (fun_of S) xs (y#ys) = (\z zs. xs = z#zs \ (z,y) \ S \ list_all2 (fun_of S) zs ys)" (*<*)by (cases xs) auto(*>*) lemma rel_list_all2_refl: "(\x. (x,x) \ S) \ list_all2 (fun_of S) xs xs" (*<*)by (simp add: list_all2_refl)(*>*) lemma rel_list_all2_antisym: "\ (\x y. \(x,y) \ S; (y,x) \ T\ \ x = y); list_all2 (fun_of S) xs ys; list_all2 (fun_of T) ys xs \ \ xs = ys" (*<*)by (rule list_all2_antisym) auto(*>*) lemma rel_list_all2_trans: "\ \a b c. \(a,b) \ R; (b,c) \ S\ \ (a,c) \ T; list_all2 (fun_of R) as bs; list_all2 (fun_of S) bs cs\ \ list_all2 (fun_of T) as cs" (*<*)by (rule list_all2_trans) auto(*>*) lemma rel_list_all2_update_cong: "\ i S \ \ list_all2 (fun_of S) (xs[i:=x]) (ys[i:=y])" (*<*)by (simp add: list_all2_update_cong)(*>*) lemma rel_list_all2_nthD: "\ list_all2 (fun_of S) xs ys; p < size xs \ \ (xs!p,ys!p) \ S" (*<*)by (drule list_all2_nthD) auto(*>*) lemma rel_list_all2I: "\ length a = length b; \n. n < length a \ (a!n,b!n) \ S \ \ list_all2 (fun_of S) a b" (*<*)by (erule list_all2_all_nthI) simp(*>*) (*<*)declare fun_of_def [simp del](*>*) -lemma list_all2_induct[consumes 1, case_names Nil Cons]: - assumes major: "list_all2 P xs ys" - and Nil: "Q [] []" - and Cons: "\x xs y ys. \ P x y; list_all2 P xs ys; Q xs ys \ \ Q (x # xs) (y # ys)" - shows "Q xs ys" -using major -by(induct xs arbitrary: ys)(auto simp add: list_all2_Cons1 Nil intro!: Cons) - lemma list_all2_split: assumes major: "list_all2 P xs ys" and split: "\x y. P x y \ \z. Q x z \ R z y" shows "\zs. list_all2 Q xs zs \ list_all2 R zs ys" using major by(induct rule: list_all2_induct)(auto dest: split) -lemma list_all2_refl_conv: - "list_all2 P xs xs \ (\x\set xs. P x x)" -unfolding list_all2_conv_all_nth Ball_def in_set_conv_nth -by auto - lemma list_all2_op_eq [simp]: "list_all2 (=) xs ys \ xs = ys" by(induct xs arbitrary: ys)(auto simp add: list_all2_Cons1) lemmas filter_replicate_conv = filter_replicate lemma length_greater_Suc_0_conv: "Suc 0 < length xs \ (\x x' xs'. xs = x # x' # xs')" by(cases xs, auto simp add: neq_Nil_conv) lemmas zip_same_conv = zip_same_conv_map -lemma nth_Cons_subtract: "0 < n \ (x # xs) ! n = xs ! (n - 1)" -by(auto simp add: nth_Cons split: nat.split) - lemma f_nth_set: "\ f (xs ! n) = v; n < length xs \ \ v \ f ` set xs" unfolding set_conv_nth by auto lemma nth_concat_eqI: "\ n = sum_list (map length (take i xss)) + k; i < length xss; k < length (xss ! i); x = xss ! i ! k \ \ concat xss ! n = x" apply(induct xss arbitrary: n i k) apply simp apply simp apply(case_tac i) apply(simp add: nth_append) apply(simp add: nth_append) done lemma replicate_eq_append_conv: "(replicate n x = xs @ ys) = (\m\n. xs = replicate m x \ ys = replicate (n-m) x)" proof(induct n arbitrary: xs ys) case 0 thus ?case by simp next case (Suc n xs ys) have IH: "\xs ys. (replicate n x = xs @ ys) = (\m\n. xs = replicate m x \ ys = replicate (n - m) x)" by fact show ?case proof(unfold replicate_Suc, rule iffI) assume a: "x # replicate n x = xs @ ys" show "\m\Suc n. xs = replicate m x \ ys = replicate (Suc n - m) x" proof(cases xs) case Nil thus ?thesis using a by(auto) next case (Cons X XS) with a have x: "x = X" and "replicate n x = XS @ ys" by auto hence "\m\n. XS = replicate m x \ ys = replicate (n - m) x" by -(rule IH[THEN iffD1]) then obtain m where "m \ n" and XS: "XS = replicate m x" and ys: "ys = replicate (n - m) x" by blast with x Cons show ?thesis by(fastforce) qed next assume "\m\Suc n. xs = replicate m x \ ys = replicate (Suc n - m) x" then obtain m where m: "m \ Suc n" and xs: "xs = replicate m x" and ys: "ys = replicate (Suc n - m) x" by blast thus "x # replicate n x = xs @ ys" by(simp add: replicate_add[THEN sym]) qed qed lemma replicate_Suc_snoc: "replicate (Suc n) x = replicate n x @ [x]" by (metis replicate_Suc replicate_append_same) lemma map_eq_append_conv: "map f xs = ys @ zs \ (\ys' zs'. map f ys' = ys \ map f zs' = zs \ xs = ys' @ zs')" apply(rule iffI) apply(metis append_eq_conv_conj append_take_drop_id drop_map take_map) by(clarsimp) lemma append_eq_map_conv: "ys @ zs = map f xs \ (\ys' zs'. map f ys' = ys \ map f zs' = zs \ xs = ys' @ zs')" unfolding map_eq_append_conv[symmetric] by auto lemma map_eq_map_conv: "map f xs = map g ys \ list_all2 (\x y. f x = g y) xs ys" apply(induct xs arbitrary: ys) apply(auto simp add: list_all2_Cons1 Cons_eq_map_conv) done lemma map_eq_all_nth_conv: "map f xs = ys \ length xs = length ys \ (\n < length xs. f (xs ! n) = ys ! n)" apply(induct xs arbitrary: ys) apply(fastforce simp add: nth_Cons Suc_length_conv split: nat.splits)+ done lemma take_eq_take_le_eq: "\ take n xs = take n ys; m \ n \ \ take m xs = take m ys" by(metis min.absorb_iff1 take_take) -lemma take_list_update_beyond: - "n \ m \ take n (xs[m := x]) = take n xs" -by(cases "n \ length xs")(rule nth_take_lemma, simp_all) - -lemma hd_drop_conv_nth: - "n < length xs \ hd (drop n xs) = xs ! n" -by(rule hd_drop_conv_nth) (metis list.size(3) not_less0) - lemma set_tl_subset: "set (tl xs) \ set xs" by(cases xs) auto lemma tl_conv_drop: "tl xs = drop 1 xs" by(cases xs) simp_all lemma takeWhile_eq_Nil_dropWhile_eq_Nil_imp_Nil: "\ takeWhile P xs = []; dropWhile P xs = [] \ \ xs = []" by (metis dropWhile_eq_drop drop_0 list.size(3)) -lemma takeWhile_eq_Nil_conv: - "takeWhile P xs = [] \ (xs = [] \ \ P (hd xs))" -by(cases xs) simp_all - lemma dropWhile_append1': "dropWhile P xs \ [] \ dropWhile P (xs @ ys) = dropWhile P xs @ ys" by(cases xs) auto lemma dropWhile_append2': "dropWhile P xs = [] \ dropWhile P (xs @ ys) = dropWhile P ys" by(simp) lemma takeWhile_append1': "dropWhile P xs \ [] \ takeWhile P (xs @ ys) = takeWhile P xs" by auto -lemma takeWhile_takeWhile: "takeWhile P (takeWhile Q xs) = takeWhile (\x. P x \ Q x) xs" -by(induct xs) simp_all - lemma dropWhile_eq_ConsD: "dropWhile P xs = y # ys \ y \ set xs \ \ P y" by(induct xs)(auto split: if_split_asm) lemma dropWhile_eq_hd_conv: "dropWhile P xs = hd xs # rest \ xs \ [] \ rest = tl xs \ \ P (hd xs)" -by (metis append_Nil append_is_Nil_conv dropWhile_eq_Cons_conv list.sel(1) neq_Nil_conv takeWhile_dropWhile_id takeWhile_eq_Nil_conv list.sel(3)) - -lemma dropWhile_eq_same_conv: "dropWhile P xs = xs \ (xs = [] \ \ P (hd xs))" -by (metis dropWhile.simps(1) eq_Nil_appendI hd_dropWhile takeWhile_dropWhile_id takeWhile_eq_Nil_conv) +by (metis append_Nil dropWhile_eq_Cons_conv list.sel(1) neq_Nil_conv takeWhile_dropWhile_id takeWhile_eq_Nil_iff list.sel(3)) lemma subset_code [code_unfold]: "set xs \ set ys \ (\x \ set xs. x \ set ys)" by(rule Set.subset_eq) lemma eval_bot [simp]: "Predicate.eval bot = (\_. False)" by(auto simp add: fun_eq_iff) lemma not_is_emptyE: assumes "\ Predicate.is_empty P" obtains x where "Predicate.eval P x" using assms by(fastforce simp add: Predicate.is_empty_def bot_pred_def intro!: pred_iffI) lemma is_emptyD: assumes "Predicate.is_empty P" shows "Predicate.eval P x \ False" using assms by(simp add: Predicate.is_empty_def bot_pred_def bot_apply Set.empty_def) lemma eval_bind_conv: "Predicate.eval (P \ R) y = (\x. Predicate.eval P x \ Predicate.eval (R x) y)" by(blast elim: bindE intro: bindI) lemma eval_single_conv: "Predicate.eval (Predicate.single a) b \ a = b" by(blast intro: singleI elim: singleE) lemma conj_asm_conv_imp: "(A \ B \ PROP C) \ (A \ B \ PROP C)" apply(rule equal_intr_rule) apply(erule meta_mp) apply(erule (1) conjI) apply(erule meta_impE) apply(erule conjunct1) apply(erule meta_mp) apply(erule conjunct2) done lemma meta_all_eq_conv: "(\a. a = b \ PROP P a) \ PROP P b" apply(rule equal_intr_rule) apply(erule meta_allE) apply(erule meta_mp) apply(rule refl) apply(hypsubst) apply assumption done lemma meta_all_eq_conv2: "(\a. b = a \ PROP P a) \ PROP P b" apply(rule equal_intr_rule) apply(erule meta_allE) apply(erule meta_mp) apply(rule refl) apply(hypsubst) apply assumption done lemma disj_split: "P (a \ b) \ (a \ P True) \ (b \ P True) \ (\ a \ \ b \ P False)" apply(cases a) apply(case_tac [!] b) apply auto done lemma disj_split_asm: "P (a \ b) \ \ (a \ \ P True \ b \ \ P True \ \ a \ \ b \ \ P False)" apply(auto simp add: disj_split[of P]) done lemma disjCE: assumes "P \ Q" obtains P | "Q" "\ P" using assms by blast lemma case_option_conv_if: "(case v of None \ f | Some x \ g x) = (if \a. v = Some a then g (the v) else f)" by(simp) lemma LetI: "(\x. x = t \ P x) \ let x = t in P x" by(simp) (* rearrange parameters and premises to allow application of one-point-rules *) (* adapted from Tools/induct.ML and Isabelle Developer Workshop 2010 *) simproc_setup rearrange_eqs ("Pure.all t") = \ let fun swap_params_conv ctxt i j cv = let fun conv1 0 ctxt = Conv.forall_conv (cv o snd) ctxt | conv1 k ctxt = Conv.rewr_conv @{thm swap_params} then_conv Conv.forall_conv (conv1 (k - 1) o snd) ctxt fun conv2 0 ctxt = conv1 j ctxt | conv2 k ctxt = Conv.forall_conv (conv2 (k - 1) o snd) ctxt in conv2 i ctxt end; fun swap_prems_conv 0 = Conv.all_conv | swap_prems_conv i = Conv.implies_concl_conv (swap_prems_conv (i - 1)) then_conv Conv.rewr_conv Drule.swap_prems_eq; fun find_eq ctxt t = let val l = length (Logic.strip_params t); val Hs = Logic.strip_assums_hyp t; fun find (i, (_ $ (Const ("HOL.eq", _) $ Bound j $ _))) = SOME (i, j) | find (i, (_ $ (Const ("HOL.eq", _) $ _ $ Bound j))) = SOME (i, j) | find _ = NONE in (case get_first find (map_index I Hs) of NONE => NONE | SOME (0, 0) => NONE | SOME (i, j) => SOME (i, l - j - 1, j)) end; fun mk_swap_rrule ctxt ct = (case find_eq ctxt (Thm.term_of ct) of NONE => NONE | SOME (i, k, j) => SOME (swap_params_conv ctxt k j (K (swap_prems_conv i)) ct)) in fn _ => mk_swap_rrule end \ declare [[simproc del: rearrange_eqs]] lemmas meta_onepoint = meta_all_eq_conv meta_all_eq_conv2 lemma meta_all2_eq_conv: "(\a b. a = c \ PROP P a b) \ (\b. PROP P c b)" apply(rule equal_intr_rule) apply(erule meta_allE)+ apply(erule meta_mp) apply(rule refl) apply(erule meta_allE) apply simp done lemma meta_all3_eq_conv: "(\a b c. a = d \ PROP P a b c) \ (\b c. PROP P d b c)" apply(rule equal_intr_rule) apply(erule meta_allE)+ apply(erule meta_mp) apply(rule refl) apply(erule meta_allE)+ apply simp done lemma meta_all4_eq_conv: "(\a b c d. a = e \ PROP P a b c d) \ (\b c d. PROP P e b c d)" apply(rule equal_intr_rule) apply(erule meta_allE)+ apply(erule meta_mp) apply(rule refl) apply(erule meta_allE)+ apply simp done lemma meta_all5_eq_conv: "(\a b c d e. a = f \ PROP P a b c d e) \ (\b c d e. PROP P f b c d e)" apply(rule equal_intr_rule) apply(erule meta_allE)+ apply(erule meta_mp) apply(rule refl) apply(erule meta_allE)+ apply simp done -lemma inj_on_image_mem_iff: - "\ inj_on f A; B \ A; a \ A \ \ f a \ f ` B \ a \ B" -by(metis inv_into_f_eq inv_into_image_cancel rev_image_eqI) - -lemma sum_hom: - assumes hom_add [simp]: "\a b. f (a + b) = f a + f b" - and hom_0 [simp]: "f 0 = 0" - shows "sum (f \ h) A = f (sum h A)" -proof(cases "finite A") - case False thus ?thesis by simp -next - case True thus ?thesis - by(induct) simp_all -qed - lemma sum_upto_add_nat: "a \ b \ sum f {..<(a :: nat)} + sum f {a.. nat" shows "(\a. f a + g a) = (\a. 0) \ f = (\a .0) \ g = (\a. 0)" by(auto simp add: fun_eq_iff) lemma in_ran_conv: "v \ ran m \ (\k. m k = Some v)" by(simp add: ran_def) lemma map_le_dom_eq_conv_eq: "\ m \\<^sub>m m'; dom m = dom m' \ \ m = m'" by (metis map_le_antisym map_le_def) lemma map_leI: "(\k v. f k = Some v \ g k = Some v) \ f \\<^sub>m g" unfolding map_le_def by auto lemma map_le_SomeD: "\ m \\<^sub>m m'; m x = \y\ \ \ m' x = \y\" by(auto simp add: map_le_def dest: bspec) lemma map_le_same_upd: "f x = None \ f \\<^sub>m f(x \ y)" by(auto simp add: map_le_def) lemma map_upd_map_add: "X(V \ v) = (X ++ [V \ v])" by(simp) lemma foldr_filter_conv: "foldr f (filter P xs) = foldr (\x s. if P x then f x s else s) xs" by(induct xs)(auto intro: ext) lemma foldr_insert_conv_set: "foldr insert xs A = A \ set xs" by(induct xs arbitrary: A) auto lemma snd_o_Pair_conv_id: "snd o Pair a = id" by(simp add: o_def id_def) lemma if_intro: "\ P \ A; \ P \ B \ \ if P then A else B" by(auto) lemma ex_set_conv: "(\x. x \ set xs) \ xs \ []" apply(auto) apply(auto simp add: neq_Nil_conv) done lemma subset_Un1: "A \ B \ A \ B \ C" by blast lemma subset_Un2: "A \ C \ A \ B \ C" by blast lemma subset_insert: "A \ B \ A \ insert a B" by blast lemma UNION_subsetD: "\ (\x\A. f x) \ B; a \ A \ \ f a \ B" by blast lemma Collect_eq_singleton_conv: "{a. P a} = {a} \ P a \ (\a'. P a' \ a = a')" by(auto) lemma Collect_conv_UN_singleton: "{f x|x. x \ P} = (\x\P. {f x})" by blast lemma if_else_if_else_eq_if_else [simp]: "(if b then x else if b then y else z) = (if b then x else z)" by(simp) lemma rec_prod_split [simp]: "old.rec_prod = case_prod" by(simp add: fun_eq_iff) lemma inj_Pair_snd [simp]: "inj (Pair x)" by(rule injI) auto lemma rtranclp_False [simp]: "(\a b. False)\<^sup>*\<^sup>* = (=)" by(auto simp add: fun_eq_iff elim: rtranclp_induct) lemmas rtranclp_induct3 = rtranclp_induct[where a="(ax, ay, az)" and b="(bx, by, bz)", split_rule, consumes 1, case_names refl step] lemmas tranclp_induct3 = tranclp_induct[where a="(ax, ay, az)" and b="(bx, by, bz)", split_rule, consumes 1, case_names refl step] lemmas rtranclp_induct4 = rtranclp_induct[where a="(ax, ay, az, aw)" and b="(bx, by, bz, bw)", split_rule, consumes 1, case_names refl step] lemmas tranclp_induct4 = tranclp_induct[where a="(ax, ay, az, aw)" and b="(bx, by, bz, bw)", split_rule, consumes 1, case_names refl step] lemmas converse_tranclp_induct2 = converse_tranclp_induct [of _ "(ax,ay)" "(bx,by)", split_rule, consumes 1, case_names base step] lemma wfP_induct' [consumes 1, case_names wfP]: "\wfP r; \x. (\y. r y x \ P y) \ P x\ \ P a" by(blast intro: wfP_induct) lemma wfP_induct2 [consumes 1, case_names wfP]: "\wfP r; \x x'. (\y y'. r (y, y') (x, x') \ P y y') \ P x x' \ \ P x x'" by(drule wfP_induct'[where P="\(x, y). P x y"]) blast+ lemma wfP_minimalE: assumes "wfP r" and "P x" obtains z where "P z" "r^** z x" "\y. r y z \ \ P y" proof - let ?Q = "\x'. P x' \ r^** x' x" from \P x\ have "?Q x" by blast from \wfP r\ have "\Q. x \ Q \ (\z\Q. \y. r y z \ y \ Q)" unfolding wfP_eq_minimal by blast from this[rule_format, of "Collect ?Q"] \?Q x\ obtain z where "?Q z" and min: "\y. r y z \ \ ?Q y" by auto from \?Q z\ have "P z" "r^** z x" by auto moreover { fix y assume "r y z" hence "\ ?Q y" by(rule min) moreover from \r y z\ \r^** z x\ have "r^** y x" by(rule converse_rtranclp_into_rtranclp) ultimately have "\ P y" by blast } ultimately show thesis .. qed lemma coinduct_set_wf [consumes 3, case_names coinduct, case_conclusion coinduct wait step]: assumes "mono f" "wf r" "(a, b) \ X" and step: "\x b. (x, b) \ X \ (\b'. (b', b) \ r \ (x, b') \ X) \ (x \ f (fst ` X \ gfp f))" shows "a \ gfp f" proof - from \(a, b) \ X\ have "a \ fst ` X" by(auto intro: rev_image_eqI) moreover { fix a b assume "(a, b) \ X" with \wf r\ have "a \ f (fst ` X \ gfp f)" by(induct)(blast dest: step) } hence "fst ` X \ f (fst ` X \ gfp f)" by(auto) ultimately show ?thesis by(rule coinduct_set[OF \mono f\]) qed subsection \reflexive transitive closure for label relations\ inductive converse3p :: "('a \ 'b \ 'c \ bool) \ 'c \ 'b \ 'a \ bool" for r :: "'a \ 'b \ 'c \ bool" where converse3pI: "r a b c \ converse3p r c b a" lemma converse3p_converse3p: "converse3p (converse3p r) = r" by(auto intro: converse3pI intro!: ext elim: converse3p.cases) lemma converse3pD: "converse3p r c b a \ r a b c" by(auto elim: converse3p.cases) inductive rtrancl3p :: "('a \ 'b \ 'a \ bool) \ 'a \ 'b list \ 'a \ bool" for r :: "'a \ 'b \ 'a \ bool" where rtrancl3p_refl [intro!, simp]: "rtrancl3p r a [] a" | rtrancl3p_step: "\ rtrancl3p r a bs a'; r a' b a'' \ \ rtrancl3p r a (bs @ [b]) a''" lemmas rtrancl3p_induct3 = rtrancl3p.induct[of _ "(ax,ay,az)" _ "(cx,cy,cz)", split_format (complete), consumes 1, case_names refl step] lemmas rtrancl3p_induct4 = rtrancl3p.induct[of _ "(ax,ay,az,aw)" _ "(cx,cy,cz,cw)", split_format (complete), consumes 1, case_names refl step] lemma rtrancl3p_induct4' [consumes 1, case_names refl step]: assumes major: "rtrancl3p r (ax, (ay, az), aw) bs (cx, (cy, cz), cw)" and refl: "\a aa b ba. P a aa b ba [] a aa b ba" and step: "\a aa b ba bs ab ac bb bc bd ad ae be bf. \ rtrancl3p r (a, (aa, b), ba) bs (ab, (ac, bb), bc); P a aa b ba bs ab ac bb bc; r (ab, (ac, bb), bc) bd (ad, (ae, be), bf) \ \ P a aa b ba (bs @ [bd]) ad ae be bf" shows "P ax ay az aw bs cx cy cz cw" using major apply - apply(drule_tac P="\(a, (aa, b), ba) bs (cx, (cy, cz), cw). P a aa b ba bs cx cy cz cw" in rtrancl3p.induct) by(auto intro: refl step) lemma rtrancl3p_induct1: "\ rtrancl3p r a bs c; P a; \bs c b d. \ rtrancl3p r a bs c; r c b d; P c \ \ P d \ \ P c" apply(induct rule: rtrancl3p.induct) apply(auto) done inductive_cases rtrancl3p_cases: "rtrancl3p r x [] y" "rtrancl3p r x (b # bs) y" lemma rtrancl3p_trans [trans]: assumes one: "rtrancl3p r a bs a'" and two: "rtrancl3p r a' bs' a''" shows "rtrancl3p r a (bs @ bs') a''" using two one proof(induct rule: rtrancl3p.induct) case rtrancl3p_refl thus ?case by simp next case rtrancl3p_step thus ?case by(auto simp only: append_assoc[symmetric] intro: rtrancl3p.rtrancl3p_step) qed lemma rtrancl3p_step_converse: assumes step: "r a b a'" and stepify: "rtrancl3p r a' bs a''" shows "rtrancl3p r a (b # bs) a''" using stepify step proof(induct rule: rtrancl3p.induct) case rtrancl3p_refl with rtrancl3p.rtrancl3p_refl[where r=r and a=a] show ?case by(auto dest: rtrancl3p.rtrancl3p_step simp del: rtrancl3p.rtrancl3p_refl) next case rtrancl3p_step thus ?case unfolding append_Cons[symmetric] by -(rule rtrancl3p.rtrancl3p_step) qed lemma converse_rtrancl3p_step: "rtrancl3p r a (b # bs) a'' \ \a'. r a b a' \ rtrancl3p r a' bs a''" apply(induct bs arbitrary: a'' rule: rev_induct) apply(erule rtrancl3p.cases) apply(clarsimp)+ apply(erule rtrancl3p.cases) apply(clarsimp) apply(rule_tac x="a''a" in exI) apply(clarsimp) apply(clarsimp) apply(erule rtrancl3p.cases) apply(clarsimp) apply(clarsimp) by(fastforce intro: rtrancl3p_step) lemma converse_rtrancl3pD: "rtrancl3p (converse3p r) a' bs a \ rtrancl3p r a (rev bs) a'" apply(induct rule: rtrancl3p.induct) apply(fastforce intro: rtrancl3p.intros) apply(auto dest: converse3pD intro: rtrancl3p_step_converse) done lemma rtrancl3p_converseD: "rtrancl3p r a bs a' \ rtrancl3p (converse3p r) a' (rev bs) a" proof(induct rule: rtrancl3p.induct) case rtrancl3p_refl thus ?case by(auto intro: rtrancl3p.intros) next case rtrancl3p_step thus ?case by(auto intro: rtrancl3p_step_converse converse3p.intros) qed lemma rtrancl3p_converse_induct [consumes 1, case_names refl step]: assumes ih: "rtrancl3p r a bs a''" and refl: "\a. P a [] a" and step: "\a b a' bs a''. \ rtrancl3p r a' bs a''; r a b a'; P a' bs a'' \ \ P a (b # bs) a''" shows "P a bs a''" using ih proof(induct bs arbitrary: a a'') case Nil thus ?case by(auto elim: rtrancl3p.cases intro: refl) next case Cons thus ?case by(auto dest!: converse_rtrancl3p_step intro: step) qed lemma rtrancl3p_converse_induct' [consumes 1, case_names refl step]: assumes ih: "rtrancl3p r a bs a''" and refl: "P a'' []" and step: "\a b a' bs. \ rtrancl3p r a' bs a''; r a b a'; P a' bs \ \ P a (b # bs)" shows "P a bs" using rtrancl3p_converse_induct[OF ih, where P="\a bs a'. a' = a'' \ P a bs"] by(auto intro: refl step) lemma rtrancl3p_converseE[consumes 1, case_names refl step]: "\ rtrancl3p r a bs a''; \ a = a''; bs = [] \ \ thesis; \bs' b a'. \ bs = b # bs'; r a b a'; rtrancl3p r a' bs' a'' \ \ thesis \ \ thesis" by(induct rule: rtrancl3p_converse_induct)(auto) lemma rtrancl3p_induct' [consumes 1, case_names refl step]: assumes major: "rtrancl3p r a bs c" and refl: "P a [] a" and step: "\bs a' b a''. \ rtrancl3p r a bs a'; P a bs a'; r a' b a'' \ \ P a (bs @ [b]) a''" shows "P a bs c" proof - from major have "(P a [] a \ (\bs a' b a''. rtrancl3p r a bs a' \ P a bs a' \ r a' b a'' \ P a (bs @ [b]) a'')) \ P a bs c" by-(erule rtrancl3p.induct, blast+) with refl step show ?thesis by blast qed lemma r_into_rtrancl3p: "r a b a' \ rtrancl3p r a [b] a'" by(rule rtrancl3p_step_converse) auto lemma rtrancl3p_appendE: assumes "rtrancl3p r a (bs @ bs') a''" obtains a' where "rtrancl3p r a bs a'" "rtrancl3p r a' bs' a''" using assms apply(induct a "bs @ bs'" arbitrary: bs rule: rtrancl3p_converse_induct') apply(fastforce intro: rtrancl3p_step_converse simp add: Cons_eq_append_conv)+ done lemma rtrancl3p_Cons: "rtrancl3p r a (b # bs) a' \ (\a''. r a b a'' \ rtrancl3p r a'' bs a')" by(auto intro: rtrancl3p_step_converse converse_rtrancl3p_step) lemma rtrancl3p_Nil: "rtrancl3p r a [] a' \ a = a'" by(auto elim: rtrancl3p_cases) definition invariant3p :: "('a \ 'b \ 'a \ bool) \ 'a set \ bool" where "invariant3p r I \ (\s tl s'. s \ I \ r s tl s' \ s' \ I)" lemma invariant3pI: "(\s tl s'. \ s \ I; r s tl s' \ \ s' \ I) \ invariant3p r I" unfolding invariant3p_def by blast lemma invariant3pD: "\tl. \ invariant3p r I; r s tl s'; s \ I \ \ s' \ I" unfolding invariant3p_def by(blast) lemma invariant3p_rtrancl3p: assumes inv: "invariant3p r I" and rtrancl: "rtrancl3p r a bs a'" and start: "a \ I" shows "a' \ I" using rtrancl start by(induct)(auto dest: invariant3pD[OF inv]) lemma invariant3p_UNIV [simp, intro!]: "invariant3p r UNIV" by(blast intro: invariant3pI) lemma invarinat3p_empty [simp, intro!]: "invariant3p r {}" by(blast intro: invariant3pI) lemma invariant3p_IntI [simp, intro]: "\ invariant3p r I; invariant3p r J \ \ invariant3p r (I \ J)" by(blast dest: invariant3pD intro: invariant3pI) subsection \Concatenation for @{typ String.literal}\ definition concat :: "String.literal list \ String.literal" where [code_abbrev, code del]: "concat = sum_list" lemma explode_add [simp]: "String.explode (s + t) = String.explode s @ String.explode t" by (fact plus_literal.rep_eq) code_printing constant concat \ (SML) "String.concat" and (Haskell) "concat" and (OCaml) "String.concat \"\"" hide_const (open) concat end diff --git a/thys/JinjaThreads/Common/ExternalCallWF.thy b/thys/JinjaThreads/Common/ExternalCallWF.thy --- a/thys/JinjaThreads/Common/ExternalCallWF.thy +++ b/thys/JinjaThreads/Common/ExternalCallWF.thy @@ -1,765 +1,765 @@ (* Title: JinjaThreads/Common/ExternalCallWF.thy Author: Andreas Lochbihler *) section \Properties of external calls in well-formed programs\ theory ExternalCallWF imports WellForm "../Framework/FWSemantics" begin lemma external_WT_defs_is_type: assumes "wf_prog wf_md P" and "C\M(Ts) :: T" shows "is_class P C" and "is_type P T" "set Ts \ types P" using assms by(auto elim: external_WT_defs.cases) context heap_base begin lemma WT_red_external_aggr_imp_red_external: "\ wf_prog wf_md P; (ta, va, h') \ red_external_aggr P t a M vs h; P,h \ a\M(vs) : U; P,h \ t \t \ \ P,t \ \a\M(vs), h\ -ta\ext \va, h'\" apply(drule tconfD) apply(erule external_WT'.cases) apply(clarsimp) apply(drule (1) sees_wf_native) apply(erule external_WT_defs.cases) apply(case_tac [!] hT) apply(auto 4 4 simp add: red_external_aggr_def widen_Class intro: red_external.intros heap_base.red_external.intros[where addr2thread_id=addr2thread_id and thread_id2addr=thread_id2addr and spurious_wakeups=True and empty_heap=empty_heap and allocate=allocate and typeof_addr=typeof_addr and heap_read=heap_read and heap_write=heap_write] heap_base.red_external.intros[where addr2thread_id=addr2thread_id and thread_id2addr=thread_id2addr and spurious_wakeups=False and empty_heap=empty_heap and allocate=allocate and typeof_addr=typeof_addr and heap_read=heap_read and heap_write=heap_write] split: if_split_asm dest: sees_method_decl_above) done lemma WT_red_external_list_conv: "\ wf_prog wf_md P; P,h \ a\M(vs) : U; P,h \ t \t \ \ P,t \ \a\M(vs), h\ -ta\ext \va, h'\ \ (ta, va, h') \ red_external_aggr P t a M vs h" by(blast intro: WT_red_external_aggr_imp_red_external red_external_imp_red_external_aggr) lemma red_external_new_thread_sees: "\ wf_prog wf_md P; P,t \ \a\M(vs), h\ -ta\ext \va, h'\; NewThread t' (C, M', a') h'' \ set \ta\\<^bsub>t\<^esub> \ \ typeof_addr h' a' = \Class_type C\ \ (\T meth D. P \ C sees M':[]\T = \meth\ in D)" by(fastforce elim!: red_external.cases simp add: widen_Class ta_upd_simps dest: sub_Thread_sees_run) end subsection \Preservation of heap conformance\ context heap_conf_read begin lemma hconf_heap_copy_loc_mono: assumes "heap_copy_loc a a' al h obs h'" and "hconf h" and "P,h \ a@al : T" "P,h \ a'@al : T" shows "hconf h'" proof - from \heap_copy_loc a a' al h obs h'\ obtain v where read: "heap_read h a al v" and "write": "heap_write h a' al v h'" by cases auto from read \P,h \ a@al : T\ \hconf h\ have "P,h \ v :\ T" by(rule heap_read_conf) with "write" \hconf h\ \P,h \ a'@al : T\ show ?thesis by(rule hconf_heap_write_mono) qed lemma hconf_heap_copies_mono: assumes "heap_copies a a' als h obs h'" and "hconf h" and "list_all2 (\al T. P,h \ a@al : T) als Ts" and "list_all2 (\al T. P,h \ a'@al : T) als Ts" shows "hconf h'" using assms proof(induct arbitrary: Ts) case Nil thus ?case by simp next case (Cons al h ob h' als obs h'') note step = \heap_copy_loc a a' al h ob h'\ from \list_all2 (\al T. P,h \ a@al : T) (al # als) Ts\ obtain T Ts' where [simp]: "Ts = T # Ts'" and "P,h \ a@al : T" "list_all2 (\al T. P,h \ a@al : T) als Ts'" by(auto simp add: list_all2_Cons1) from \list_all2 (\al T. P,h \ a'@al : T) (al # als) Ts\ have "P,h \ a'@al : T" "list_all2 (\al T. P,h \ a'@al : T) als Ts'" by simp_all from step \hconf h\ \P,h \ a@al : T\ \P,h \ a'@al : T\ have "hconf h'" by(rule hconf_heap_copy_loc_mono) moreover from step have "h \ h'" by(rule hext_heap_copy_loc) from \list_all2 (\al T. P,h \ a@al : T) als Ts'\ have "list_all2 (\al T. P,h' \ a@al : T) als Ts'" by(rule list_all2_mono)(rule addr_loc_type_hext_mono[OF _ \h \ h'\]) moreover from \list_all2 (\al T. P,h \ a'@al : T) als Ts'\ have "list_all2 (\al T. P,h' \ a'@al : T) als Ts'" by(rule list_all2_mono)(rule addr_loc_type_hext_mono[OF _ \h \ h'\]) ultimately show ?case by(rule Cons) qed lemma hconf_heap_clone_mono: assumes "heap_clone P h a h' res" and "hconf h" shows "hconf h'" using \heap_clone P h a h' res\ proof cases case CloneFail thus ?thesis using \hconf h\ by(fastforce intro: hconf_heap_ops_mono dest: typeof_addr_is_type) next case (ObjClone C h'' a' FDTs obs) note FDTs = \P \ C has_fields FDTs\ let ?als = "map (\((F, D), Tfm). CField D F) FDTs" let ?Ts = "map (\(FD, T). fst (the (map_of FDTs FD))) FDTs" note \heap_copies a a' ?als h'' obs h'\ moreover from \typeof_addr h a = \Class_type C\\ \hconf h\ have "is_class P C" by(auto dest: typeof_addr_is_type) from \(h'', a') \ allocate h (Class_type C)\ have "h \ h''" "hconf h''" by(rule hext_heap_ops hconf_allocate_mono)+(simp_all add: \hconf h\ \is_class P C\) note \hconf h''\ moreover from \typeof_addr h a = \Class_type C\\ FDTs have "list_all2 (\al T. P,h \ a@al : T) ?als ?Ts" - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce intro: addr_loc_type.intros simp add: has_field_def dest: weak_map_of_SomeI) hence "list_all2 (\al T. P,h'' \ a@al : T) ?als ?Ts" by(rule list_all2_mono)(rule addr_loc_type_hext_mono[OF _ \h \ h''\]) moreover from \(h'', a') \ allocate h (Class_type C)\ \is_class P C\ have "typeof_addr h'' a' = \Class_type C\" by(auto dest: allocate_SomeD) with FDTs have "list_all2 (\al T. P,h'' \ a'@al : T) ?als ?Ts" - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce intro: addr_loc_type.intros simp add: has_field_def dest: weak_map_of_SomeI) ultimately have "hconf h'" by(rule hconf_heap_copies_mono) thus ?thesis using ObjClone by simp next case (ArrClone T n h'' a' FDTs obs) let ?als = "map (\((F, D), Tfm). CField D F) FDTs @ map ACell [0..heap_copies a a' ?als h'' obs h'\ moreover from \typeof_addr h a = \Array_type T n\\ \hconf h\ have "is_type P (T\\)" by(auto dest: typeof_addr_is_type) from \(h'', a') \ allocate h (Array_type T n)\ have "h \ h''" "hconf h''" by(rule hext_heap_ops hconf_allocate_mono)+(simp_all add: \hconf h\ \is_type P (T\\)\[simplified]) note \hconf h''\ moreover from \h \ h''\ \typeof_addr h a = \Array_type T n\\ have type'a: "typeof_addr h'' a = \Array_type T n\" by(auto intro: hext_arrD) note FDTs = \P \ Object has_fields FDTs\ from type'a FDTs have "list_all2 (\al T. P,h'' \ a@al : T) ?als ?Ts" - by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def distinct_fst_def list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv dest: weak_map_of_SomeI) + by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def distinct_fst_def list_all2_append list_all2_map1 list_all2_map2 list_all2_same dest: weak_map_of_SomeI) moreover from \(h'', a') \ allocate h (Array_type T n)\ \is_type P (T\\)\ have "typeof_addr h'' a' = \Array_type T n\" by(auto dest: allocate_SomeD) hence "list_all2 (\al T. P,h'' \ a'@al : T) ?als ?Ts" using FDTs - by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def distinct_fst_def list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv dest: weak_map_of_SomeI) + by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def distinct_fst_def list_all2_append list_all2_map1 list_all2_map2 list_all2_same dest: weak_map_of_SomeI) ultimately have "hconf h'" by(rule hconf_heap_copies_mono) thus ?thesis using ArrClone by simp qed theorem external_call_hconf: assumes major: "P,t \ \a\M(vs), h\ -ta\ext \va, h'\" and minor: "P,h \ a\M(vs) : U" "hconf h" shows "hconf h'" using major minor by cases(fastforce intro: hconf_heap_clone_mono)+ end context heap_base begin primrec conf_extRet :: "'m prog \ 'heap \ 'addr extCallRet \ ty \ bool" where "conf_extRet P h (RetVal v) T = (P,h \ v :\ T)" | "conf_extRet P h (RetExc a) T = (P,h \ Addr a :\ Class Throwable)" | "conf_extRet P h RetStaySame T = True" end context heap_conf begin lemma red_external_conf_extRet: assumes wf: "wf_prog wf_md P" shows "\P,t \ \a\M(vs), h\ -ta\ext \va, h'\; P,h \ a\M(vs) : U; hconf h; preallocated h; P,h \ t \t \ \ conf_extRet P h' va U" using wf apply - apply(frule red_external_hext) apply(drule (1) preallocated_hext) apply(auto elim!: red_external.cases external_WT'.cases external_WT_defs_cases dest!: sees_wf_native[OF wf]) apply(auto simp add: conf_def tconf_def intro: xcpt_subcls_Throwable dest!: hext_heap_write) apply(case_tac hT) apply(auto 4 4 dest!: typeof_addr_heap_clone dest: typeof_addr_is_type intro: widen_array_object subcls_C_Object) done end subsection \Progress theorems for external calls\ context heap_progress begin lemma heap_copy_loc_progress: assumes hconf: "hconf h" and alconfa: "P,h \ a@al : T" and alconfa': "P,h \ a'@al : T" shows "\v h'. heap_copy_loc a a' al h ([ReadMem a al v, WriteMem a' al v]) h' \ P,h \ v :\ T \ hconf h'" proof - from heap_read_total[OF hconf alconfa] obtain v where "heap_read h a al v" "P,h \ v :\ T" by blast moreover from heap_write_total[OF hconf alconfa' \P,h \ v :\ T\] obtain h' where "heap_write h a' al v h'" .. moreover hence "hconf h'" using hconf alconfa' \P,h \ v :\ T\ by(rule hconf_heap_write_mono) ultimately show ?thesis by(blast intro: heap_copy_loc.intros) qed lemma heap_copies_progress: assumes "hconf h" and "list_all2 (\al T. P,h \ a@al : T) als Ts" and "list_all2 (\al T. P,h \ a'@al : T) als Ts" shows "\vs h'. heap_copies a a' als h (concat (map (\(al, v). [ReadMem a al v, WriteMem a' al v]) (zip als vs))) h' \ hconf h'" using assms proof(induct als arbitrary: h Ts) case Nil thus ?case by(auto intro: heap_copies.Nil) next case (Cons al als) from \list_all2 (\al T. P,h \ a@al : T) (al # als) Ts\ obtain T' Ts' where [simp]: "Ts = T' # Ts'" and "P,h \ a@al : T'" "list_all2 (\al T. P,h \ a@al : T) als Ts'" by(auto simp add: list_all2_Cons1) from \list_all2 (\al T. P,h \ a'@al : T) (al # als) Ts\ have "P,h \ a'@al : T'" and "list_all2 (\al T. P,h \ a'@al : T) als Ts'" by simp_all from \hconf h\ \P,h \ a@al : T'\ \P,h \ a'@al : T'\ obtain v h' where "heap_copy_loc a a' al h [ReadMem a al v, WriteMem a' al v] h'" and "hconf h'" by(fastforce dest: heap_copy_loc_progress) moreover hence "h \ h'" by-(rule hext_heap_copy_loc) { note \hconf h'\ moreover from \list_all2 (\al T. P,h \ a@al : T) als Ts'\ have "list_all2 (\al T. P,h' \ a@al : T) als Ts'" by(rule list_all2_mono)(rule addr_loc_type_hext_mono[OF _ \h \ h'\]) moreover from \list_all2 (\al T. P,h \ a'@al : T) als Ts'\ have "list_all2 (\al T. P,h' \ a'@al : T) als Ts'" by(rule list_all2_mono)(rule addr_loc_type_hext_mono[OF _ \h \ h'\]) ultimately have "\vs h''. heap_copies a a' als h' (concat (map (\(al, v). [ReadMem a al v, WriteMem a' al v]) (zip als vs))) h'' \ hconf h''" by(rule Cons) } then obtain vs h'' where "heap_copies a a' als h' (concat (map (\(al, v). [ReadMem a al v, WriteMem a' al v]) (zip als vs))) h''" and "hconf h''" by blast ultimately have "heap_copies a a' (al # als) h ([ReadMem a al v, WriteMem a' al v] @ (concat (map (\(al, v). [ReadMem a al v, WriteMem a' al v]) (zip als vs)))) h''" by- (rule heap_copies.Cons) also have "[ReadMem a al v, WriteMem a' al v] @ (concat (map (\(al, v). [ReadMem a al v, WriteMem a' al v]) (zip als vs))) = (concat (map (\(al, v). [ReadMem a al v, WriteMem a' al v]) (zip (al # als) (v # vs))))" by simp finally show ?case using \hconf h''\ by blast qed lemma heap_clone_progress: assumes wf: "wf_prog wf_md P" and typea: "typeof_addr h a = \hT\" and hconf: "hconf h" shows "\h' res. heap_clone P h a h' res" proof - from typea hconf have "is_htype P hT" by(rule typeof_addr_is_type) show ?thesis proof(cases "allocate h hT = {}") case True with typea CloneFail[of h a hT P] show ?thesis by auto next case False then obtain h' a' where new: "(h', a') \ allocate h hT" by(rule not_empty_pairE) hence "h \ h'" by(rule hext_allocate) have "hconf h'" using new hconf \is_htype P hT\ by(rule hconf_allocate_mono) show ?thesis proof(cases hT) case [simp]: (Class_type C) from \is_htype P hT\ have "is_class P C" by simp from wf_Fields_Ex[OF wf this] obtain FDTs where FDTs: "P \ C has_fields FDTs" .. let ?als = "map (\((F, D), Tfm). CField D F) FDTs" let ?Ts = "map (\(FD, T). fst (the (map_of FDTs FD))) FDTs" from typea FDTs have "list_all2 (\al T. P,h \ a@al : T) ?als ?Ts" - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce intro: addr_loc_type.intros simp add: has_field_def dest: weak_map_of_SomeI) hence "list_all2 (\al T. P,h' \ a@al : T) ?als ?Ts" by(rule list_all2_mono)(simp add: addr_loc_type_hext_mono[OF _ \h \ h'\] split_def) moreover from new \is_class P C\ have "typeof_addr h' a' = \Class_type C\" by(auto dest: allocate_SomeD) with FDTs have "list_all2 (\al T. P,h' \ a'@al : T) ?als ?Ts" - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce intro: addr_loc_type.intros map_of_SomeI simp add: has_field_def dest: weak_map_of_SomeI) ultimately obtain obs h'' where "heap_copies a a' ?als h' obs h''" "hconf h''" by(blast dest: heap_copies_progress[OF \hconf h'\]) with typea new FDTs ObjClone[of h a C h' a' P FDTs obs h''] show ?thesis by auto next case [simp]: (Array_type T n) from wf obtain FDTs where FDTs: "P \ Object has_fields FDTs" by(blast dest: wf_Fields_Ex is_class_Object) let ?als = "map (\((F, D), Tfm). CField D F) FDTs @ map ACell [0..h \ h'\ typea have type'a: "typeof_addr h' a = \Array_type T n\" by(auto intro: hext_arrD) from type'a FDTs have "list_all2 (\al T. P,h' \ a@al : T) ?als ?Ts" - by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv dest: weak_map_of_SomeI) + by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_same dest: weak_map_of_SomeI) moreover from new \is_htype P hT\ have "typeof_addr h' a' = \Array_type T n\" by(auto dest: allocate_SomeD) hence "list_all2 (\al T. P,h' \ a'@al : T) ?als ?Ts" using FDTs - by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv dest: weak_map_of_SomeI) + by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_same dest: weak_map_of_SomeI) ultimately obtain obs h'' where "heap_copies a a' ?als h' obs h''" "hconf h''" by(blast dest: heap_copies_progress[OF \hconf h'\]) with typea new FDTs ArrClone[of h a T n h' a' P FDTs obs h''] show ?thesis by auto qed qed qed theorem external_call_progress: assumes wf: "wf_prog wf_md P" and wt: "P,h \ a\M(vs) : U" and hconf: "hconf h" shows "\ta va h'. P,t \ \a\M(vs), h\ -ta\ext \va, h'\" proof - note [simp del] = split_paired_Ex from wt obtain hT Ts Ts' D where T: "typeof_addr h a = \hT\" and Ts: "map typeof\<^bsub>h\<^esub> vs = map Some Ts" and "P \ class_type_of hT sees M:Ts'\U = Native in D" and subTs: "P \ Ts [\] Ts'" unfolding external_WT'_iff by blast from wf \P \ class_type_of hT sees M:Ts'\U = Native in D\ have "D\M(Ts') :: U" by(rule sees_wf_native) moreover from \P \ class_type_of hT sees M:Ts'\U = Native in D\ have "P \ ty_of_htype hT \ Class D" by(cases hT)(auto dest: sees_method_decl_above intro: widen_trans widen_array_object) ultimately show ?thesis using T Ts subTs proof cases assume [simp]: "D = Object" "M = clone" "Ts' = []" "U = Class Object" from heap_clone_progress[OF wf T hconf] obtain h' res where "heap_clone P h a h' res" by blast thus ?thesis using subTs Ts by(cases res)(auto intro: red_external.intros) qed(auto simp add: widen_Class intro: red_external.intros) qed end subsection \Lemmas for preservation of deadlocked threads\ context heap_progress begin lemma red_external_wt_hconf_hext: assumes wf: "wf_prog wf_md P" and red: "P,t \ \a\M(vs),h\ -ta\ext \va,h'\" and hext: "h'' \ h" and wt: "P,h'' \ a\M(vs) : U" and tconf: "P,h'' \ t \t" and hconf: "hconf h''" shows "\ta' va' h'''. P,t \ \a\M(vs),h''\ -ta'\ext \va', h'''\ \ collect_locks \ta\\<^bsub>l\<^esub> = collect_locks \ta'\\<^bsub>l\<^esub> \ collect_cond_actions \ta\\<^bsub>c\<^esub> = collect_cond_actions \ta'\\<^bsub>c\<^esub> \ collect_interrupts \ta\\<^bsub>i\<^esub> = collect_interrupts \ta'\\<^bsub>i\<^esub>" using red wt hext proof cases case (RedClone obs a') from wt obtain hT C Ts Ts' D where T: "typeof_addr h'' a = \hT\" unfolding external_WT'_iff by blast from heap_clone_progress[OF wf T hconf] obtain h''' res where "heap_clone P h'' a h''' res" by blast thus ?thesis using RedClone by(cases res)(fastforce intro: red_external.intros)+ next case RedCloneFail from wt obtain hT Ts Ts' where T: "typeof_addr h'' a = \hT\" unfolding external_WT'_iff by blast from heap_clone_progress[OF wf T hconf] obtain h''' res where "heap_clone P h'' a h''' res" by blast thus ?thesis using RedCloneFail by(cases res)(fastforce intro: red_external.intros)+ qed(fastforce simp add: ta_upd_simps elim!: external_WT'.cases intro: red_external.intros[simplified] dest: typeof_addr_hext_mono)+ lemma red_external_wf_red: assumes wf: "wf_prog wf_md P" and red: "P,t \ \a\M(vs), h\ -ta\ext \va, h'\" and tconf: "P,h \ t \t" and hconf: "hconf h" and wst: "wset s t = None \ (M = wait \ (\w. wset s t = \PostWS w\))" obtains ta' va' h'' where "P,t \ \a\M(vs), h\ -ta'\ext \va', h''\" and "final_thread.actions_ok final s t ta' \ final_thread.actions_ok' s t ta' \ final_thread.actions_subset ta' ta" proof(atomize_elim) let ?a_t = "thread_id2addr t" let ?t_a = "addr2thread_id a" from tconf obtain C where ht: "typeof_addr h ?a_t = \Class_type C\" and sub: "P \ C \\<^sup>* Thread" by(fastforce dest: tconfD) show "\ta' va' h'. P,t \ \a\M(vs), h\ -ta'\ext \va', h'\ \ (final_thread.actions_ok final s t ta' \ final_thread.actions_ok' s t ta' \ final_thread.actions_subset ta' ta)" proof(cases "final_thread.actions_ok' s t ta") case True have "final_thread.actions_subset ta ta" by(rule final_thread.actions_subset_refl) with True red show ?thesis by blast next case False note [simp] = final_thread.actions_ok'_iff lock_ok_las'_def final_thread.cond_action_oks'_subset_Join final_thread.actions_subset_iff ta_upd_simps collect_cond_actions_def collect_interrupts_def note [rule del] = subsetI note [intro] = collect_locks'_subset_collect_locks red_external.intros[simplified] show ?thesis proof(cases "wset s t") case [simp]: (Some w) with wst obtain w' where [simp]: "w = PostWS w'" "M = wait" by auto from red have [simp]: "vs = []" by(auto elim: red_external.cases) show ?thesis proof(cases w') case [simp]: WSWokenUp let ?ta' = "\WokenUp, ClearInterrupt t, ObsInterrupted t\" have "final_thread.actions_ok' s t ?ta'" by(simp add: wset_actions_ok_def) moreover have "final_thread.actions_subset ?ta' ta" by(auto simp add: collect_locks'_def finfun_upd_apply) moreover from RedWaitInterrupted have "\va h'. P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by auto ultimately show ?thesis by blast next case [simp]: WSNotified let ?ta' = "\Notified\" have "final_thread.actions_ok' s t ?ta'" by(simp add: wset_actions_ok_def) moreover have "final_thread.actions_subset ?ta' ta" by(auto simp add: collect_locks'_def finfun_upd_apply) moreover from RedWaitNotified have "\va h'. P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by auto ultimately show ?thesis by blast qed next case None from red False show ?thesis proof cases case (RedNewThread C) note ta = \ta = \NewThread ?t_a (C, run, a) h, ThreadStart ?t_a\\ let ?ta' = "\ThreadExists ?t_a True\" from ta False None have "final_thread.actions_ok' s t ?ta'" by(auto) moreover from ta have "final_thread.actions_subset ?ta' ta" by(auto) ultimately show ?thesis using RedNewThread by(fastforce) next case RedNewThreadFail then obtain va' h' x where "P,t \ \a\M(vs), h\ -\NewThread ?t_a x h', ThreadStart ?t_a\\ext \va', h'\" by(fastforce) moreover from \ta = \ThreadExists ?t_a True\\ False None have "final_thread.actions_ok' s t \NewThread ?t_a x h', ThreadStart ?t_a\" by(auto) moreover from \ta = \ThreadExists ?t_a True\\ have "final_thread.actions_subset \NewThread ?t_a x h', ThreadStart ?t_a\ ta" by(auto) ultimately show ?thesis by blast next case RedJoin let ?ta = "\IsInterrupted t True, ClearInterrupt t, ObsInterrupted t\" from \ta = \Join (addr2thread_id a), IsInterrupted t False, ThreadJoin (addr2thread_id a)\\ None False have "t \ interrupts s" by(auto) hence "final_thread.actions_ok final s t ?ta" using None by(auto simp add: final_thread.actions_ok_iff final_thread.cond_action_oks.simps) moreover obtain va h' where "P,t \ \a\M(vs),h\ -?ta\ext \va,h'\" using RedJoinInterrupt RedJoin by auto ultimately show ?thesis by blast next case RedJoinInterrupt hence False using False None by(auto) thus ?thesis .. next case RedInterrupt let ?ta = "\ThreadExists (addr2thread_id a) False\" from RedInterrupt None False have "free_thread_id (thr s) (addr2thread_id a)" by(auto simp add: wset_actions_ok_def) hence "final_thread.actions_ok final s t ?ta" using None by(auto simp add: final_thread.actions_ok_iff final_thread.cond_action_oks.simps) moreover obtain va h' where "P,t \ \a\M(vs),h\ -?ta\ext \va,h'\" using RedInterruptInexist RedInterrupt by auto ultimately show ?thesis by blast next case RedInterruptInexist let ?ta = "\ThreadExists (addr2thread_id a) True, WakeUp (addr2thread_id a), Interrupt (addr2thread_id a), ObsInterrupt (addr2thread_id a)\" from RedInterruptInexist None False have "\ free_thread_id (thr s) (addr2thread_id a)" by(auto simp add: wset_actions_ok_def) hence "final_thread.actions_ok final s t ?ta" using None by(auto simp add: final_thread.actions_ok_iff final_thread.cond_action_oks.simps wset_actions_ok_def) moreover obtain va h' where "P,t \ \a\M(vs),h\ -?ta\ext \va,h'\" using RedInterruptInexist RedInterrupt by auto ultimately show ?thesis by blast next case (RedIsInterruptedTrue C) let ?ta' = "\IsInterrupted ?t_a False\" from RedIsInterruptedTrue False None have "?t_a \ interrupts s" by(auto) hence "final_thread.actions_ok' s t ?ta'" using None by auto moreover from RedIsInterruptedTrue have "final_thread.actions_subset ?ta' ta" by auto moreover from RedIsInterruptedTrue RedIsInterruptedFalse obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by auto ultimately show ?thesis by blast next case (RedIsInterruptedFalse C) let ?ta' = "\IsInterrupted ?t_a True, ObsInterrupted ?t_a\" from RedIsInterruptedFalse have "?t_a \ interrupts s" using False None by(auto) hence "final_thread.actions_ok final s t ?ta'" using None by(auto simp add: final_thread.actions_ok_iff final_thread.cond_action_oks.simps) moreover obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" using RedIsInterruptedFalse RedIsInterruptedTrue by auto ultimately show ?thesis by blast next case RedWaitInterrupt note ta = \ta = \Unlock\a, Lock\a, IsInterrupted t True, ClearInterrupt t, ObsInterrupted t\\ from ta False None have hli: "\ has_lock (locks s $ a) t \ t \ interrupts s" by(fastforce simp add: lock_actions_ok'_iff finfun_upd_apply split: if_split_asm dest: may_lock_t_may_lock_unlock_lock_t dest: has_lock_may_lock) show ?thesis proof(cases "has_lock (locks s $ a) t") case True let ?ta' = "\Suspend a, Unlock\a, Lock\a, ReleaseAcquire\a, IsInterrupted t False, SyncUnlock a \" from True hli have "t \ interrupts s" by simp with True False have "final_thread.actions_ok' s t ?ta'" using None by(auto simp add: lock_actions_ok'_iff finfun_upd_apply wset_actions_ok_def Cons_eq_append_conv) moreover from ta have "final_thread.actions_subset ?ta' ta" by(auto simp add: collect_locks'_def finfun_upd_apply) moreover from RedWait RedWaitInterrupt obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by auto ultimately show ?thesis by blast next case False let ?ta' = "\UnlockFail\a\" from False have "final_thread.actions_ok' s t ?ta'" using None by(auto simp add: lock_actions_ok'_iff finfun_upd_apply) moreover from ta have "final_thread.actions_subset ?ta' ta" by(auto simp add: collect_locks'_def finfun_upd_apply) moreover from RedWaitInterrupt obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by(fastforce) ultimately show ?thesis by blast qed next case RedWait note ta = \ta = \Suspend a, Unlock\a, Lock\a, ReleaseAcquire\a, IsInterrupted t False, SyncUnlock a\\ from ta False None have hli: "\ has_lock (locks s $ a) t \ t \ interrupts s" by(auto simp add: lock_actions_ok'_iff finfun_upd_apply wset_actions_ok_def Cons_eq_append_conv split: if_split_asm dest: may_lock_t_may_lock_unlock_lock_t dest: has_lock_may_lock) show ?thesis proof(cases "has_lock (locks s $ a) t") case True let ?ta' = "\Unlock\a, Lock\a, IsInterrupted t True, ClearInterrupt t, ObsInterrupted t\" from True hli have "t \ interrupts s" by simp with True False have "final_thread.actions_ok final s t ?ta'" using None by(auto simp add: final_thread.actions_ok_iff final_thread.cond_action_oks.simps lock_ok_las_def finfun_upd_apply has_lock_may_lock) moreover from RedWait RedWaitInterrupt obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by auto ultimately show ?thesis by blast next case False let ?ta' = "\UnlockFail\a\" from False have "final_thread.actions_ok' s t ?ta'" using None by(auto simp add: lock_actions_ok'_iff finfun_upd_apply) moreover from ta have "final_thread.actions_subset ?ta' ta" by(auto simp add: collect_locks'_def finfun_upd_apply) moreover from RedWait RedWaitFail obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by(fastforce) ultimately show ?thesis by blast qed next case RedWaitFail note ta = \ta = \UnlockFail\a\\ let ?ta' = "if t \ interrupts s then \Unlock\a, Lock\a, IsInterrupted t True, ClearInterrupt t, ObsInterrupted t\ else \Suspend a, Unlock\a, Lock\a, ReleaseAcquire\a, IsInterrupted t False, SyncUnlock a \" from ta False None have "has_lock (locks s $ a) t" by(auto simp add: finfun_upd_apply split: if_split_asm) hence "final_thread.actions_ok final s t ?ta'" using None by(auto simp add: final_thread.actions_ok_iff final_thread.cond_action_oks.simps lock_ok_las_def finfun_upd_apply has_lock_may_lock wset_actions_ok_def) moreover from RedWaitFail RedWait RedWaitInterrupt obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by(cases "t \ interrupts s") (auto) ultimately show ?thesis by blast next case RedWaitNotified note ta = \ta = \Notified\\ let ?ta' = "if has_lock (locks s $ a) t then (if t \ interrupts s then \Unlock\a, Lock\a, IsInterrupted t True, ClearInterrupt t, ObsInterrupted t\ else \Suspend a, Unlock\a, Lock\a, ReleaseAcquire\a, IsInterrupted t False, SyncUnlock a \) else \UnlockFail\a\" have "final_thread.actions_ok final s t ?ta'" using None by(auto simp add: final_thread.actions_ok_iff final_thread.cond_action_oks.simps lock_ok_las_def finfun_upd_apply has_lock_may_lock wset_actions_ok_def) moreover from RedWaitNotified RedWait RedWaitInterrupt RedWaitFail have "\va h'. P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by auto ultimately show ?thesis by blast next case RedWaitInterrupted note ta = \ta = \WokenUp, ClearInterrupt t, ObsInterrupted t\\ let ?ta' = "if has_lock (locks s $ a) t then (if t \ interrupts s then \Unlock\a, Lock\a, IsInterrupted t True, ClearInterrupt t, ObsInterrupted t\ else \Suspend a, Unlock\a, Lock\a, ReleaseAcquire\a, IsInterrupted t False, SyncUnlock a \) else \UnlockFail\a\" have "final_thread.actions_ok final s t ?ta'" using None by(auto simp add: final_thread.actions_ok_iff final_thread.cond_action_oks.simps lock_ok_las_def finfun_upd_apply has_lock_may_lock wset_actions_ok_def) moreover from RedWaitInterrupted RedWait RedWaitInterrupt RedWaitFail have "\va h'. P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by auto ultimately show ?thesis by blast next case RedWaitSpurious note ta = \ta = \Unlock\a, Lock\a, ReleaseAcquire\a, IsInterrupted t False, SyncUnlock a\\ from ta False None have hli: "\ has_lock (locks s $ a) t \ t \ interrupts s" by(auto simp add: lock_actions_ok'_iff finfun_upd_apply wset_actions_ok_def Cons_eq_append_conv split: if_split_asm dest: may_lock_t_may_lock_unlock_lock_t dest: has_lock_may_lock) show ?thesis proof(cases "has_lock (locks s $ a) t") case True let ?ta' = "\Unlock\a, Lock\a, IsInterrupted t True, ClearInterrupt t, ObsInterrupted t\" from True hli have "t \ interrupts s" by simp with True False have "final_thread.actions_ok final s t ?ta'" using None by(auto simp add: final_thread.actions_ok_iff final_thread.cond_action_oks.simps lock_ok_las_def finfun_upd_apply has_lock_may_lock) moreover from RedWaitInterrupt RedWaitSpurious(1-5) obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by auto ultimately show ?thesis by blast next case False let ?ta' = "\UnlockFail\a\" from False have "final_thread.actions_ok' s t ?ta'" using None by(auto simp add: lock_actions_ok'_iff finfun_upd_apply) moreover from ta have "final_thread.actions_subset ?ta' ta" by(auto simp add: collect_locks'_def finfun_upd_apply) moreover from RedWaitSpurious(1-5) RedWaitFail obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by(fastforce) ultimately show ?thesis by blast qed next case RedNotify note ta = \ta = \Notify a, Unlock\a, Lock\a\\ let ?ta' = "\UnlockFail\a\" from ta False None have "\ has_lock (locks s $ a) t" by(fastforce simp add: lock_actions_ok'_iff finfun_upd_apply wset_actions_ok_def Cons_eq_append_conv split: if_split_asm dest: may_lock_t_may_lock_unlock_lock_t has_lock_may_lock) hence "final_thread.actions_ok' s t ?ta'" using None by(auto simp add: lock_actions_ok'_iff finfun_upd_apply) moreover from ta have "final_thread.actions_subset ?ta' ta" by(auto simp add: collect_locks'_def finfun_upd_apply) moreover from RedNotify obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by(fastforce) ultimately show ?thesis by blast next case RedNotifyFail note ta = \ta = \UnlockFail\a\\ let ?ta' = "\Notify a, Unlock\a, Lock\a\" from ta False None have "has_lock (locks s $ a) t" by(auto simp add: finfun_upd_apply split: if_split_asm) hence "final_thread.actions_ok' s t ?ta'" using None by(auto simp add: finfun_upd_apply simp add: wset_actions_ok_def intro: has_lock_may_lock) moreover from ta have "final_thread.actions_subset ?ta' ta" by(auto simp add: collect_locks'_def finfun_upd_apply) moreover from RedNotifyFail obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by(fastforce) ultimately show ?thesis by blast next case RedNotifyAll note ta = \ta = \NotifyAll a, Unlock\a, Lock\a\\ let ?ta' = "\UnlockFail\a\" from ta False None have "\ has_lock (locks s $ a) t" by(auto simp add: lock_actions_ok'_iff finfun_upd_apply wset_actions_ok_def Cons_eq_append_conv split: if_split_asm dest: may_lock_t_may_lock_unlock_lock_t) hence "final_thread.actions_ok' s t ?ta'" using None by(auto simp add: lock_actions_ok'_iff finfun_upd_apply) moreover from ta have "final_thread.actions_subset ?ta' ta" by(auto simp add: collect_locks'_def finfun_upd_apply) moreover from RedNotifyAll obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by(fastforce) ultimately show ?thesis by blast next case RedNotifyAllFail note ta = \ta = \UnlockFail\a\\ let ?ta' = "\NotifyAll a, Unlock\a, Lock\a\" from ta False None have "has_lock (locks s $ a) t" by(auto simp add: finfun_upd_apply split: if_split_asm) hence "final_thread.actions_ok' s t ?ta'" using None by(auto simp add: finfun_upd_apply wset_actions_ok_def intro: has_lock_may_lock) moreover from ta have "final_thread.actions_subset ?ta' ta" by(auto simp add: collect_locks'_def finfun_upd_apply) moreover from RedNotifyAllFail obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" by(fastforce) ultimately show ?thesis by blast next case RedInterruptedTrue let ?ta' = "\IsInterrupted t False\" from RedInterruptedTrue have "final_thread.actions_ok final s t ?ta'" using None False by(auto simp add: final_thread.actions_ok_iff final_thread.cond_action_oks.simps) moreover obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" using RedInterruptedFalse RedInterruptedTrue by auto ultimately show ?thesis by blast next case RedInterruptedFalse let ?ta' = "\IsInterrupted t True, ClearInterrupt t, ObsInterrupted t\" from RedInterruptedFalse have "final_thread.actions_ok final s t ?ta'" using None False by(auto simp add: final_thread.actions_ok_iff final_thread.cond_action_oks.simps) moreover obtain va h' where "P,t \ \a\M(vs),h\ -?ta'\ext \va,h'\" using RedInterruptedFalse RedInterruptedTrue by auto ultimately show ?thesis by blast qed(auto simp add: None) qed qed qed end context heap_base begin lemma red_external_ta_satisfiable: fixes final assumes "P,t \ \a\M(vs), h\ -ta\ext \va, h'\" shows "\s. final_thread.actions_ok final s t ta" proof - note [simp] = final_thread.actions_ok_iff final_thread.cond_action_oks.simps final_thread.cond_action_ok.simps lock_ok_las_def finfun_upd_apply wset_actions_ok_def has_lock_may_lock and [intro] = free_thread_id.intros and [cong] = conj_cong from assms show ?thesis by cases(fastforce intro: exI[where x="(K$ None)(a $:= \(t, 0)\)"] exI[where x="(K$ None)"])+ qed lemma red_external_aggr_ta_satisfiable: fixes final assumes "(ta, va, h') \ red_external_aggr P t a M vs h" shows "\s. final_thread.actions_ok final s t ta" proof - note [simp] = final_thread.actions_ok_iff final_thread.cond_action_oks.simps final_thread.cond_action_ok.simps lock_ok_las_def finfun_upd_apply wset_actions_ok_def has_lock_may_lock and [intro] = free_thread_id.intros and [cong] = conj_cong from assms show ?thesis by(fastforce simp add: red_external_aggr_def split_beta ta_upd_simps split: if_split_asm intro: exI[where x="Map.empty"] exI[where x="(K$ None)(a $:= \(t, 0)\)"] exI[where x="K$ None"]) qed end subsection \Determinism\ context heap_base begin lemma heap_copy_loc_deterministic: assumes det: "deterministic_heap_ops" and copy: "heap_copy_loc a a' al h ops h'" "heap_copy_loc a a' al h ops' h''" shows "ops = ops' \ h' = h''" using copy by(auto elim!: heap_copy_loc.cases dest: deterministic_heap_ops_readD[OF det] deterministic_heap_ops_writeD[OF det]) lemma heap_copies_deterministic: assumes det: "deterministic_heap_ops" and copy: "heap_copies a a' als h ops h'" "heap_copies a a' als h ops' h''" shows "ops = ops' \ h' = h''" using copy apply(induct arbitrary: ops' h'') apply(fastforce elim!: heap_copies_cases) apply(erule heap_copies_cases) apply clarify apply(drule (1) heap_copy_loc_deterministic[OF det]) apply clarify apply(unfold same_append_eq) apply blast done lemma heap_clone_deterministic: assumes det: "deterministic_heap_ops" and clone: "heap_clone P h a h' obs" "heap_clone P h a h'' obs'" shows "h' = h'' \ obs = obs'" using clone by(auto 4 4 elim!: heap_clone.cases dest: heap_copies_deterministic[OF det] deterministic_heap_ops_allocateD[OF det] has_fields_fun) lemma red_external_deterministic: fixes final assumes det: "deterministic_heap_ops" and red: "P,t \ \a\M(vs), (shr s)\ -ta\ext \va, h'\" "P,t \ \a\M(vs), (shr s)\ -ta'\ext \va', h''\" and aok: "final_thread.actions_ok final s t ta" "final_thread.actions_ok final s t ta'" shows "ta = ta' \ va = va' \ h' = h''" using red aok apply(simp add: final_thread.actions_ok_iff lock_ok_las_def) apply(erule red_external.cases) apply(erule_tac [!] red_external.cases) apply simp_all apply(auto simp add: finfun_upd_apply wset_actions_ok_def dest: heap_clone_deterministic[OF det] split: if_split_asm) using deterministic_heap_ops_no_spurious_wakeups[OF det] apply simp_all done end end diff --git a/thys/JinjaThreads/Compiler/Correctness1.thy b/thys/JinjaThreads/Compiler/Correctness1.thy --- a/thys/JinjaThreads/Compiler/Correctness1.thy +++ b/thys/JinjaThreads/Compiler/Correctness1.thy @@ -1,2385 +1,2385 @@ (* Title: JinjaThreads/Compiler/Correctness1.thy Author: Andreas Lochbihler, Tobias Nipkow reminiscent of the Jinja theory Compiler/Correctness1 *) section \Semantic Correctness of Stage 1\ theory Correctness1 imports J0J1Bisim "../J/DefAssPreservation" begin lemma finals_map_Val [simp]: "finals (map Val vs)" by(simp add: finals_iff) context J_heap_base begin lemma \red0r_preserves_defass: assumes wf: "wf_J_prog P" shows "\ \red0r extTA P t h (e, xs) (e', xs'); \ e \dom xs\ \ \ \ e' \dom xs'\" by(induct rule: rtranclp_induct2)(auto dest: red_preserves_defass[OF wf]) lemma \red0t_preserves_defass: assumes wf: "wf_J_prog P" shows "\ \red0t extTA P t h (e, xs) (e', xs'); \ e \dom xs\ \ \ \ e' \dom xs'\" by(rule \red0r_preserves_defass[OF wf])(rule tranclp_into_rtranclp) end lemma LAss_lem: "\x \ set xs; size xs \ size ys \ \ m1 \\<^sub>m m2(xs[\]ys) \ m1(x\y) \\<^sub>m m2(xs[\]ys[index xs x := y])" apply(simp add:map_le_def) apply(simp add:fun_upds_apply index_less_aux eq_sym_conv) done lemma Block_lem: fixes l :: "'a \ 'b" assumes 0: "l \\<^sub>m [Vs [\] ls]" and 1: "l' \\<^sub>m [Vs [\] ls', V\v]" and hidden: "V \ set Vs \ ls ! index Vs V = ls' ! index Vs V" and size: "size ls = size ls'" "size Vs < size ls'" shows "l'(V := l V) \\<^sub>m [Vs [\] ls']" proof - have "l'(V := l V) \\<^sub>m [Vs [\] ls', V\v](V := l V)" using 1 by(rule map_le_upd) also have "\ = [Vs [\] ls'](V := l V)" by simp also have "\ \\<^sub>m [Vs [\] ls']" proof (cases "l V") case None thus ?thesis by simp next case (Some w) hence "[Vs [\] ls] V = Some w" using 0 by(force simp add: map_le_def split:if_splits) hence VinVs: "V \ set Vs" and w: "w = ls ! index Vs V" using size by(auto simp add:fun_upds_apply split:if_splits) hence "w = ls' ! index Vs V" using hidden[OF VinVs] by simp hence "[Vs [\] ls'](V := l V) = [Vs [\] ls']" using Some size VinVs by(simp add:index_less_aux map_upds_upd_conv_index) thus ?thesis by simp qed finally show ?thesis . qed subsection \Correctness proof\ locale J0_J1_heap_base = J?: J_heap_base + J1?: J1_heap_base + constrains addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" begin lemma ta_bisim01_extTA2J0_extTA2J1: assumes wf: "wf_J_prog P" and nt: "\n T C M a h. \ n < length \ta\\<^bsub>t\<^esub>; \ta\\<^bsub>t\<^esub> ! n = NewThread T (C, M, a) h \ \ typeof_addr h a = \Class_type C\ \ (\T meth D. P \ C sees M:[]\T =\meth\ in D)" shows "ta_bisim01 (extTA2J0 P ta) (extTA2J1 (compP1 P) ta)" apply(simp add: ta_bisim_def ta_upd_simps) apply(auto intro!: list_all2_all_nthI) apply(case_tac "\ta\\<^bsub>t\<^esub> ! n") apply(auto simp add: bisim_red0_Red1_def) apply(drule (1) nt) apply(clarify) apply(erule bisim_list_extTA2J0_extTA2J1[OF wf, simplified]) done lemma red_external_ta_bisim01: "\ wf_J_prog P; P,t \ \a\M(vs), h\ -ta\ext \va, h'\ \ \ ta_bisim01 (extTA2J0 P ta) (extTA2J1 (compP1 P) ta)" apply(rule ta_bisim01_extTA2J0_extTA2J1, assumption) apply(drule (1) red_external_new_thread_sees, auto simp add: in_set_conv_nth) apply(drule red_ext_new_thread_heap, auto simp add: in_set_conv_nth) done lemmas \red1t_expr = NewArray_\red1t_xt Cast_\red1t_xt InstanceOf_\red1t_xt BinOp_\red1t_xt1 BinOp_\red1t_xt2 LAss_\red1t AAcc_\red1t_xt1 AAcc_\red1t_xt2 AAss_\red1t_xt1 AAss_\red1t_xt2 AAss_\red1t_xt3 ALength_\red1t_xt FAcc_\red1t_xt FAss_\red1t_xt1 FAss_\red1t_xt2 CAS_\red1t_xt1 CAS_\red1t_xt2 CAS_\red1t_xt3 Call_\red1t_obj Call_\red1t_param Block_None_\red1t_xt Block_\red1t_Some Sync_\red1t_xt InSync_\red1t_xt Seq_\red1t_xt Cond_\red1t_xt Throw_\red1t_xt Try_\red1t_xt lemmas \red1r_expr = NewArray_\red1r_xt Cast_\red1r_xt InstanceOf_\red1r_xt BinOp_\red1r_xt1 BinOp_\red1r_xt2 LAss_\red1r AAcc_\red1r_xt1 AAcc_\red1r_xt2 AAss_\red1r_xt1 AAss_\red1r_xt2 AAss_\red1r_xt3 ALength_\red1r_xt FAcc_\red1r_xt FAss_\red1r_xt1 FAss_\red1r_xt2 CAS_\red1r_xt1 CAS_\red1r_xt2 CAS_\red1r_xt3 Call_\red1r_obj Call_\red1r_param Block_None_\red1r_xt Block_\red1r_Some Sync_\red1r_xt InSync_\red1r_xt Seq_\red1r_xt Cond_\red1r_xt Throw_\red1r_xt Try_\red1r_xt definition sim_move01 :: "'addr J1_prog \ 'thread_id \ ('addr, 'thread_id, 'heap) J0_thread_action \ 'addr expr \ 'addr expr1 \ 'heap \ 'addr locals1 \ ('addr, 'thread_id, 'heap) external_thread_action \ 'addr expr1 \ 'heap \ 'addr locals1 \ bool" where "sim_move01 P t ta0 e0 e h xs ta e' h' xs' \ \ final e0 \ (if \move0 P h e0 then h' = h \ ta0 = \ \ ta = \ \ \red1't P t h (e, xs) (e', xs') else ta_bisim01 ta0 (extTA2J1 P ta) \ (if call e0 = None \ call1 e = None then (\e'' xs''. \red1'r P t h (e, xs) (e'', xs'') \ False,P,t \1 \e'', (h, xs'')\ -ta\ \e', (h', xs')\ \ \ \move1 P h e'') else False,P,t \1 \e, (h, xs)\ -ta\ \e', (h', xs')\ \ \ \move1 P h e))" definition sim_moves01 :: "'addr J1_prog \ 'thread_id \ ('addr, 'thread_id, 'heap) J0_thread_action \ 'addr expr list \ 'addr expr1 list \ 'heap \ 'addr locals1 \ ('addr, 'thread_id, 'heap) external_thread_action \ 'addr expr1 list \ 'heap \ 'addr locals1 \ bool" where "sim_moves01 P t ta0 es0 es h xs ta es' h' xs' \ \ finals es0 \ (if \moves0 P h es0 then h' = h \ ta0 = \ \ ta = \ \ \reds1't P t h (es, xs) (es', xs') else ta_bisim01 ta0 (extTA2J1 P ta) \ (if calls es0 = None \ calls1 es = None then (\es'' xs''. \reds1'r P t h (es, xs) (es'', xs'') \ False,P,t \1 \es'', (h, xs'')\ [-ta\] \es', (h', xs')\ \ \ \moves1 P h es'') else False,P,t \1 \es, (h, xs)\ [-ta\] \es', (h', xs')\ \ \ \moves1 P h es))" declare \red1t_expr [elim!] \red1r_expr[elim!] lemma sim_move01_expr: assumes "sim_move01 P t ta0 e0 e h xs ta e' h' xs'" shows "sim_move01 P t ta0 (newA T\e0\) (newA T\e\) h xs ta (newA T\e'\) h' xs'" "sim_move01 P t ta0 (Cast T e0) (Cast T e) h xs ta (Cast T e') h' xs'" "sim_move01 P t ta0 (e0 instanceof T) (e instanceof T) h xs ta (e' instanceof T) h' xs'" "sim_move01 P t ta0 (e0 \bop\ e2) (e \bop\ e2') h xs ta (e' \bop\ e2') h' xs'" "sim_move01 P t ta0 (Val v \bop\ e0) (Val v \bop\ e) h xs ta (Val v \bop\ e') h' xs'" "sim_move01 P t ta0 (V := e0) (V' := e) h xs ta (V' := e') h' xs'" "sim_move01 P t ta0 (e0\e2\) (e\e2'\) h xs ta (e'\e2'\) h' xs'" "sim_move01 P t ta0 (Val v\e0\) (Val v\e\) h xs ta (Val v\e'\) h' xs'" "sim_move01 P t ta0 (e0\e2\ := e3) (e\e2'\ := e3') h xs ta (e'\e2'\ := e3') h' xs'" "sim_move01 P t ta0 (Val v\e0\ := e3) (Val v\e\ := e3') h xs ta (Val v\e'\ := e3') h' xs'" "sim_move01 P t ta0 (AAss (Val v) (Val v') e0) (AAss (Val v) (Val v') e) h xs ta (AAss (Val v) (Val v') e') h' xs'" "sim_move01 P t ta0 (e0\length) (e\length) h xs ta (e'\length) h' xs'" "sim_move01 P t ta0 (e0\F{D}) (e\F'{D'}) h xs ta (e'\F'{D'}) h' xs'" "sim_move01 P t ta0 (FAss e0 F D e2) (FAss e F' D' e2') h xs ta (FAss e' F' D' e2') h' xs'" "sim_move01 P t ta0 (FAss (Val v) F D e0) (FAss (Val v) F' D' e) h xs ta (FAss (Val v) F' D' e') h' xs'" "sim_move01 P t ta0 (CompareAndSwap e0 D F e2 e3) (CompareAndSwap e D F e2' e3') h xs ta (CompareAndSwap e' D F e2' e3') h' xs'" "sim_move01 P t ta0 (CompareAndSwap (Val v) D F e0 e3) (CompareAndSwap (Val v) D F e e3') h xs ta (CompareAndSwap (Val v) D F e' e3') h' xs'" "sim_move01 P t ta0 (CompareAndSwap (Val v) D F (Val v') e0) (CompareAndSwap (Val v) D F (Val v') e) h xs ta (CompareAndSwap (Val v) D F (Val v') e') h' xs'" "sim_move01 P t ta0 (e0\M(es)) (e\M(es')) h xs ta (e'\M(es')) h' xs'" "sim_move01 P t ta0 ({V:T=vo; e0}) ({V':T=None; e}) h xs ta ({V':T=None; e'}) h' xs'" "sim_move01 P t ta0 (sync(e0) e2) (sync\<^bsub>V'\<^esub>(e) e2') h xs ta (sync\<^bsub>V'\<^esub>(e') e2') h' xs'" "sim_move01 P t ta0 (insync(a) e0) (insync\<^bsub>V'\<^esub>(a') e) h xs ta (insync\<^bsub>V'\<^esub>(a') e') h' xs'" "sim_move01 P t ta0 (e0;;e2) (e;;e2') h xs ta (e';;e2') h' xs'" "sim_move01 P t ta0 (if (e0) e2 else e3) (if (e) e2' else e3') h xs ta (if (e') e2' else e3') h' xs'" "sim_move01 P t ta0 (throw e0) (throw e) h xs ta (throw e') h' xs'" "sim_move01 P t ta0 (try e0 catch(C V) e2) (try e catch(C' V') e2') h xs ta (try e' catch(C' V') e2') h' xs'" using assms apply(simp_all add: sim_move01_def final_iff \red1r_Val \red1t_Val split: if_split_asm split del: if_split) apply(fastforce simp add: final_iff \red1r_Val \red1t_Val split!: if_splits intro: red1_reds1.intros)+ done lemma sim_moves01_expr: "sim_move01 P t ta0 e0 e h xs ta e' h' xs' \ sim_moves01 P t ta0 (e0 # es2) (e # es2') h xs ta (e' # es2') h' xs'" "sim_moves01 P t ta0 es0 es h xs ta es' h' xs' \ sim_moves01 P t ta0 (Val v # es0) (Val v # es) h xs ta (Val v # es') h' xs'" apply(simp_all add: sim_move01_def sim_moves01_def final_iff finals_iff Cons_eq_append_conv \red1t_Val \red1r_Val split: if_split_asm split del: if_split) apply(auto simp add: Cons_eq_append_conv \red1t_Val \red1r_Val split!: if_splits intro: List1Red1 List1Red2 \red1t_inj_\reds1t \red1r_inj_\reds1r \reds1t_cons_\reds1t \reds1r_cons_\reds1r) apply(force elim!: \red1r_inj_\reds1r List1Red1) apply(force elim!: \red1r_inj_\reds1r List1Red1) apply(force elim!: \red1r_inj_\reds1r List1Red1) apply(force elim!: \red1r_inj_\reds1r List1Red1) apply(force elim!: \reds1r_cons_\reds1r intro!: List1Red2) apply(force elim!: \reds1r_cons_\reds1r intro!: List1Red2) done lemma sim_move01_CallParams: "sim_moves01 P t ta0 es0 es h xs ta es' h' xs' \ sim_move01 P t ta0 (Val v\M(es0)) (Val v\M(es)) h xs ta (Val v\M(es')) h' xs'" apply(clarsimp simp add: sim_move01_def sim_moves01_def \reds1r_map_Val \reds1t_map_Val is_vals_conv split: if_split_asm split del: if_split) apply(fastforce simp add: sim_move01_def sim_moves01_def \reds1r_map_Val \reds1t_map_Val intro: Call_\red1r_param Call1Params) apply(rule conjI, fastforce) apply(split if_split) apply(rule conjI) apply(clarsimp simp add: finals_iff) apply(clarify) apply(split if_split) apply(rule conjI) apply(simp del: call.simps calls.simps call1.simps calls1.simps) apply(fastforce simp add: sim_move01_def sim_moves01_def \red1r_Val \red1t_Val \reds1r_map_Val_Throw intro: Call_\red1r_param Call1Params split: if_split_asm) apply(fastforce split: if_split_asm simp add: is_vals_conv \reds1r_map_Val \reds1r_map_Val_Throw) apply(rule conjI, fastforce) apply(fastforce simp add: sim_move01_def sim_moves01_def \red1r_Val \red1t_Val \reds1t_map_Val \reds1r_map_Val is_vals_conv intro: Call_\red1r_param Call1Params split: if_split_asm) done lemma sim_move01_reds: "\ (h', a) \ allocate h (Class_type C); ta0 = \NewHeapElem a (Class_type C)\; ta = \NewHeapElem a (Class_type C)\ \ \ sim_move01 P t ta0 (new C) (new C) h xs ta (addr a) h' xs" "allocate h (Class_type C) = {} \ sim_move01 P t \ (new C) (new C) h xs \ (THROW OutOfMemory) h xs" "\ (h', a) \ allocate h (Array_type T (nat (sint i))); 0 <=s i; ta0 = \NewHeapElem a (Array_type T (nat (sint i)))\; ta = \NewHeapElem a (Array_type T (nat (sint i)))\ \ \ sim_move01 P t ta0 (newA T\Val (Intg i)\) (newA T\Val (Intg i)\) h xs ta (addr a) h' xs" "i sim_move01 P t \ (newA T\Val (Intg i)\) (newA T\Val (Intg i)\) h xs \ (THROW NegativeArraySize) h xs" "\ allocate h (Array_type T (nat (sint i))) = {}; 0 <=s i \ \ sim_move01 P t \ (newA T\Val (Intg i)\) (newA T\Val (Intg i)\) h xs \ (THROW OutOfMemory) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; P \ U \ T \ \ sim_move01 P t \ (Cast T (Val v)) (Cast T (Val v)) h xs \ (Val v) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; \ P \ U \ T \ \ sim_move01 P t \ (Cast T (Val v)) (Cast T (Val v)) h xs \ (THROW ClassCast) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; b \ v \ Null \ P \ U \ T \ \ sim_move01 P t \ ((Val v) instanceof T) ((Val v) instanceof T) h xs \ (Val (Bool b)) h xs" "binop bop v1 v2 = Some (Inl v) \ sim_move01 P t \ ((Val v1) \bop\ (Val v2)) (Val v1 \bop\ Val v2) h xs \ (Val v) h xs" "binop bop v1 v2 = Some (Inr a) \ sim_move01 P t \ ((Val v1) \bop\ (Val v2)) (Val v1 \bop\ Val v2) h xs \ (Throw a) h xs" "\ xs!V = v; V < size xs \ \ sim_move01 P t \ (Var V') (Var V) h xs \ (Val v) h xs" "V < length xs \ sim_move01 P t \ (V' := Val v) (V := Val v) h xs \ unit h (xs[V := v])" "sim_move01 P t \ (null\Val v\) (null\Val v\) h xs \ (THROW NullPointer) h xs" "\ typeof_addr h a = \Array_type T n\; i sint i \ int n \ \ sim_move01 P t \ (addr a\Val (Intg i)\) ((addr a)\Val (Intg i)\) h xs \ (THROW ArrayIndexOutOfBounds) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; heap_read h a (ACell (nat (sint i))) v; ta0 = \ReadMem a (ACell (nat (sint i))) v\; ta = \ReadMem a (ACell (nat (sint i))) v\ \ \ sim_move01 P t ta0 (addr a\Val (Intg i)\) ((addr a)\Val (Intg i)\) h xs ta (Val v) h xs" "sim_move01 P t \ (null\Val v\ := Val v') (null\Val v\ := Val v') h xs \ (THROW NullPointer) h xs" "\ typeof_addr h a = \Array_type T n\; i sint i \ int n \ \ sim_move01 P t \ (AAss (addr a) (Val (Intg i)) (Val v)) (AAss (addr a) (Val (Intg i)) (Val v)) h xs \ (THROW ArrayIndexOutOfBounds) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>h\<^esub> v = \U\; \ (P \ U \ T) \ \ sim_move01 P t \ (AAss (addr a) (Val (Intg i)) (Val v)) (AAss (addr a) (Val (Intg i)) (Val v)) h xs \ (THROW ArrayStore) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>h\<^esub> v = Some U; P \ U \ T; heap_write h a (ACell (nat (sint i))) v h'; ta0 = \WriteMem a (ACell (nat (sint i))) v\; ta = \WriteMem a (ACell (nat (sint i))) v \ \ \ sim_move01 P t ta0 (AAss (addr a) (Val (Intg i)) (Val v)) (AAss (addr a) (Val (Intg i)) (Val v)) h xs ta unit h' xs" "typeof_addr h a = \Array_type T n\ \ sim_move01 P t \ (addr a\length) (addr a\length) h xs \ (Val (Intg (word_of_int (int n)))) h xs" "sim_move01 P t \ (null\length) (null\length) h xs \ (THROW NullPointer) h xs" "\ heap_read h a (CField D F) v; ta0 = \ReadMem a (CField D F) v\; ta = \ReadMem a (CField D F) v\ \ \ sim_move01 P t ta0 (addr a\F{D}) (addr a\F{D}) h xs ta (Val v) h xs" "sim_move01 P t \ (null\F{D}) (null\F{D}) h xs \ (THROW NullPointer) h xs" "\ heap_write h a (CField D F) v h'; ta0 = \WriteMem a (CField D F) v\; ta = \WriteMem a (CField D F) v\ \ \ sim_move01 P t ta0 (addr a\F{D} := Val v) (addr a\F{D} := Val v) h xs ta unit h' xs" "sim_move01 P t \ (null\compareAndSwap(D\F, Val v, Val v')) (null\compareAndSwap(D\F, Val v, Val v')) h xs \ (THROW NullPointer) h xs" "\ heap_read h a (CField D F) v''; heap_write h a (CField D F) v' h'; v'' = v; ta0 = \ReadMem a (CField D F) v'', WriteMem a (CField D F) v'\; ta = \ReadMem a (CField D F) v'', WriteMem a (CField D F) v'\ \ \ sim_move01 P t ta0 (addr a\compareAndSwap(D\F, Val v, Val v')) (addr a\compareAndSwap(D\F, Val v, Val v')) h xs ta true h' xs" "\ heap_read h a (CField D F) v''; v'' \ v; ta0 = \ReadMem a (CField D F) v''\; ta = \ReadMem a (CField D F) v''\ \ \ sim_move01 P t ta0 (addr a\compareAndSwap(D\F, Val v, Val v')) (addr a\compareAndSwap(D\F, Val v, Val v')) h xs ta false h xs" "sim_move01 P t \ (null\F{D} := Val v) (null\F{D} := Val v) h xs \ (THROW NullPointer) h xs" "sim_move01 P t \ ({V':T=vo; Val u}) ({V:T=None; Val u}) h xs \ (Val u) h xs" "V < length xs \ sim_move01 P t \ (sync(null) e0) (sync\<^bsub>V\<^esub> (null) e1) h xs \ (THROW NullPointer) h (xs[V := Null])" "sim_move01 P t \ (Val v;;e0) (Val v;; e1) h xs \ e1 h xs" "sim_move01 P t \ (if (true) e0 else e0') (if (true) e1 else e1') h xs \ e1 h xs" "sim_move01 P t \ (if (false) e0 else e0') (if (false) e1 else e1') h xs \ e1' h xs" "sim_move01 P t \ (throw null) (throw null) h xs \ (THROW NullPointer) h xs" "sim_move01 P t \ (try (Val v) catch(C V') e0) (try (Val v) catch(C V) e1) h xs \ (Val v) h xs" "\ typeof_addr h a = \Class_type D\; P \ D \\<^sup>* C; V < length xs \ \ sim_move01 P t \ (try (Throw a) catch(C V') e0) (try (Throw a) catch(C V) e1) h xs \ ({V:Class C=None; e1}) h (xs[V := Addr a])" "sim_move01 P t \ (newA T\Throw a\) (newA T\Throw a\) h xs \ (Throw a) h xs" "sim_move01 P t \ (Cast T (Throw a)) (Cast T (Throw a)) h xs \ (Throw a) h xs" "sim_move01 P t \ ((Throw a) instanceof T) ((Throw a) instanceof T) h xs \ (Throw a) h xs" "sim_move01 P t \ ((Throw a) \bop\ e0) ((Throw a) \bop\ e1) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v \bop\ (Throw a)) (Val v \bop\ (Throw a)) h xs \ (Throw a) h xs" "sim_move01 P t \ (V' := Throw a) (V := Throw a) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\e0\) (Throw a\e1\) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\Throw a\) (Val v\Throw a\) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\e0\ := e0') (Throw a\e1\ := e1') h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\Throw a\ := e0) (Val v\Throw a\ := e1) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\Val v'\ := Throw a) (Val v\Val v'\ := Throw a) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\length) (Throw a\length) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\F{D}) (Throw a\F{D}) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\F{D} := e0) (Throw a\F{D} := e1) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\F{D} := Throw a) (Val v\F{D} := Throw a) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\compareAndSwap(D\F, e2, e3)) (Throw a\compareAndSwap(D\F, e2', e3')) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\compareAndSwap(D\F, Throw a, e3)) (Val v\compareAndSwap(D\F, Throw a, e3')) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\compareAndSwap(D\F, Val v', Throw a)) (Val v\compareAndSwap(D\F, Val v', Throw a)) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\M(es0)) (Throw a\M(es1)) h xs \ (Throw a) h xs" "sim_move01 P t \ ({V':T=vo; Throw a}) ({V:T=None; Throw a}) h xs \ (Throw a) h xs" "sim_move01 P t \ (sync(Throw a) e0) (sync\<^bsub>V\<^esub>(Throw a) e1) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a;;e0) (Throw a;;e1) h xs \ (Throw a) h xs" "sim_move01 P t \ (if (Throw a) e0 else e0') (if (Throw a) e1 else e1') h xs \ (Throw a) h xs" "sim_move01 P t \ (throw (Throw a)) (throw (Throw a)) h xs \ (Throw a) h xs" apply(simp_all add: sim_move01_def ta_bisim_def split: if_split_asm split del: if_split) apply(fastforce intro: red1_reds1.intros)+ done lemma sim_move01_ThrowParams: "sim_move01 P t \ (Val v\M(map Val vs @ Throw a # es0)) (Val v\M(map Val vs @ Throw a # es1)) h xs \ (Throw a) h xs" apply(simp add: sim_move01_def split del: if_split) apply(rule conjI, fastforce) apply(split if_split) apply(rule conjI) apply(fastforce intro: red1_reds1.intros) apply(fastforce simp add: sim_move01_def intro: red1_reds1.intros) done lemma sim_move01_CallNull: "sim_move01 P t \ (null\M(map Val vs)) (null\M(map Val vs)) h xs \ (THROW NullPointer) h xs" by(fastforce simp add: sim_move01_def map_eq_append_conv intro: red1_reds1.intros) lemma sim_move01_SyncLocks: "\ V < length xs; ta0 = \Lock\a, SyncLock a\; ta = \Lock\a, SyncLock a\ \ \ sim_move01 P t ta0 (sync(addr a) e0) (sync\<^bsub>V\<^esub> (addr a) e1) h xs ta (insync\<^bsub>V\<^esub> (a) e1) h (xs[V := Addr a])" "\ xs ! V = Addr a'; V < length xs; ta0 = \Unlock\a', SyncUnlock a'\; ta = \Unlock\a', SyncUnlock a'\ \ \ sim_move01 P t ta0 (insync(a') (Val v)) (insync\<^bsub>V\<^esub> (a) (Val v)) h xs ta (Val v) h xs" "\ xs ! V = Addr a'; V < length xs; ta0 = \Unlock\a', SyncUnlock a'\; ta = \Unlock\a', SyncUnlock a'\ \ \ sim_move01 P t ta0 (insync(a') (Throw a'')) (insync\<^bsub>V\<^esub> (a) (Throw a'')) h xs ta (Throw a'') h xs" by(fastforce simp add: sim_move01_def ta_bisim_def expand_finfun_eq fun_eq_iff finfun_upd_apply ta_upd_simps intro: red1_reds1.intros[simplified] split: if_split_asm)+ lemma sim_move01_TryFail: "\ typeof_addr h a = \Class_type D\; \ P \ D \\<^sup>* C \ \ sim_move01 P t \ (try (Throw a) catch(C V') e0) (try (Throw a) catch(C V) e1) h xs \ (Throw a) h xs" by(auto simp add: sim_move01_def intro!: Red1TryFail) lemma sim_move01_BlockSome: "\ sim_move01 P t ta0 e0 e h (xs[V := v]) ta e' h' xs'; V < length xs \ \ sim_move01 P t ta0 ({V':T=\v\; e0}) ({V:T=\v\; e}) h xs ta ({V:T=None; e'}) h' xs'" "V < length xs \ sim_move01 P t \ ({V':T=\v\; Val u}) ({V:T=\v\; Val u}) h xs \ (Val u) h (xs[V := v])" "V < length xs \ sim_move01 P t \ ({V':T=\v\; Throw a}) ({V:T=\v\; Throw a}) h xs \ (Throw a) h (xs[V := v])" apply(auto simp add: sim_move01_def) apply(split if_split_asm) apply(fastforce intro: intro: converse_rtranclp_into_rtranclp Block1Some Block1Red Block_\red1r_Some) apply(fastforce intro: intro: converse_rtranclp_into_rtranclp Block1Some Block1Red Block_\red1r_Some) apply(fastforce simp add: sim_move01_def intro!: \red1t_2step[OF Block1Some] \red1r_1step[OF Block1Some] Red1Block Block1Throw)+ done lemmas sim_move01_intros = sim_move01_expr sim_move01_reds sim_move01_ThrowParams sim_move01_CallNull sim_move01_TryFail sim_move01_BlockSome sim_move01_CallParams declare sim_move01_intros[intro] lemma sim_move01_preserves_len: "sim_move01 P t ta0 e0 e h xs ta e' h' xs' \ length xs' = length xs" by(fastforce simp add: sim_move01_def split: if_split_asm dest: \red1r_preserves_len \red1t_preserves_len red1_preserves_len) lemma sim_move01_preserves_unmod: "\ sim_move01 P t ta0 e0 e h xs ta e' h' xs'; unmod e i; i < length xs \ \ xs' ! i = xs ! i" apply(auto simp add: sim_move01_def split: if_split_asm dest: \red1t_preserves_unmod) apply(frule (2) \red1'r_preserves_unmod) apply(frule (1) \red1r_unmod_preserved) apply(frule \red1r_preserves_len) apply(auto dest: red1_preserves_unmod) apply(frule (2) \red1'r_preserves_unmod) apply(frule (1) \red1r_unmod_preserved) apply(frule \red1r_preserves_len) apply(auto dest: red1_preserves_unmod) done lemma assumes wf: "wf_J_prog P" shows red1_simulates_red_aux: "\ extTA2J0 P,P,t \ \e1, S\ -TA\ \e1', S'\; bisim vs e1 e2 XS; fv e1 \ set vs; lcl S \\<^sub>m [vs [\] XS]; length vs + max_vars e1 \ length XS; \aMvs. call e1 = \aMvs\ \ synthesized_call P (hp S) aMvs \ \ \ta e2' XS'. sim_move01 (compP1 P) t TA e1 e2 (hp S) XS ta e2' (hp S') XS' \ bisim vs e1' e2' XS' \ lcl S' \\<^sub>m [vs [\] XS']" (is "\ _; _; _; _; _; ?synth e1 S \ \ ?concl e1 e2 S XS e1' S' TA vs") and reds1_simulates_reds_aux: "\ extTA2J0 P,P,t \ \es1, S\ [-TA\] \es1', S'\; bisims vs es1 es2 XS; fvs es1 \ set vs; lcl S \\<^sub>m [vs [\] XS]; length vs + max_varss es1 \ length XS; \aMvs. calls es1 = \aMvs\ \ synthesized_call P (hp S) aMvs \ \ \ta es2' xs'. sim_moves01 (compP1 P) t TA es1 es2 (hp S) XS ta es2' (hp S') xs' \ bisims vs es1' es2' xs' \ lcl S' \\<^sub>m [vs [\] xs']" (is "\ _; _; _; _; _; ?synths es1 S \ \ ?concls es1 es2 S XS es1' S' TA vs") proof(induct arbitrary: vs e2 XS and vs es2 XS rule: red_reds.inducts) case (BinOpRed1 e s ta e' s' bop e2 Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs (e \bop\ e2) E2 xs\ obtain E where "E2 = E \bop\ compE1 Vs e2" and "bisim Vs e E xs" and "\ contains_insync e2" by auto with IH[of Vs E xs] \fv (e \bop\ e2) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val e\ \length Vs + max_vars (e \bop\ e2) \ length xs\ \?synth (e \bop\ e2) s\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ show ?case by(cases "is_val e'")(fastforce elim!: sim_move01_expr)+ next case (BinOpRed2 e s ta e' s' v bop Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \bisim Vs (Val v \bop\ e) E2 xs\ obtain E where "E2 = Val v \bop\ E" and "bisim Vs e E xs" by auto with IH[of Vs E xs] \fv (Val v \bop\ e) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars (Val v \bop\ e) \ length xs\ \?synth (Val v \bop\ e) s\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ show ?case by(fastforce elim!: sim_move01_expr) next case RedVar thus ?case by(fastforce simp add: index_less_aux map_le_def fun_upds_apply intro!: exI dest: bspec) next case RedLAss thus ?case by(fastforce intro: index_less_aux LAss_lem intro!: exI simp del: fun_upd_apply) next case (AAccRed1 a s ta a' s' i Vs E2 xs) note IH = \\vs e2 XS. \bisim vs a e2 XS; fv a \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars a \ length XS; ?synth a s \ \ ?concl a e2 s XS a' s' ta vs\ from \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ have "\ is_val a" by auto with \bisim Vs (a\i\) E2 xs\ obtain E where "E2 = E\compE1 Vs i\" and "bisim Vs a E xs" and "\ contains_insync i" by auto with IH[of Vs E xs] \fv (a\i\) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val a\ \length Vs + max_vars (a\i\) \ length xs\ \?synth (a\i\) s\ \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ show ?case by(cases "is_val a'")(fastforce elim!: sim_move01_expr)+ next case (AAccRed2 i s ta i' s' a Vs E2 xs) note IH = \\vs e2 XS. \bisim vs i e2 XS; fv i \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars i \ length XS; ?synth i s \ \ ?concl i e2 s XS i' s' ta vs\ from \bisim Vs (Val a\i\) E2 xs\ obtain E where "E2 = Val a\E\" and "bisim Vs i E xs" by auto with IH[of Vs E xs] \fv (Val a\i\) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars (Val a\i\) \ length xs\ \?synth (Val a\i\) s\ \extTA2J0 P,P,t \ \i,s\ -ta\ \i',s'\\ show ?case by(fastforce elim!: sim_move01_expr) next case RedAAcc thus ?case by(auto simp del: split_paired_Ex) next case (AAssRed1 a s ta a' s' i e Vs E2 xs) note IH = \\vs e2 XS. \bisim vs a e2 XS; fv a \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars a \ length XS; ?synth a s \ \ ?concl a e2 s XS a' s' ta vs\ from \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ have "\ is_val a" by auto with \bisim Vs (a\i\:=e) E2 xs\ obtain E where E2: "E2 = E\compE1 Vs i\:=compE1 Vs e" and "bisim Vs a E xs" and sync: "\ contains_insync i" "\ contains_insync e" by auto with IH[of Vs E xs] \fv (a\i\:=e) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val a\ \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ \length Vs + max_vars (a\i\:=e) \ length xs\ \?synth (a\i\:=e) s\ obtain ta' e2' xs' where IH': "sim_move01 (compP1 P) t ta a E (hp s) xs ta' e2' (hp s') xs'" "bisim Vs a' e2' xs'" "lcl s' \\<^sub>m [Vs [\] xs']" by auto show ?case proof(cases "is_val a'") case True from \fv (a\i\:=e) \ set Vs\ sync have "bisim Vs i (compE1 Vs i) xs'" "bisim Vs e (compE1 Vs e) xs'" by auto with IH' E2 True sync \\ is_val a\ \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ show ?thesis by(cases "is_val i")(fastforce elim!: sim_move01_expr)+ next case False with IH' E2 sync \\ is_val a\ \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ show ?thesis by(fastforce elim!: sim_move01_expr) qed next case (AAssRed2 i s ta i' s' a e Vs E2 xs) note IH = \\vs e2 XS. \bisim vs i e2 XS; fv i \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars i \ length XS; ?synth i s \ \ ?concl i e2 s XS i' s' ta vs\ from \extTA2J0 P,P,t \ \i,s\ -ta\ \i',s'\\ have "\ is_val i" by auto with \bisim Vs (Val a\i\ := e) E2 xs\ obtain E where "E2 = Val a\E\:=compE1 Vs e" and "bisim Vs i E xs" and "\ contains_insync e" by auto with IH[of Vs E xs] \fv (Val a\i\:=e) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val i\ \extTA2J0 P,P,t \ \i,s\ -ta\ \i',s'\\ \length Vs + max_vars (Val a\i\:=e) \ length xs\ \?synth (Val a\i\:=e) s\ show ?case by(cases "is_val i'")(fastforce elim!: sim_move01_expr)+ next case (AAssRed3 e s ta e' s' a i Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \bisim Vs (Val a\Val i\ := e) E2 xs\ obtain E where "E2 = Val a\Val i\:=E" and "bisim Vs e E xs" by auto with IH[of Vs E xs] \fv (Val a\Val i\:=e) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars (Val a\Val i\:=e) \ length xs\ \?synth (Val a\Val i\:=e) s\ show ?case by(fastforce elim!: sim_move01_expr) next case RedAAssStore thus ?case by(auto intro!: exI) next case (FAssRed1 e s ta e' s' F D e2 Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs (e\F{D} := e2) E2 xs\ obtain E where "E2 = E\F{D} := compE1 Vs e2" and "bisim Vs e E xs" and "\ contains_insync e2" by auto with IH[of Vs E xs] \fv (e\F{D} := e2) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val e\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars (e\F{D} := e2) \ length xs\ \?synth (e\F{D} := e2) s\ show ?case by(cases "is_val e'")(fastforce elim!: sim_move01_expr)+ next case (FAssRed2 e s ta e' s' v F D Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \bisim Vs (Val v\F{D} := e) E2 xs\ obtain E where "E2 = Val v\F{D} := E" and "bisim Vs e E xs" by auto with IH[of Vs E xs] \fv (Val v\F{D} := e) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars (Val v\F{D} := e) \ length xs\ \?synth (Val v\F{D} := e) s\ show ?case by(fastforce elim!: sim_move01_expr) next case (CASRed1 e s ta e' s' D F e2 e3 Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs _ E2 xs\ obtain E where E2: "E2 = E\compareAndSwap(D\F, compE1 Vs e2, compE1 Vs e3)" and "bisim Vs e E xs" and sync: "\ contains_insync e2" "\ contains_insync e3" by(auto) with IH[of Vs E xs] \fv _ \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val e\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars _ \ length xs\ \?synth _ s\ obtain ta' e2' xs' where IH': "sim_move01 (compP1 P) t ta e E (hp s) xs ta' e2' (hp s') xs'" "bisim Vs e' e2' xs'" "lcl s' \\<^sub>m [Vs [\] xs']" by auto show ?case proof(cases "is_val e'") case True from \fv _ \ set Vs\ sync have "bisim Vs e2 (compE1 Vs e2) xs'" "bisim Vs e3 (compE1 Vs e3) xs'" by auto with IH' E2 True sync \\ is_val e\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ show ?thesis by(cases "is_val e2")(fastforce elim!: sim_move01_expr)+ next case False with IH' E2 sync \\ is_val e\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ show ?thesis by(fastforce elim!: sim_move01_expr) qed next case (CASRed2 e s ta e' s' v D F e3 Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs _ E2 xs\ obtain E where "E2 = Val v\compareAndSwap(D\F, E, compE1 Vs e3)" and "bisim Vs e E xs" and "\ contains_insync e3" by(auto) with IH[of Vs E xs] \fv _ \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val e\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars _ \ length xs\ \?synth _ s\ show ?case by(cases "is_val e'")(fastforce elim!: sim_move01_expr)+ next case (CASRed3 e s ta e' s' v D F v' Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \bisim Vs _ E2 xs\ obtain E where "E2 = Val v\compareAndSwap(D\F, Val v', E)" and "bisim Vs e E xs" by auto with IH[of Vs E xs] \fv _ \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars _ \ length xs\ \?synth _ s\ show ?case by(fastforce elim!: sim_move01_expr) next case (CallObj e s ta e' s' M es Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs (e\M(es)) E2 xs\ obtain E where "E2 = E\M(compEs1 Vs es)" and "bisim Vs e E xs" and "\ contains_insyncs es" by(auto simp add: compEs1_conv_map) with IH[of Vs E xs] \fv (e\M(es)) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars (e\M(es)) \ length xs\ \?synth (e\M(es)) s\ show ?case by(cases "is_val e'")(fastforce elim!: sim_move01_expr split: if_split_asm)+ next case (CallParams es s ta es' s' v M Vs E2 xs) note IH = \\vs es2 XS. \bisims vs es es2 XS; fvs es \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_varss es \ length XS; ?synths es s \ \ ?concls es es2 s XS es' s' ta vs\ from \bisim Vs (Val v\M(es)) E2 xs\ obtain Es where "E2 = Val v\M(Es)" and "bisims Vs es Es xs" by auto moreover from \extTA2J0 P,P,t \ \es,s\ [-ta\] \es',s'\\ have "\ is_vals es" by auto with \?synth (Val v\M(es)) s\ have "?synths es s" by(auto) moreover note IH[of Vs Es xs] \fv (Val v\M(es)) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars (Val v\M(es)) \ length xs\ ultimately show ?case by(fastforce elim!: sim_move01_CallParams) next case (RedCall s a U M Ts T pns body D vs Vs E2 xs) from \typeof_addr (hp s) a = \U\\ have "call (addr a\M(map Val vs)) = \(a, M, vs)\" by auto with \?synth (addr a\M(map Val vs)) s\ have "synthesized_call P (hp s) (a, M, vs)" by auto with \typeof_addr (hp s) a = \U\\ \P \ class_type_of U sees M: Ts\T = \(pns, body)\ in D\ have False by(auto simp add: synthesized_call_conv dest: sees_method_fun) thus ?case .. next case (RedCallExternal s a T M Ts Tr D vs ta va h' ta' e' s' Vs E2 xs) from \bisim Vs (addr a\M(map Val vs)) E2 xs\ have "E2 = addr a\M(map Val vs)" by auto moreover note \P \ class_type_of T sees M: Ts\Tr = Native in D\ \typeof_addr (hp s) a = \T\\ \ta' = extTA2J0 P ta\ \e' = extRet2J (addr a\M(map Val vs)) va\ \s' = (h', lcl s)\ \P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ \lcl s \\<^sub>m [Vs [\] xs]\ moreover from wf \P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ have "ta_bisim01 (extTA2J0 P ta) (extTA2J1 (compP1 P) ta)" by(rule red_external_ta_bisim01) moreover from \P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ \typeof_addr (hp s) a = \T\\ \P \ class_type_of T sees M: Ts\Tr = Native in D\ have "\external_defs D M \ h' = hp s \ ta = \" by(fastforce dest: \external'_red_external_heap_unchanged \external'_red_external_TA_empty simp add: \external'_def \external_def) ultimately show ?case by(cases va)(fastforce intro!: exI[where x=ta] intro: Red1CallExternal simp add: map_eq_append_conv sim_move01_def dest: sees_method_fun simp del: split_paired_Ex)+ next case (BlockRed e h x V vo ta e' h' x' T Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl (h, x(V := vo)) \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e (h, (x(V := vo)))\ \ ?concl e e2 (h, x(V := vo)) XS e' (h', x') ta vs\ note red = \extTA2J0 P,P,t \ \e,(h, x(V := vo))\ -ta\ \e',(h', x')\\ note len = \length Vs + max_vars {V:T=vo; e} \ length xs\ from \fv {V:T=vo; e} \ set Vs\ have fv: "fv e \ set (Vs@[V])" by auto from \bisim Vs {V:T=vo; e} E2 xs\ show ?case proof(cases rule: bisim_cases(7)[consumes 1, case_names BlockNone BlockSome BlockSomeNone]) case (BlockNone E') with red IH[of "Vs@[V]" E' xs] fv \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ \?synth {V:T=vo; e} (h, x)\ obtain TA' e2' xs' where red': "sim_move01 (compP1 P) t ta e E' h xs TA' e2' h' xs'" and bisim': "bisim (Vs @ [V]) e' e2' xs'" "x' \\<^sub>m [Vs @ [V] [\] xs']" by auto from red' \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length (Vs@[V]) + max_vars e \ length xs'" by(fastforce simp add: sim_move01_def dest: red1_preserves_len \red1t_preserves_len \red1r_preserves_len split: if_split_asm) with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "x' \\<^sub>m [Vs [\] xs', V \ xs' ! length Vs]" by(simp) moreover { assume "V \ set Vs" hence "hidden (Vs @ [V]) (index Vs V)" by(rule hidden_index) with \bisim (Vs @ [V]) e E' xs\ have "unmod E' (index Vs V)" by -(rule hidden_bisim_unmod) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ \V \ set Vs\ have "index Vs V < length xs" by(auto intro: index_less_aux) ultimately have "xs ! index Vs V = xs' ! index Vs V" using sim_move01_preserves_unmod[OF red'] by(simp) } moreover from red' have "length xs = length xs'" by(rule sim_move01_preserves_len[symmetric]) ultimately have rel: "x'(V := x V) \\<^sub>m [Vs [\] xs']" using \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ by(auto intro: Block_lem) show ?thesis proof(cases "x' V") case None with red' bisim' BlockNone len show ?thesis by(fastforce simp del: split_paired_Ex fun_upd_apply intro: rel) next case (Some v) moreover with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "[Vs @ [V] [\] xs'] V = \v\" by(auto simp add: map_le_def dest: bspec) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length Vs < length xs" by auto ultimately have "xs' ! length Vs = v" using \length xs = length xs'\ by(simp) with red' bisim' BlockNone Some len show ?thesis by(fastforce simp del: fun_upd_apply intro: rel) qed next case (BlockSome E' v) with red IH[of "Vs@[V]" E' "xs[length Vs := v]"] fv \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ \?synth {V:T=vo; e} (h, x)\ obtain TA' e2' xs' where red': "sim_move01 (compP1 P) t ta e E' h (xs[length Vs := v]) TA' e2' h' xs'" and bisim': "bisim (Vs @ [V]) e' e2' xs'" "x' \\<^sub>m [Vs @ [V] [\] xs']" by auto from red' \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length (Vs@[V]) + max_vars e \ length xs'" by(auto dest: sim_move01_preserves_len) with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "x' \\<^sub>m [Vs [\] xs', V \ xs' ! length Vs]" by(simp) moreover { assume "V \ set Vs" hence "hidden (Vs @ [V]) (index Vs V)" by(rule hidden_index) with \bisim (Vs @ [V]) e E' (xs[length Vs := v])\ have "unmod E' (index Vs V)" by -(rule hidden_bisim_unmod) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ \V \ set Vs\ have "index Vs V < length xs" by(auto intro: index_less_aux) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ \V \ set Vs\ have "(xs[length Vs := v]) ! index Vs V = xs ! index Vs V" by(simp) ultimately have "xs ! index Vs V = xs' ! index Vs V" using sim_move01_preserves_unmod[OF red', of "index Vs V"] by(simp) } moreover from red' have "length xs = length xs'" by(auto dest: sim_move01_preserves_len) ultimately have rel: "x'(V := x V) \\<^sub>m [Vs [\] xs']" using \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ by(auto intro: Block_lem) from BlockSome red obtain v' where Some: "x' V = \v'\" by(auto dest!: red_lcl_incr) with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "[Vs @ [V] [\] xs'] V = \v'\" by(auto simp add: map_le_def dest: bspec) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length Vs < length xs" by auto ultimately have "xs' ! length Vs = v'" using \length xs = length xs'\ by(simp) with red' bisim' BlockSome Some \length Vs < length xs\ show ?thesis by(fastforce simp del: fun_upd_apply intro: rel) next case (BlockSomeNone E') with red IH[of "Vs@[V]" E' xs] fv \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ \?synth {V:T=vo; e} (h, x)\ obtain TA' e2' xs' where red': "sim_move01 (compP1 P) t ta e E' h xs TA' e2' h' xs'" and IH': "bisim (Vs @ [V]) e' e2' xs'" "x' \\<^sub>m [Vs @ [V] [\] xs']" by auto from red' \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length (Vs@[V]) + max_vars e \ length xs'" by(auto dest: sim_move01_preserves_len) with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "x' \\<^sub>m [Vs [\] xs', V \ xs' ! length Vs]" by(simp) moreover { assume "V \ set Vs" hence "hidden (Vs @ [V]) (index Vs V)" by(rule hidden_index) with \bisim (Vs @ [V]) e E' xs\ have "unmod E' (index Vs V)" by -(rule hidden_bisim_unmod) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ \V \ set Vs\ have "index Vs V < length xs" by(auto intro: index_less_aux) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ \V \ set Vs\ have "(xs[length Vs := v]) ! index Vs V = xs ! index Vs V" by(simp) ultimately have "xs ! index Vs V = xs' ! index Vs V" using sim_move01_preserves_unmod[OF red', of "index Vs V"] by(simp) } moreover from red' have "length xs = length xs'" by(auto dest: sim_move01_preserves_len) ultimately have rel: "x'(V := x V) \\<^sub>m [Vs [\] xs']" using \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ by(auto intro: Block_lem) from BlockSomeNone red obtain v' where Some: "x' V = \v'\" by(auto dest!: red_lcl_incr) with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "[Vs @ [V] [\] xs'] V = \v'\" by(auto simp add: map_le_def dest: bspec) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length Vs < length xs" by auto ultimately have "xs' ! length Vs = v'" using \length xs = length xs'\ by(simp) with red' IH' BlockSomeNone Some \length Vs < length xs\ show ?thesis by(fastforce simp del: fun_upd_apply intro: rel) qed next case (RedBlock V T vo u s Vs E2 xs) from \bisim Vs {V:T=vo; Val u} E2 xs\ obtain vo' where [simp]: "E2 = {length Vs:T=vo'; Val u}" by auto from RedBlock show ?case proof(cases vo) case (Some v) with \bisim Vs {V:T=vo; Val u} E2 xs\ have vo': "case vo' of None \ xs ! length Vs = v | Some v' \ v = v'" by auto have "sim_move01 (compP1 P) t \ {V:T=vo; Val u} E2 (hp s) xs \ (Val u) (hp s) (xs[length Vs := v])" proof(cases vo') case None with vo' have "xs[length Vs := v] = xs" by auto thus ?thesis using Some None by auto next case Some from \length Vs + max_vars {V:T=vo; Val u} \ length xs\ have "length Vs < length xs" by simp with vo' Some show ?thesis using \vo = Some v\ by auto qed thus ?thesis using RedBlock by fastforce qed fastforce next case SynchronizedNull thus ?case by fastforce next case (LockSynchronized a e s Vs E2 xs) from \bisim Vs (sync(addr a) e) E2 xs\ have E2: "E2 = sync\<^bsub>length Vs\<^esub> (addr a) (compE1 (Vs@[fresh_var Vs]) e)" and sync: "\ contains_insync e" by auto moreover have "fresh_var Vs \ set Vs" by(rule fresh_var_fresh) with \fv (sync(addr a) e) \ set Vs\ have "fresh_var Vs \ fv e" by auto from E2 \fv (sync(addr a) e) \ set Vs\ sync have "bisim (Vs@[fresh_var Vs]) e (compE1 (Vs@[fresh_var Vs]) e) (xs[length Vs := Addr a])" by(auto intro!: compE1_bisim) hence "bisim Vs (insync(a) e) (insync\<^bsub>length Vs\<^esub> (a) (compE1 (Vs@[fresh_var Vs]) e)) (xs[length Vs := Addr a])" using \fresh_var Vs \ fv e\ \length Vs + max_vars (sync(addr a) e) \ length xs\ by auto moreover from \length Vs + max_vars (sync(addr a) e) \ length xs\ have "False,compP1 P,t \1 \sync\<^bsub>length Vs\<^esub> (addr a) (compE1 (Vs@[fresh_var Vs]) e), (hp s, xs)\ -\Lock\a, SyncLock a\\ \insync\<^bsub>length Vs\<^esub> (a) (compE1 (Vs@[fresh_var Vs]) e), (hp s, xs[length Vs := Addr a])\" by -(rule Lock1Synchronized, auto) hence "sim_move01 (compP1 P) t \Lock\a, SyncLock a\ (sync(addr a) e) E2 (hp s) xs \Lock\a, SyncLock a\ (insync\<^bsub>length Vs\<^esub> (a) (compE1 (Vs@[fresh_var Vs]) e)) (hp s) (xs[length Vs := Addr a])" using E2 by(fastforce simp add: sim_move01_def ta_bisim_def) moreover have "zip Vs (xs[length Vs := Addr a]) = (zip Vs xs)[length Vs := (arbitrary, Addr a)]" by(rule sym)(simp add: update_zip) hence "zip Vs (xs[length Vs := Addr a]) = zip Vs xs" by simp with \lcl s \\<^sub>m [Vs [\] xs]\ have "lcl s \\<^sub>m [Vs [\] xs[length Vs := Addr a]]" by(auto simp add: map_le_def map_upds_def) ultimately show ?case using \lcl s \\<^sub>m [Vs [\] xs]\ by fastforce next case (SynchronizedRed2 e s ta e' s' a Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \bisim Vs (insync(a) e) E2 xs\ obtain E where E2: "E2 = insync\<^bsub>length Vs\<^esub> (a) E" and bisim: "bisim (Vs@[fresh_var Vs]) e E xs" and xsa: "xs ! length Vs = Addr a" by auto from \fv (insync(a) e) \ set Vs\ fresh_var_fresh[of Vs] have fv: "fresh_var Vs \ fv e" by auto from \length Vs + max_vars (insync(a) e) \ length xs\ have "length Vs < length xs" by simp { assume "lcl s (fresh_var Vs) \ None" then obtain v where "lcl s (fresh_var Vs) = \v\" by auto with \lcl s \\<^sub>m [Vs [\] xs]\ have "[Vs [\] xs] (fresh_var Vs) = \v\" by(auto simp add: map_le_def dest: bspec) hence "fresh_var Vs \ set Vs" by(auto simp add: map_upds_def set_zip dest!: map_of_SomeD ) moreover have "fresh_var Vs \ set Vs" by(rule fresh_var_fresh) ultimately have False by contradiction } hence "lcl s (fresh_var Vs) = None" by(cases "lcl s (fresh_var Vs)", auto) hence "(lcl s)(fresh_var Vs := None) = lcl s" by(auto intro: ext) moreover from \lcl s \\<^sub>m [Vs [\] xs]\ have "(lcl s)(fresh_var Vs := None) \\<^sub>m [Vs [\] xs, fresh_var Vs \ xs ! length Vs]" by(simp) ultimately have "lcl s \\<^sub>m [Vs @ [fresh_var Vs] [\] xs]" using \length Vs < length xs\ by(auto) with IH[of "Vs@[fresh_var Vs]" E xs] \fv (insync(a) e) \ set Vs\ bisim \length Vs + max_vars (insync(a) e) \ length xs\ \?synth (insync(a) e) s\ obtain TA' e2' xs' where IH': "sim_move01 (compP1 P) t ta e E (hp s) xs TA' e2' (hp s') xs'" "bisim (Vs @ [fresh_var Vs]) e' e2' xs'" "lcl s' \\<^sub>m [Vs @ [fresh_var Vs] [\] xs']" by auto from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "dom (lcl s') \ dom (lcl s) \ fv e" by(auto dest: red_dom_lcl) with fv \lcl s (fresh_var Vs) = None\ have "(fresh_var Vs) \ dom (lcl s')" by auto hence "lcl s' (fresh_var Vs) = None" by auto moreover from IH' have "length xs = length xs'" by(auto dest: sim_move01_preserves_len) moreover note \lcl s' \\<^sub>m [Vs @ [fresh_var Vs] [\] xs']\ \length Vs < length xs\ ultimately have "lcl s' \\<^sub>m [Vs [\] xs']" by(auto simp add: map_le_def dest: bspec) moreover from bisim fv have "unmod E (length Vs)" by(auto intro: bisim_fv_unmod) with IH' \length Vs < length xs\ have "xs ! length Vs = xs' ! length Vs" by(auto dest!: sim_move01_preserves_unmod) with xsa have "xs' ! length Vs = Addr a" by simp ultimately show ?case using IH' E2 by(fastforce) next case (UnlockSynchronized a v s Vs E2 xs) from \bisim Vs (insync(a) Val v) E2 xs\ have E2: "E2 = insync\<^bsub>length Vs\<^esub> (a) Val v" and xsa: "xs ! length Vs = Addr a" by auto moreover from \length Vs + max_vars (insync(a) Val v) \ length xs\ xsa have "False,compP1 P,t \1 \insync\<^bsub>length Vs\<^esub> (a) (Val v), (hp s, xs)\ -\Unlock\a, SyncUnlock a\\ \Val v, (hp s, xs)\" by-(rule Unlock1Synchronized, auto) hence "sim_move01 (compP1 P) t \Unlock\a, SyncUnlock a\ (insync(a) Val v) (insync\<^bsub>length Vs\<^esub> (a) Val v) (hp s) xs \Unlock\a, SyncUnlock a\ (Val v) (hp s) xs" by(fastforce simp add: sim_move01_def ta_bisim_def) ultimately show ?case using \lcl s \\<^sub>m [Vs [\] xs]\ by fastforce next case (RedWhile b c s Vs E2 xs) from \bisim Vs (while (b) c) E2 xs\ have "E2 = while (compE1 Vs b) (compE1 Vs c)" and sync: "\ contains_insync b" "\ contains_insync c" by auto moreover have "False,compP1 P,t \1 \while(compE1 Vs b) (compE1 Vs c), (hp s, xs)\ -\\ \if (compE1 Vs b) (compE1 Vs c;;while(compE1 Vs b) (compE1 Vs c)) else unit, (hp s, xs)\" by(rule Red1While) hence "sim_move01 (compP1 P) t \ (while (b) c) (while (compE1 Vs b) (compE1 Vs c)) (hp s) xs \ (if (compE1 Vs b) (compE1 Vs c;;while(compE1 Vs b) (compE1 Vs c)) else unit) (hp s) xs" by(auto simp add: sim_move01_def) moreover from \fv (while (b) c) \ set Vs\ sync have "bisim Vs (if (b) (c;; while (b) c) else unit) (if (compE1 Vs b) (compE1 Vs (c;; while(b) c)) else (compE1 Vs unit)) xs" by -(rule bisimCond, auto) ultimately show ?case using \lcl s \\<^sub>m [Vs [\] xs]\ by(simp)((rule exI)+, erule conjI, auto) next case (RedTryCatch s a D C V e2 Vs E2 xs) thus ?case by(auto intro!: exI)(auto intro!: compE1_bisim) next case RedTryFail thus ?case by(auto intro!: exI) next case (ListRed1 e s ta e' s' es Vs ES2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s\ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisims Vs (e # es) ES2 xs\ obtain E' where "bisim Vs e E' xs" and ES2: "ES2 = E' # compEs1 Vs es" and sync: "\ contains_insyncs es" by(auto simp add: compEs1_conv_map) with IH[of Vs E' xs] \fvs (e # es) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_varss (e # es) \ length xs\ \?synths (e # es) s\ \\ is_val e\ show ?case by(cases "is_val e'")(fastforce elim!: sim_moves01_expr split: if_split_asm)+ next case (ListRed2 es s ta es' s' v Vs ES2 xs) thus ?case by(fastforce elim!: bisims_cases elim!: sim_moves01_expr) next case CallThrowParams thus ?case by(fastforce elim!:bisim_cases simp add: bisims_map_Val_Throw) next case (BlockThrow V T vo a s Vs e2 xs) thus ?case by(cases vo)(fastforce elim!: bisim_cases)+ next case (SynchronizedThrow2 a ad s Vs E2 xs) from \bisim Vs (insync(a) Throw ad) E2 xs\ have "xs ! length Vs = Addr a" by auto with \length Vs + max_vars (insync(a) Throw ad) \ length xs\ have "False,compP1 P,t \1 \insync\<^bsub>length Vs\<^esub> (a) Throw ad, (hp s, xs)\ -\Unlock\a, SyncUnlock a\\ \Throw ad, (hp s, xs)\" by-(rule Synchronized1Throw2, auto) hence "sim_move01 (compP1 P) t \Unlock\a, SyncUnlock a\ (insync(a) Throw ad) (insync\<^bsub>length Vs\<^esub> (a) Throw ad) (hp s) xs \Unlock\a, SyncUnlock a\ (Throw ad) (hp s) xs" by(fastforce simp add: sim_move01_def ta_bisim_def fun_eq_iff expand_finfun_eq finfun_upd_apply ta_upd_simps split: if_split_asm) moreover note \lcl s \\<^sub>m [Vs [\] xs]\ \bisim Vs (insync(a) Throw ad) E2 xs\ ultimately show ?case by(fastforce) next case InstanceOfRed thus ?case by(fastforce) next case RedInstanceOf thus ?case by(auto intro!: exI) next case InstanceOfThrow thus ?case by fastforce qed(fastforce simp del: fun_upd_apply split: if_split_asm)+ end declare max_dest [iff del] declare split_paired_Ex [simp del] primrec countInitBlock :: "('a, 'b, 'addr) exp \ nat" and countInitBlocks :: "('a, 'b, 'addr) exp list \ nat" where "countInitBlock (new C) = 0" | "countInitBlock (newA T\e\) = countInitBlock e" | "countInitBlock (Cast T e) = countInitBlock e" | "countInitBlock (e instanceof T) = countInitBlock e" | "countInitBlock (Val v) = 0" | "countInitBlock (Var V) = 0" | "countInitBlock (V := e) = countInitBlock e" | "countInitBlock (e \bop\ e') = countInitBlock e + countInitBlock e'" | "countInitBlock (a\i\) = countInitBlock a + countInitBlock i" | "countInitBlock (AAss a i e) = countInitBlock a + countInitBlock i + countInitBlock e" | "countInitBlock (a\length) = countInitBlock a" | "countInitBlock (e\F{D}) = countInitBlock e" | "countInitBlock (FAss e F D e') = countInitBlock e + countInitBlock e'" | "countInitBlock (e\compareAndSwap(D\F, e', e'')) = countInitBlock e + countInitBlock e' + countInitBlock e''" | "countInitBlock (e\M(es)) = countInitBlock e + countInitBlocks es" | "countInitBlock ({V:T=vo; e}) = (case vo of None \ 0 | Some v \ 1) + countInitBlock e" | "countInitBlock (sync\<^bsub>V'\<^esub> (e) e') = countInitBlock e + countInitBlock e'" | "countInitBlock (insync\<^bsub>V'\<^esub> (ad) e) = countInitBlock e" | "countInitBlock (e;;e') = countInitBlock e + countInitBlock e'" | "countInitBlock (if (e) e1 else e2) = countInitBlock e + countInitBlock e1 + countInitBlock e2" | "countInitBlock (while(b) e) = countInitBlock b + countInitBlock e" | "countInitBlock (throw e) = countInitBlock e" | "countInitBlock (try e catch(C V) e') = countInitBlock e + countInitBlock e'" | "countInitBlocks [] = 0" | "countInitBlocks (e # es) = countInitBlock e + countInitBlocks es" context J0_J1_heap_base begin lemmas \red0r_expr = NewArray_\red0r_xt Cast_\red0r_xt InstanceOf_\red0r_xt BinOp_\red0r_xt1 BinOp_\red0r_xt2 LAss_\red0r AAcc_\red0r_xt1 AAcc_\red0r_xt2 AAss_\red0r_xt1 AAss_\red0r_xt2 AAss_\red0r_xt3 ALength_\red0r_xt FAcc_\red0r_xt FAss_\red0r_xt1 FAss_\red0r_xt2 CAS_\red0r_xt1 CAS_\red0r_xt2 CAS_\red0r_xt3 Call_\red0r_obj Call_\red0r_param Block_\red0r_xt Sync_\red0r_xt InSync_\red0r_xt Seq_\red0r_xt Cond_\red0r_xt Throw_\red0r_xt Try_\red0r_xt lemmas \red0t_expr = NewArray_\red0t_xt Cast_\red0t_xt InstanceOf_\red0t_xt BinOp_\red0t_xt1 BinOp_\red0t_xt2 LAss_\red0t AAcc_\red0t_xt1 AAcc_\red0t_xt2 AAss_\red0t_xt1 AAss_\red0t_xt2 AAss_\red0t_xt3 ALength_\red0t_xt FAcc_\red0t_xt FAss_\red0t_xt1 FAss_\red0t_xt2 CAS_\red0t_xt1 CAS_\red0t_xt2 CAS_\red0t_xt3 Call_\red0t_obj Call_\red0t_param Block_\red0t_xt Sync_\red0t_xt InSync_\red0t_xt Seq_\red0t_xt Cond_\red0t_xt Throw_\red0t_xt Try_\red0t_xt declare \red0r_expr [elim!] declare \red0t_expr [elim!] definition sim_move10 :: "'addr J_prog \ 'thread_id \ ('addr, 'thread_id, 'heap) external_thread_action \ 'addr expr1 \ 'addr expr1 \ 'addr expr \ 'heap \ 'addr locals \ ('addr, 'thread_id, 'heap) J0_thread_action \ 'addr expr \ 'heap \ 'addr locals \ bool" where "sim_move10 P t ta1 e1 e1' e h xs ta e' h' xs' \ \ final e1 \ (if \move1 P h e1 then (\red0t (extTA2J0 P) P t h (e, xs) (e', xs') \ countInitBlock e1' < countInitBlock e1 \ e' = e \ xs' = xs) \ h' = h \ ta1 = \ \ ta = \ else ta_bisim01 ta (extTA2J1 (compP1 P) ta1) \ (if call e = None \ call1 e1 = None then (\e'' xs''. \red0r (extTA2J0 P) P t h (e, xs) (e'', xs'') \ extTA2J0 P,P,t \ \e'', (h, xs'')\ -ta\ \e', (h', xs')\ \ no_call P h e'' \ \ \move0 P h e'') else extTA2J0 P,P,t \ \e, (h, xs)\ -ta\ \e', (h', xs')\ \ no_call P h e \ \ \move0 P h e))" definition sim_moves10 :: "'addr J_prog \ 'thread_id \ ('addr, 'thread_id, 'heap) external_thread_action \ 'addr expr1 list \ 'addr expr1 list \ 'addr expr list \ 'heap \ 'addr locals \ ('addr, 'thread_id, 'heap) J0_thread_action \ 'addr expr list \ 'heap \ 'addr locals \ bool" where "sim_moves10 P t ta1 es1 es1' es h xs ta es' h' xs' \ \ finals es1 \ (if \moves1 P h es1 then (\reds0t (extTA2J0 P) P t h (es, xs) (es', xs') \ countInitBlocks es1' < countInitBlocks es1 \ es' = es \ xs' = xs) \ h' = h \ ta1 = \ \ ta = \ else ta_bisim01 ta (extTA2J1 (compP1 P) ta1) \ (if calls es = None \ calls1 es1 = None then (\es'' xs''. \reds0r (extTA2J0 P) P t h (es, xs) (es'', xs'') \ extTA2J0 P,P,t \ \es'', (h, xs'')\ [-ta\] \es', (h', xs')\ \ no_calls P h es'' \ \ \moves0 P h es'') else extTA2J0 P,P,t \ \es, (h, xs)\ [-ta\] \es', (h', xs')\ \ no_calls P h es \ \ \moves0 P h es))" lemma sim_move10_expr: assumes "sim_move10 P t ta1 e1 e1' e h xs ta e' h' xs'" shows "sim_move10 P t ta1 (newA T\e1\) (newA T\e1'\) (newA T\e\) h xs ta (newA T\e'\) h' xs'" "sim_move10 P t ta1 (Cast T e1) (Cast T e1') (Cast T e) h xs ta (Cast T e') h' xs'" "sim_move10 P t ta1 (e1 instanceof T) (e1' instanceof T) (e instanceof T) h xs ta (e' instanceof T) h' xs'" "sim_move10 P t ta1 (e1 \bop\ e2) (e1' \bop\ e2) (e \bop\ e2') h xs ta (e' \bop\ e2') h' xs'" "sim_move10 P t ta1 (Val v \bop\ e1) (Val v \bop\ e1') (Val v \bop\ e) h xs ta (Val v \bop\ e') h' xs'" "sim_move10 P t ta1 (V := e1) (V := e1') (V' := e) h xs ta (V' := e') h' xs'" "sim_move10 P t ta1 (e1\e2\) (e1'\e2\) (e\e2'\) h xs ta (e'\e2'\) h' xs'" "sim_move10 P t ta1 (Val v\e1\) (Val v\e1'\) (Val v\e\) h xs ta (Val v\e'\) h' xs'" "sim_move10 P t ta1 (e1\e2\ := e3) (e1'\e2\ := e3) (e\e2'\ := e3') h xs ta (e'\e2'\ := e3') h' xs'" "sim_move10 P t ta1 (Val v\e1\ := e3) (Val v\e1'\ := e3) (Val v\e\ := e3') h xs ta (Val v\e'\ := e3') h' xs'" "sim_move10 P t ta1 (AAss (Val v) (Val v') e1) (AAss (Val v) (Val v') e1') (AAss (Val v) (Val v') e) h xs ta (AAss (Val v) (Val v') e') h' xs'" "sim_move10 P t ta1 (e1\length) (e1'\length) (e\length) h xs ta (e'\length) h' xs'" "sim_move10 P t ta1 (e1\F{D}) (e1'\F{D}) (e\F'{D'}) h xs ta (e'\F'{D'}) h' xs'" "sim_move10 P t ta1 (FAss e1 F D e2) (FAss e1' F D e2) (FAss e F' D' e2') h xs ta (FAss e' F' D' e2') h' xs'" "sim_move10 P t ta1 (FAss (Val v) F D e1) (FAss (Val v) F D e1') (FAss (Val v) F' D' e) h xs ta (FAss (Val v) F' D' e') h' xs'" "sim_move10 P t ta1 (CompareAndSwap e1 F D e2 e3) (CompareAndSwap e1' F D e2 e3) (CompareAndSwap e F' D' e2' e3') h xs ta (CompareAndSwap e' F' D' e2' e3') h' xs'" "sim_move10 P t ta1 (CompareAndSwap (Val v) F D e1 e3) (CompareAndSwap (Val v) F D e1' e3) (CompareAndSwap (Val v) F' D' e e3') h xs ta (CompareAndSwap (Val v) F' D' e' e3') h' xs'" "sim_move10 P t ta1 (CompareAndSwap (Val v) F D (Val v') e1) (CompareAndSwap (Val v) F D (Val v') e1') (CompareAndSwap (Val v) F' D' (Val v') e) h xs ta (CompareAndSwap (Val v) F' D' (Val v') e') h' xs'" "sim_move10 P t ta1 (e1\M(es)) (e1'\M(es)) (e\M(es')) h xs ta (e'\M(es')) h' xs'" "sim_move10 P t ta1 (sync\<^bsub>V\<^esub>(e1) e2) (sync\<^bsub>V\<^esub>(e1') e2) (sync(e) e2') h xs ta (sync(e') e2') h' xs'" "sim_move10 P t ta1 (insync\<^bsub>V\<^esub>(a) e1) (insync\<^bsub>V\<^esub>(a) e1') (insync(a') e) h xs ta (insync(a') e') h' xs'" "sim_move10 P t ta1 (e1;;e2) (e1';;e2) (e;;e2') h xs ta (e';;e2') h' xs'" "sim_move10 P t ta1 (if (e1) e2 else e3) (if (e1') e2 else e3) (if (e) e2' else e3') h xs ta (if (e') e2' else e3') h' xs'" "sim_move10 P t ta1 (throw e1) (throw e1') (throw e) h xs ta (throw e') h' xs'" "sim_move10 P t ta1 (try e1 catch(C V) e2) (try e1' catch(C V) e2) (try e catch(C' V') e2') h xs ta (try e' catch(C' V') e2') h' xs'" using assms apply(simp_all add: sim_move10_def final_iff split del: if_split split: if_split_asm) apply(fastforce simp: \red0t_Val \red0r_Val intro: red_reds.intros split!: if_splits)+ done lemma sim_moves10_expr: "sim_move10 P t ta1 e1 e1' e h xs ta e' h' xs' \ sim_moves10 P t ta1 (e1 # es2) (e1' # es2) (e # es2') h xs ta (e' # es2') h' xs'" "sim_moves10 P t ta1 es1 es1' es h xs ta es' h' xs' \ sim_moves10 P t ta1 (Val v # es1) (Val v # es1') (Val v # es) h xs ta (Val v # es') h' xs'" unfolding sim_moves10_def sim_move10_def final_iff finals_iff apply(simp_all add: Cons_eq_append_conv split del: if_split split: if_split_asm) apply(safe intro!: if_split) apply(fastforce simp add: is_vals_conv \reds0t_map_Val \reds0r_map_Val \red0t_Val \red0r_Val intro!: \red0r_inj_\reds0r \reds0r_cons_\reds0r \red0t_inj_\reds0t \reds0t_cons_\reds0t ListRed1 ListRed2 split: if_split_asm)+ done lemma sim_move10_CallParams: "sim_moves10 P t ta1 es1 es1' es h xs ta es' h' xs' \ sim_move10 P t ta1 (Val v\M(es1)) (Val v\M(es1')) (Val v\M(es)) h xs ta (Val v\M(es')) h' xs'" unfolding sim_move10_def sim_moves10_def apply(simp split: if_split_asm split del: if_split add: is_vals_conv) apply(fastforce simp add: \red0t_Val \red0r_Val \reds0t_map_Val \reds0r_map_Val intro: Call_\red0r_param Call_\red0t_param CallParams split: if_split_asm split del: if_split intro!: if_split) apply(rule conjI) apply fastforce apply(rule if_intro) apply fastforce apply(clarsimp split del: if_split) apply(rule if_intro) apply(clarsimp split: if_split_asm simp add: is_vals_conv) apply(rule exI conjI)+ apply(erule Call_\red0r_param) apply(fastforce intro: CallParams) apply(rule exI conjI)+ apply(erule Call_\red0r_param) apply(fastforce intro!: CallParams) apply(clarsimp split del: if_split split: if_split_asm simp add: is_vals_conv \reds0r_map_Val) apply fastforce apply(rule conjI) apply fastforce apply(rule if_intro) apply fastforce apply(rule conjI) apply fastforce apply(rule if_intro) apply(clarsimp split: if_split_asm) apply(clarsimp split: if_split_asm split del: if_split simp add: is_vals_conv) apply(fastforce intro: CallParams) done lemma sim_move10_Block: "sim_move10 P t ta1 e1 e1' e h (xs(V' := vo)) ta e' h' xs' \ sim_move10 P t ta1 ({V:T=None; e1}) ({V:T=None; e1'}) ({V':T=vo; e}) h xs ta ({V':T=xs' V'; e'}) h' (xs'(V' := xs V'))" proof - assume "sim_move10 P t ta1 e1 e1' e h (xs(V' := vo)) ta e' h' xs'" moreover { fix e'' xs'' assume "extTA2J0 P,P,t \ \e'',(h, xs'')\ -ta\ \e',(h', xs')\" hence "extTA2J0 P,P,t \ \e'',(h, xs''(V' := xs V', V' := xs'' V'))\ -ta\ \e',(h', xs')\" by simp from BlockRed[OF this, of T] have "extTA2J0 P,P,t \ \{V':T=xs'' V'; e''},(h, xs''(V' := xs V'))\ -ta\ \{V':T=xs' V'; e'},(h', xs'(V' := xs V'))\" by simp } ultimately show ?thesis by(fastforce simp add: sim_move10_def final_iff split: if_split_asm) qed lemma sim_move10_reds: "\ (h', a) \ allocate h (Class_type C); ta1 = \NewHeapElem a (Class_type C)\; ta = \NewHeapElem a (Class_type C)\ \ \ sim_move10 P t ta1 (new C) e1' (new C) h xs ta (addr a) h' xs" "allocate h (Class_type C) = {} \ sim_move10 P t \ (new C) e1' (new C) h xs \ (THROW OutOfMemory) h xs" "\ (h', a) \ allocate h (Array_type T (nat (sint i))); 0 <=s i; ta1 = \NewHeapElem a (Array_type T (nat (sint i)))\; ta = \NewHeapElem a (Array_type T (nat (sint i)))\ \ \ sim_move10 P t ta1 (newA T\Val (Intg i)\) e1' (newA T\Val (Intg i)\) h xs ta (addr a) h' xs" "i sim_move10 P t \ (newA T\Val (Intg i)\) e1' (newA T\Val (Intg i)\) h xs \ (THROW NegativeArraySize) h xs" "\ allocate h (Array_type T (nat (sint i))) = {}; 0 <=s i \ \ sim_move10 P t \ (newA T\Val (Intg i)\) e1' (newA T\Val (Intg i)\) h xs \ (THROW OutOfMemory) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; P \ U \ T \ \ sim_move10 P t \ (Cast T (Val v)) e1' (Cast T (Val v)) h xs \ (Val v) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; \ P \ U \ T \ \ sim_move10 P t \ (Cast T (Val v)) e1' (Cast T (Val v)) h xs \ (THROW ClassCast) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; b \ v \ Null \ P \ U \ T \ \ sim_move10 P t \ ((Val v) instanceof T) e1' ((Val v) instanceof T) h xs \ (Val (Bool b)) h xs" "binop bop v1 v2 = Some (Inl v) \ sim_move10 P t \ ((Val v1) \bop\ (Val v2)) e1' (Val v1 \bop\ Val v2) h xs \ (Val v) h xs" "binop bop v1 v2 = Some (Inr a) \ sim_move10 P t \ ((Val v1) \bop\ (Val v2)) e1' (Val v1 \bop\ Val v2) h xs \ (Throw a) h xs" "xs V = \v\ \ sim_move10 P t \ (Var V') e1' (Var V) h xs \ (Val v) h xs" "sim_move10 P t \ (V' := Val v) e1' (V := Val v) h xs \ unit h (xs(V \ v))" "sim_move10 P t \ (null\Val v\) e1' (null\Val v\) h xs \ (THROW NullPointer) h xs" "\ typeof_addr h a = \Array_type T n\; i sint i \ int n \ \ sim_move10 P t \ (addr a\Val (Intg i)\) e1' ((addr a)\Val (Intg i)\) h xs \ (THROW ArrayIndexOutOfBounds) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; heap_read h a (ACell (nat (sint i))) v; ta1 = \ReadMem a (ACell (nat (sint i))) v\; ta = \ReadMem a (ACell (nat (sint i))) v\ \ \ sim_move10 P t ta1 (addr a\Val (Intg i)\) e1' ((addr a)\Val (Intg i)\) h xs ta (Val v) h xs" "sim_move10 P t \ (null\Val v\ := Val v') e1' (null\Val v\ := Val v') h xs \ (THROW NullPointer) h xs" "\ typeof_addr h a = \Array_type T n\; i sint i \ int n \ \ sim_move10 P t \ (AAss (addr a) (Val (Intg i)) (Val v)) e1' (AAss (addr a) (Val (Intg i)) (Val v)) h xs \ (THROW ArrayIndexOutOfBounds) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>h\<^esub> v = \U\; \ (P \ U \ T) \ \ sim_move10 P t \ (AAss (addr a) (Val (Intg i)) (Val v)) e1' (AAss (addr a) (Val (Intg i)) (Val v)) h xs \ (THROW ArrayStore) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>h\<^esub> v = Some U; P \ U \ T; heap_write h a (ACell (nat (sint i))) v h'; ta1 = \WriteMem a (ACell (nat (sint i))) v\; ta = \WriteMem a (ACell (nat (sint i))) v\ \ \ sim_move10 P t ta1 (AAss (addr a) (Val (Intg i)) (Val v)) e1' (AAss (addr a) (Val (Intg i)) (Val v)) h xs ta unit h' xs" "typeof_addr h a = \Array_type T n\ \ sim_move10 P t \ (addr a\length) e1' (addr a\length) h xs \ (Val (Intg (word_of_nat n))) h xs" "sim_move10 P t \ (null\length) e1' (null\length) h xs \ (THROW NullPointer) h xs" "\ heap_read h a (CField D F) v; ta1 = \ReadMem a (CField D F) v\; ta = \ReadMem a (CField D F) v\ \ \ sim_move10 P t ta1 (addr a\F{D}) e1' (addr a\F{D}) h xs ta (Val v) h xs" "sim_move10 P t \ (null\F{D}) e1' (null\F{D}) h xs \ (THROW NullPointer) h xs" "\ heap_write h a (CField D F) v h'; ta1 = \WriteMem a (CField D F) v\; ta = \WriteMem a (CField D F) v\ \ \ sim_move10 P t ta1 (addr a\F{D} := Val v) e1' (addr a\F{D} := Val v) h xs ta unit h' xs" "sim_move10 P t \ (null\compareAndSwap(D\F, Val v, Val v')) e1' (null\compareAndSwap(D\F, Val v, Val v')) h xs \ (THROW NullPointer) h xs" "\ heap_read h a (CField D F) v''; heap_write h a (CField D F) v' h'; v'' = v; ta1 = \ ReadMem a (CField D F) v'', WriteMem a (CField D F) v' \; ta = \ ReadMem a (CField D F) v'', WriteMem a (CField D F) v' \ \ \ sim_move10 P t ta1 (addr a\compareAndSwap(D\F, Val v, Val v')) e1' (addr a\compareAndSwap(D\F, Val v, Val v')) h xs ta true h' xs" "\ heap_read h a (CField D F) v''; v'' \ v; ta1 = \ ReadMem a (CField D F) v'' \; ta = \ ReadMem a (CField D F) v'' \ \ \ sim_move10 P t ta1 (addr a\compareAndSwap(D\F, Val v, Val v')) e1' (addr a\compareAndSwap(D\F, Val v, Val v')) h xs ta false h xs" "sim_move10 P t \ (null\F{D} := Val v) e1' (null\F{D} := Val v) h xs \ (THROW NullPointer) h xs" "sim_move10 P t \ ({V':T=None; Val u}) e1' ({V:T=vo; Val u}) h xs \ (Val u) h xs" "sim_move10 P t \ ({V':T=\v\; e}) ({V':T=None; e}) ({V:T=vo; e'}) h xs \ ({V:T=vo; e'}) h xs" "sim_move10 P t \ (sync\<^bsub>V'\<^esub>(null) e0) e1' (sync(null) e1) h xs \ (THROW NullPointer) h xs" "sim_move10 P t \ (Val v;;e0) e1' (Val v;; e1) h xs \ e1 h xs" "sim_move10 P t \ (if (true) e0 else e0') e1' (if (true) e1 else e2) h xs \ e1 h xs" "sim_move10 P t \ (if (false) e0 else e0') e1' (if (false) e1 else e2) h xs \ e2 h xs" "sim_move10 P t \ (throw null) e1' (throw null) h xs \ (THROW NullPointer) h xs" "sim_move10 P t \ (try (Val v) catch(C V') e0) e1' (try (Val v) catch(C V) e1) h xs \ (Val v) h xs" "\ typeof_addr h a = \Class_type D\; P \ D \\<^sup>* C \ \ sim_move10 P t \ (try (Throw a) catch(C V') e0) e1' (try (Throw a) catch(C V) e1) h xs \ ({V:Class C=\Addr a\; e1}) h xs" "sim_move10 P t \ (newA T\Throw a\) e1' (newA T\Throw a\) h xs \ (Throw a) h xs" "sim_move10 P t \ (Cast T (Throw a)) e1' (Cast T (Throw a)) h xs \ (Throw a) h xs" "sim_move10 P t \ ((Throw a) instanceof T) e1' ((Throw a) instanceof T) h xs \ (Throw a) h xs" "sim_move10 P t \ ((Throw a) \bop\ e0) e1' ((Throw a) \bop\ e1) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v \bop\ (Throw a)) e1' (Val v \bop\ (Throw a)) h xs \ (Throw a) h xs" "sim_move10 P t \ (V' := Throw a) e1' (V := Throw a) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\e0\) e1' (Throw a\e1\) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v\Throw a\) e1' (Val v\Throw a\) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\e0\ := e0') e1' (Throw a\e1\ := e2) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v\Throw a\ := e0) e1' (Val v\Throw a\ := e1) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v\Val v'\ := Throw a) e1' (Val v\Val v'\ := Throw a) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\length) e1' (Throw a\length) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\F{D}) e1' (Throw a\F{D}) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\F{D} := e0) e1' (Throw a\F{D} := e1) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v\F{D} := Throw a) e1' (Val v\F{D} := Throw a) h xs \ (Throw a) h xs" "sim_move10 P t \ (CompareAndSwap (Throw a) D F e0 e0') e1' (Throw a\compareAndSwap(D\F, e1'', e1''')) h xs \ (Throw a) h xs" "sim_move10 P t \ (CompareAndSwap (Val v) D F (Throw a) e0') e1' (Val v\compareAndSwap(D\F, Throw a, e1'')) h xs \ (Throw a) h xs" "sim_move10 P t \ (CompareAndSwap (Val v) D F (Val v') (Throw a)) e1' (Val v\compareAndSwap(D\F, Val v', Throw a)) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\M(es0)) e1' (Throw a\M(es1)) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v\M(map Val vs @ Throw a # es0)) e1' (Val v\M(map Val vs @ Throw a # es1)) h xs \ (Throw a) h xs" "sim_move10 P t \ ({V':T=None; Throw a}) e1' ({V:T=vo; Throw a}) h xs \ (Throw a) h xs" "sim_move10 P t \ (sync\<^bsub>V'\<^esub>(Throw a) e0) e1' (sync(Throw a) e1) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a;;e0) e1' (Throw a;;e1) h xs \ (Throw a) h xs" "sim_move10 P t \ (if (Throw a) e0 else e0') e1' (if (Throw a) e1 else e2) h xs \ (Throw a) h xs" "sim_move10 P t \ (throw (Throw a)) e1' (throw (Throw a)) h xs \ (Throw a) h xs" apply(fastforce simp add: sim_move10_def no_calls_def no_call_def ta_bisim_def intro: red_reds.intros)+ done lemma sim_move10_CallNull: "sim_move10 P t \ (null\M(map Val vs)) e1' (null\M(map Val vs)) h xs \ (THROW NullPointer) h xs" by(fastforce simp add: sim_move10_def map_eq_append_conv intro: RedCallNull) lemma sim_move10_SyncLocks: "\ ta1 = \Lock\a, SyncLock a\; ta = \Lock\a, SyncLock a\ \ \ sim_move10 P t ta1 (sync\<^bsub>V\<^esub>(addr a) e0) e1' (sync(addr a) e1) h xs ta (insync(a) e1) h xs" "\ ta1 = \Unlock\a, SyncUnlock a\; ta = \Unlock\a, SyncUnlock a\ \ \ sim_move10 P t ta1 (insync\<^bsub>V\<^esub>(a') (Val v)) e1' (insync(a) (Val v)) h xs ta (Val v) h xs" "\ ta1 = \Unlock\a, SyncUnlock a\; ta = \Unlock\a, SyncUnlock a\ \ \ sim_move10 P t ta1 (insync\<^bsub>V\<^esub>(a') (Throw a'')) e1' (insync(a) (Throw a'')) h xs ta (Throw a'') h xs" by(fastforce simp add: sim_move10_def ta_bisim_def ta_upd_simps intro: red_reds.intros[simplified])+ lemma sim_move10_TryFail: "\ typeof_addr h a = \Class_type D\; \ P \ D \\<^sup>* C \ \ sim_move10 P t \ (try (Throw a) catch(C V') e0) e1' (try (Throw a) catch(C V) e1) h xs \ (Throw a) h xs" by(auto simp add: sim_move10_def intro!: RedTryFail) lemmas sim_move10_intros = sim_move10_expr sim_move10_reds sim_move10_CallNull sim_move10_TryFail sim_move10_Block sim_move10_CallParams lemma sim_move10_preserves_defass: assumes wf: "wf_J_prog P" shows "\ sim_move10 P t ta1 e1 e1' e h xs ta e' h' xs'; \ e \dom xs\ \ \ \ e' \dom xs'\" by(auto simp add: sim_move10_def split: if_split_asm dest!: \red0r_preserves_defass[OF wf] \red0t_preserves_defass[OF wf] red_preserves_defass[OF wf]) declare sim_move10_intros[intro] lemma assumes wf: "wf_J_prog P" shows red_simulates_red1_aux: "\ False,compP1 P,t \1 \e1, S\ -TA\ \e1', S'\; bisim vs e2 e1 (lcl S); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl S]; length vs + max_vars e1 \ length (lcl S); \ e2 \dom x\ \ \ \ta e2' x'. sim_move10 P t TA e1 e1' e2 (hp S) x ta e2' (hp S') x' \ bisim vs e2' e1' (lcl S') \ x' \\<^sub>m [vs [\] lcl S']" (is "\ _; _; _; _; _; _ \ \ ?concl e1 e1' e2 S x TA S' e1' vs") and reds_simulates_reds1_aux: "\ False,compP1 P,t \1 \es1, S\ [-TA\] \es1', S'\; bisims vs es2 es1 (lcl S); fvs es2 \ set vs; x \\<^sub>m [vs [\] lcl S]; length vs + max_varss es1 \ length (lcl S); \s es2 \dom x\ \ \ \ta es2' x'. sim_moves10 P t TA es1 es1' es2 (hp S) x ta es2' (hp S') x' \ bisims vs es2' es1' (lcl S') \ x' \\<^sub>m [vs [\] lcl S']" (is "\ _; _; _; _; _; _ \ \ ?concls es1 es1' es2 S x TA S' es1' vs") proof(induct arbitrary: vs e2 x and vs es2 x rule: red1_reds1.inducts) case (Bin1OpRed1 e s ta e' s' bop e2 Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs E2 (e \bop\ e2) (lcl s)\ obtain E E2' where E2: "E2 = E \bop\ E2'" "e2 = compE1 Vs E2'" and "bisim Vs E e (lcl s)" and sync: "\ contains_insync E2'" by(auto elim!: bisim_cases) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (e \bop\ e2) \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) with E2 \fv E2 \ set Vs\ sync show ?case by(cases "is_val e2'")(auto intro: BinOpRed1) next case (Bin1OpRed2 e s ta e' s' v bop Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 (Val v \bop\ e) (lcl s)\ obtain E where E2: "E2 = Val v \bop\ E" and "bisim Vs E e (lcl s)" by(auto) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (Val v \bop\ e) \ length (lcl s)\ \\ E2 \dom X\\ E2 ultimately show ?case by(auto intro: BinOpRed2) next case (Red1Var s V v Vs E2 X) from \bisim Vs E2 (Var V) (lcl s)\ \fv E2 \ set Vs\ obtain V' where "E2 = Var V'" "V' = Vs ! V" "V = index Vs V'" by(clarify, simp) from \E2 = Var V'\ \\ E2 \dom X\\ obtain v' where "X V' = \v'\" by(auto simp add: hyperset_defs) with \X \\<^sub>m [Vs [\] lcl s]\ have "[Vs [\] lcl s] V' = \v'\" by(rule map_le_SomeD) with \length Vs + max_vars (Var V) \ length (lcl s)\ have "lcl s ! (index Vs V') = v'" by(auto intro: map_upds_Some_eq_nth_index) with \V = index Vs V'\ \lcl s ! V = v\ have "v = v'" by simp with \X V' = \v'\\ \E2 = Var V'\ \X \\<^sub>m [Vs [\] lcl s]\ show ?case by(fastforce intro: RedVar) next case (LAss1Red e s ta e' s' V Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 (V:=e) (lcl s)\ obtain E V' where E2: "E2 = (V':=E)" "V = index Vs V'" and "bisim Vs E e (lcl s)" by auto with IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (V:=e) \ length (lcl s)\ \\ E2 \dom X\\ E2 show ?case by(auto intro: LAssRed) next case (Red1LAss V l v h Vs E2 X) from \bisim Vs E2 (V:=Val v) (lcl (h, l))\ obtain V' where "E2 = (V' := Val v)" "V = index Vs V'" by(auto) moreover with \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl (h, l)]\ \length Vs + max_vars (V:=Val v) \ length (lcl (h, l))\ have "X(V' \ v) \\<^sub>m [Vs [\] l[index Vs V' := v]]" by(auto intro: LAss_lem) ultimately show ?case by(auto intro: RedLAss simp del: fun_upd_apply) next case (AAcc1Red1 a s ta a' s' i Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 a (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars a \ length (lcl s); \ e2 \dom x\\ \ ?concl a a' e2 s x ta s' a' vs\ from \False,compP1 P,t \1 \a,s\ -ta\ \a',s'\\ have "\ is_val a" by auto with \bisim Vs E2 (a\i\) (lcl s)\ obtain E E2' where E2: "E2 = E\E2'\" "i = compE1 Vs E2'" and "bisim Vs E a (lcl s)" and sync: "\ contains_insync E2'" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (a\i\) \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where "sim_move10 P t ta a a' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' a' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) with E2 \fv E2 \ set Vs\ sync show ?case by(cases "is_val e2'")(auto intro: AAccRed1) next case (AAcc1Red2 i s ta i' s' a Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 i (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars i \ length (lcl s); \ e2 \dom x\\ \ ?concl i i' e2 s x ta s' i' vs\ from \bisim Vs E2 (Val a\i\) (lcl s)\ obtain E where E2: "E2 = Val a\E\" and "bisim Vs E i (lcl s)" by(auto) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ E2 \length Vs + max_vars (Val a\i\) \ length (lcl s)\ \\ E2 \dom X\\ ultimately show ?case by(auto intro: AAccRed2) next case Red1AAcc thus ?case by(fastforce intro: RedAAcc simp del: fun_upd_apply) next case (AAss1Red1 a s ta a' s' i e Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 a (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars a \ length (lcl s); \ e2 \dom x\ \ \ ?concl a a' e2 s x ta s' a' vs\ from \False,compP1 P,t \1 \a,s\ -ta\ \a',s'\\ have "\ is_val a" by auto with \bisim Vs E2 (a\i\:=e) (lcl s)\ obtain E E2' E2'' where E2: "E2 = E\E2'\:=E2''" "i = compE1 Vs E2'" "e = compE1 Vs E2''" and "bisim Vs E a (lcl s)" and sync: "\ contains_insync E2'" "\ contains_insync E2''" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (a\i\:=e) \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where IH': "sim_move10 P t ta a a' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' a' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) show ?case proof(cases "is_val e2'") case True from E2 \fv E2 \ set Vs\ sync have "bisim Vs E2' i (lcl s')" "bisim Vs E2'' e (lcl s')" by auto with IH' E2 True sync show ?thesis by(cases "is_val E2'")(fastforce intro: AAssRed1)+ next case False with IH' E2 sync show ?thesis by(fastforce intro: AAssRed1) qed next case (AAss1Red2 i s ta i' s' a e Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 i (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars i \ length (lcl s); \ e2 \dom x\\ \ ?concl i i' e2 s x ta s' i' vs\ from \False,compP1 P,t \1 \i,s\ -ta\ \i',s'\\ have "\ is_val i" by auto with \bisim Vs E2 (Val a\i\:=e) (lcl s)\ obtain E E2' where E2: "E2 = Val a\E\:=E2'" "e = compE1 Vs E2'" and "bisim Vs E i (lcl s)" and sync: "\ contains_insync E2'" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (Val a\i\:=e) \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where "sim_move10 P t ta i i' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' i' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) with E2 \fv E2 \ set Vs\ sync show ?case by(cases "is_val e2'")(fastforce intro: AAssRed2)+ next case (AAss1Red3 e s ta e' s' a i Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 (Val a\Val i\:=e) (lcl s)\ obtain E where E2: "E2 = Val a\Val i\:=E" and "bisim Vs E e (lcl s)" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ E2 \length Vs + max_vars (Val a\Val i\:=e) \ length (lcl s)\ \\ E2 \dom X\\ ultimately show ?case by(fastforce intro: AAssRed3) next case Red1AAssStore thus ?case by(auto)((rule exI conjI)+, auto) next case Red1AAss thus ?case by(fastforce simp del: fun_upd_apply) next case (FAss1Red1 e s ta e' s' F D e2 Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs E2 (e\F{D}:=e2) (lcl s)\ obtain E E2' where E2: "E2 = E\F{D}:=E2'" "e2 = compE1 Vs E2'" and "bisim Vs E e (lcl s)" and sync: "\ contains_insync E2'" by(fastforce) with IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (e\F{D}:=e2) \ length (lcl s)\ \\ E2 \dom X\\ obtain TA' e2' x' where "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(fastforce) with E2 \fv E2 \ set Vs\ sync show ?case by(cases "is_val e2'")(auto intro: FAssRed1) next case (FAss1Red2 e s ta e' s' v F D Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 (Val v\F{D}:=e) (lcl s)\ obtain E where E2: "E2 = (Val v\F{D}:=E)" and "bisim Vs E e (lcl s)" by(fastforce) with IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (Val v\F{D}:=e) \ length (lcl s)\ \\ E2 \dom X\\ E2 show ?case by(fastforce intro: FAssRed2) next case (CAS1Red1 e s ta e' s' D F e2 e3 Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\ \ \ ?concl e e' e2 s x ta s' e' vs\ from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs E2 _ (lcl s)\ obtain E E2' E2'' where E2: "E2 = E\compareAndSwap(D\F, E2', E2'')" "e2 = compE1 Vs E2'" "e3 = compE1 Vs E2''" and "bisim Vs E e (lcl s)" and sync: "\ contains_insync E2'" "\ contains_insync E2''" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars _ \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where IH': "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) show ?case proof(cases "is_val e2'") case True from E2 \fv E2 \ set Vs\ sync have "bisim Vs E2' e2 (lcl s')" "bisim Vs E2'' e3 (lcl s')" by auto with IH' E2 True sync show ?thesis by(cases "is_val E2'")(fastforce)+ next case False with IH' E2 sync show ?thesis by(fastforce) qed next case (CAS1Red2 e s ta e' s' v D F e3 Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs E2 _ (lcl s)\ obtain E E2' where E2: "E2 = (Val v\compareAndSwap(D\F, E, E2'))" "e3 = compE1 Vs E2'" and "bisim Vs E e (lcl s)" and sync: "\ contains_insync E2'" by(auto) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars _ \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) with E2 \fv E2 \ set Vs\ sync show ?case by(cases "is_val e2'")(fastforce)+ next case (CAS1Red3 e s ta e' s' v D F v' Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 _ (lcl s)\ obtain E where E2: "E2 = (Val v\compareAndSwap(D\F, Val v', E))" and "bisim Vs E e (lcl s)" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ E2 \length Vs + max_vars _ \ length (lcl s)\ \\ E2 \dom X\\ ultimately show ?case by(fastforce) next case (Call1Obj e s ta e' s' M es Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\ \ \ ?concl e e' e2 s x ta s' e' vs\ from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ \bisim Vs E2 (e\M(es)) (lcl s)\ obtain E es' where E2: "E2 = E\M(es')" "es = compEs1 Vs es'" and "bisim Vs E e (lcl s)" and sync: "\ contains_insyncs es'" by(auto elim!: bisim_cases simp add: compEs1_conv_map) with IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (e\M(es)) \ length (lcl s)\ \\ E2 \dom X\\ obtain TA' e2' x' where IH': "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(fastforce) with E2 \fv E2 \ set Vs\ \E2 = E\M(es')\ sync show ?case by(cases "is_val e2'")(auto intro: CallObj) next case (Call1Params es s ta es' s' v M Vs E2 X) note IH = \\vs es2 x. \ bisims vs es2 es (lcl s); fvs es2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_varss es \ length (lcl s); \s es2 \dom x\\ \ ?concls es es' es2 s x ta s' es' vs\ from \bisim Vs E2 (Val v\M(es)) (lcl s)\ obtain Es where "E2 = Val v \M(Es)" "bisims Vs Es es (lcl s)" by(auto) with IH[of Vs Es X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (Val v\M(es)) \ length (lcl s)\ \\ E2 \dom X\\ \E2 = Val v \M(Es)\ show ?case by(fastforce intro: CallParams) next case (Red1CallExternal s a T M Ts Tr D vs ta va h' e' s' Vs E2 X) from \bisim Vs E2 (addr a\M(map Val vs)) (lcl s)\ have E2: "E2 = addr a\M(map Val vs)" by auto moreover from \compP1 P \ class_type_of T sees M: Ts\Tr = Native in D\ have "P \ class_type_of T sees M: Ts\Tr = Native in D" by simp moreover from \compP1 P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ have "P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\" by simp moreover from wf \P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ have "ta_bisim01 (extTA2J0 P ta) (extTA2J1 (compP1 P) ta)" by(rule red_external_ta_bisim01) moreover note \typeof_addr (hp s) a = \T\\ \e' = extRet2J1 (addr a\M(map Val vs)) va\ \s' = (h', lcl s)\ moreover from \typeof_addr (hp s) a = \T\\ \P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ \P \ class_type_of T sees M: Ts\Tr = Native in D\ have "\external_defs D M \ ta = \ \ h' = hp s" by(fastforce dest: \external'_red_external_heap_unchanged \external'_red_external_TA_empty simp add: \external'_def \external_def) ultimately show ?case using \X \\<^sub>m [Vs [\] lcl s]\ by(fastforce intro!: exI[where x="extTA2J0 P ta"] intro: RedCallExternal simp add: map_eq_append_conv sim_move10_def synthesized_call_def dest: sees_method_fun del: disjCI intro: disjI1 disjI2) next case (Block1Red e h x ta e' h' x' V T Vs E2 X) note IH = \\vs e2 xa. \ bisim vs e2 e (lcl (h, x)); fv e2 \ set vs; xa \\<^sub>m [vs [\] lcl (h, x)]; length vs + max_vars e \ length (lcl (h, x)); \ e2 \dom xa\\ \ ?concl e e' e2 (h, x) xa ta (h', x') e' vs\ from \False,compP1 P,t \1 \e,(h, x)\ -ta\ \e',(h', x')\\ have "length x = length x'" by(auto dest: red1_preserves_len) with \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ have "length Vs < length x'" by simp from \bisim Vs E2 {V:T=None; e} (lcl (h, x))\ show ?case proof(cases rule: bisim_cases(14)[consumes 1, case_names BlockNone BlockSome BlockSomeNone]) case (BlockNone V' E) with \fv E2 \ set Vs\ have "fv E \ set (Vs@[V'])" by auto with IH[of "Vs@[V']" E "X(V' := None)"] BlockNone \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl (h, x)]\ \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \\ E2 \dom X\\ obtain TA' e2' X' where IH': "sim_move10 P t ta e e' E h (X(V' := None)) TA' e2' h' X'" "bisim (Vs @ [V']) e2' e' x'" "X' \\<^sub>m [Vs @ [V'] [\] x']" by(fastforce simp del: fun_upd_apply) { assume "V' \ set Vs" hence "hidden (Vs @ [V']) (index Vs V')" by(rule hidden_index) with \bisim (Vs @ [V']) E e (lcl (h, x))\ have "unmod e (index Vs V')" by(auto intro: hidden_bisim_unmod) moreover from \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \V' \ set Vs\ have "index Vs V' < length x" by(auto intro: index_less_aux) ultimately have "x ! index Vs V' = x' ! index Vs V'" using red1_preserves_unmod[OF \False,compP1 P,t \1 \e,(h, x)\ -ta\ \e',(h', x')\\] by(simp) } with \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \X' \\<^sub>m [Vs @ [V'] [\] x']\ \length x = length x'\ \X \\<^sub>m [Vs [\] lcl (h, x)]\ have rel: "X'(V' := X V') \\<^sub>m [Vs [\] x']" by(auto intro: Block_lem) show ?thesis proof(cases "X' V'") case None with BlockNone IH' rel show ?thesis by(fastforce intro: BlockRed) next case (Some v) with \X' \\<^sub>m [Vs @ [V'] [\] x']\ \length Vs < length x'\ have "x' ! length Vs = v" by(auto dest: map_le_SomeD) with BlockNone IH' rel Some show ?thesis by(fastforce intro: BlockRed) qed next case BlockSome thus ?thesis by simp next case (BlockSomeNone V' E) with \fv E2 \ set Vs\ have "fv E \ set (Vs@[V'])" by auto with IH[of "Vs@[V']" E "X(V' \ x ! length Vs)"] BlockSomeNone \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl (h, x)]\ \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \\ E2 \dom X\\ obtain TA' e2' X' where IH': "sim_move10 P t ta e e' E h (X(V' \ x ! length Vs)) TA' e2' h' X'" "bisim (Vs @ [V']) e2' e' x'" "X' \\<^sub>m [Vs @ [V'] [\] x']" by(fastforce simp del: fun_upd_apply) { assume "V' \ set Vs" hence "hidden (Vs @ [V']) (index Vs V')" by(rule hidden_index) with \bisim (Vs @ [V']) E e (lcl (h, x))\ have "unmod e (index Vs V')" by(auto intro: hidden_bisim_unmod) moreover from \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \V' \ set Vs\ have "index Vs V' < length x" by(auto intro: index_less_aux) ultimately have "x ! index Vs V' = x' ! index Vs V'" using red1_preserves_unmod[OF \False,compP1 P,t \1 \e,(h, x)\ -ta\ \e',(h', x')\\] by(simp) } with \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \X' \\<^sub>m [Vs @ [V'] [\] x']\ \length x = length x'\ \X \\<^sub>m [Vs [\] lcl (h, x)]\ have rel: "X'(V' := X V') \\<^sub>m [Vs [\] x']" by(auto intro: Block_lem) from \sim_move10 P t ta e e' E h (X(V' \ x ! length Vs)) TA' e2' h' X'\ obtain v' where "X' V' = \v'\" by(auto simp: sim_move10_def split: if_split_asm dest!: \red0t_lcl_incr \red0r_lcl_incr red_lcl_incr subsetD) with \X' \\<^sub>m [Vs @ [V'] [\] x']\ \length Vs < length x'\ have "x' ! length Vs = v'" by(auto dest: map_le_SomeD) with BlockSomeNone IH' rel \X' V' = \v'\\ show ?thesis by(fastforce intro: BlockRed) qed next case(Block1Some V xs T v e h) from \bisim vs e2 {V:T=\v\; e} (lcl (h, xs))\ obtain e' V' where "e2 = {V':T=\v\; e'}" and "V = length vs" "bisim (vs @ [V']) e' e (xs[length vs := v])" by(fastforce) moreover have "sim_move10 P t \ {length vs:T=\v\; e} {length vs:T=None; e} {V':T=\v\; e'} h x \ {V':T=\v\; e'} h x" by(auto) moreover from \bisim (vs @ [V']) e' e (xs[length vs := v])\ \length vs + max_vars {V:T=\v\; e} \ length (lcl (h, xs))\ have "bisim vs {V':T=\v\; e'} {length vs:T=None; e} (xs[length vs := v])" by auto moreover from \x \\<^sub>m [vs [\] lcl (h, xs)]\ \length vs + max_vars {V:T=\v\; e} \ length (lcl (h, xs))\ have "x \\<^sub>m [vs [\] xs[length vs := v]]" by auto ultimately show ?case by auto next case (Lock1Synchronized V xs a e h Vs E2 X) note len = \length Vs + max_vars (sync\<^bsub>V\<^esub> (addr a) e) \ length (lcl (h, xs))\ from \bisim Vs E2 (sync\<^bsub>V\<^esub> (addr a) e) (lcl (h, xs))\ obtain E where E2: "E2 = sync(addr a) E" "e = compE1 (Vs@[fresh_var Vs]) E" and sync: "\ contains_insync E" and [simp]: "V = length Vs" by auto moreover have "extTA2J0 P,P,t \ \sync(addr a) E, (h, X)\ -\Lock\a, SyncLock a\\ \insync(a) E, (h, X)\" by(rule LockSynchronized) moreover from \fv E2 \ set Vs\ fresh_var_fresh[of Vs] sync len have "bisim Vs (insync(a) E) (insync\<^bsub>length Vs\<^esub> (a) e) (xs[length Vs := Addr a])" unfolding \e = compE1 (Vs@[fresh_var Vs]) E\ \E2 = sync(addr a) E\ by -(rule bisimInSynchronized,rule compE1_bisim, auto) moreover have "zip Vs (xs[length Vs := Addr a]) = (zip Vs xs)[length Vs := (arbitrary, Addr a)]" by(rule sym)(simp add: update_zip) hence "zip Vs (xs[length Vs := Addr a]) = zip Vs xs" by simp with \X \\<^sub>m [Vs [\] (lcl (h, xs))]\ have "X \\<^sub>m [Vs [\] xs[length Vs := Addr a]]" by(auto simp add: map_le_def map_upds_def) ultimately show ?case by(fastforce intro: sim_move10_SyncLocks) next case (Synchronized1Red2 e s ta e' s' V a Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\ \ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) e) (lcl s)\ obtain E where E2: "E2 = insync(a) E" and bisim: "bisim (Vs@[fresh_var Vs]) E e (lcl s)" and xsa: "lcl s ! length Vs = Addr a" and [simp]: "V = length Vs" by auto with \fv E2 \ set Vs\ fresh_var_fresh[of Vs] have fv: "(fresh_var Vs) \ fv E" by auto from \length Vs + max_vars (insync\<^bsub>V\<^esub> (a) e) \ length (lcl s)\ have "length Vs < length (lcl s)" by simp { assume "X (fresh_var Vs) \ None" then obtain v where "X (fresh_var Vs) = \v\" by auto with \X \\<^sub>m [Vs [\] lcl s]\ have "[Vs [\] lcl s] (fresh_var Vs) = \v\" by(auto simp add: map_le_def dest: bspec) hence "(fresh_var Vs) \ set Vs" by(auto simp add: map_upds_def set_zip dest!: map_of_SomeD ) moreover have "(fresh_var Vs) \ set Vs" by(rule fresh_var_fresh) ultimately have False by contradiction } hence "X (fresh_var Vs) = None" by(cases "X (fresh_var Vs)", auto) hence "X(fresh_var Vs := None) = X" by(auto intro: ext) moreover from \X \\<^sub>m [Vs [\] lcl s]\ have "X(fresh_var Vs := None) \\<^sub>m [Vs [\] lcl s, (fresh_var Vs) \ (lcl s) ! length Vs]" by(simp) ultimately have "X \\<^sub>m [Vs @ [fresh_var Vs] [\] lcl s]" using \length Vs < length (lcl s)\ by(auto) moreover note IH[of "Vs@[fresh_var Vs]" E X] bisim E2 \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (insync\<^bsub>V\<^esub> (a) e) \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where IH': "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim (Vs @ [fresh_var Vs]) e2' e' (lcl s')" "x' \\<^sub>m [Vs @ [fresh_var Vs] [\] lcl s']" by auto hence "dom x' \ dom X \ fv E" by(fastforce iff del: domIff simp add: sim_move10_def dest: red_dom_lcl \red0r_dom_lcl[OF wf_prog_wwf_prog[OF wf]] \red0t_dom_lcl[OF wf_prog_wwf_prog[OF wf]] \red0r_fv_subset[OF wf_prog_wwf_prog[OF wf]] split: if_split_asm) with fv \X (fresh_var Vs) = None\ have "(fresh_var Vs) \ dom x'" by auto hence "x' (fresh_var Vs) = None" by auto moreover from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ have "length (lcl s) = length (lcl s')" by(auto dest: red1_preserves_len) moreover note \x' \\<^sub>m [Vs @ [fresh_var Vs] [\] lcl s']\ \length Vs < length (lcl s)\ ultimately have "x' \\<^sub>m [Vs [\] lcl s']" by(auto simp add: map_le_def dest: bspec) moreover from bisim fv have "unmod e (length Vs)" by(auto intro: bisim_fv_unmod) with \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ \length Vs < length (lcl s)\ have "lcl s ! length Vs = lcl s' ! length Vs" by(auto dest!: red1_preserves_unmod) with xsa have "lcl s' ! length Vs = Addr a" by simp ultimately show ?case using IH' E2 by(auto intro: SynchronizedRed2) next case (Unlock1Synchronized xs V a' a v h Vs E2 X) from \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) Val v) (lcl (h, xs))\ have E2: "E2 = insync(a) Val v" "V = length Vs" "xs ! length Vs = Addr a" by auto moreover with \xs ! V = Addr a'\ have [simp]: "a' = a" by simp have "extTA2J0 P,P,t \ \insync(a) (Val v), (h, X)\ -\Unlock\a, SyncUnlock a\\ \Val v, (h, X)\" by(rule UnlockSynchronized) ultimately show ?case using \X \\<^sub>m [Vs [\] lcl (h, xs)]\ by(fastforce intro: sim_move10_SyncLocks) next case (Unlock1SynchronizedNull xs V a v h Vs E2 X) from \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) Val v) (lcl (h, xs))\ have "V = length Vs" "xs ! length Vs = Addr a" by(auto) with \xs ! V = Null\ have False by simp thus ?case .. next case (Unlock1SynchronizedFail xs V A' a' v h Vs E2 X) from \False\ show ?case .. next case (Red1While b c s Vs E2 X) from \bisim Vs E2 (while (b) c) (lcl s)\ obtain B C where E2: "E2 = while (B) C" "b = compE1 Vs B" "c = compE1 Vs C" and sync: "\ contains_insync B" "\ contains_insync C" by auto moreover have "extTA2J0 P,P,t \ \while (B) C, (hp s, X)\ -\\ \if (B) (C;;while (B) C) else unit, (hp s, X)\" by(rule RedWhile) hence "sim_move10 P t \ (while (compE1 Vs B) (compE1 Vs C)) (if (compE1 Vs B) (compE1 Vs C;;while (compE1 Vs B) (compE1 Vs C)) else unit) (while (B) C) (hp s) X \ (if (B) (C;;while (B) C) else unit) (hp s) X" by(auto simp add: sim_move10_def) moreover from \fv E2 \ set Vs\ E2 sync have "bisim Vs (if (B) (C;; while (B) C) else unit) (if (compE1 Vs B) (compE1 Vs (C;; while(B) C)) else (compE1 Vs unit)) (lcl s)" by -(rule bisimCond, auto) ultimately show ?case using \X \\<^sub>m [Vs [\] lcl s]\ by(simp)(rule exI, rule exI, rule exI, erule conjI, auto) next case (Red1TryCatch h a D C V x e2 Vs E2 X) from \bisim Vs E2 (try Throw a catch(C V) e2) (lcl (h, x))\ obtain E2' V' where "E2 = try Throw a catch(C V') E2'" "V = length Vs" "e2 = compE1 (Vs @ [V']) E2'" and sync: "\ contains_insync E2'" by(auto) with \fv E2 \ set Vs\ have "fv E2' \ set (Vs @[V'])" by auto with \e2 = compE1 (Vs @ [V']) E2'\ sync have "bisim (Vs @[V']) E2' e2 (x[V := Addr a])" by(auto intro!: compE1_bisim) with \V = length Vs\ \length Vs + max_vars (try Throw a catch(C V) e2) \ length (lcl (h, x))\ have "bisim Vs {V':Class C=\Addr a\; E2'} {length Vs:Class C=None; e2} (x[V := Addr a])" by(auto intro: bisimBlockSomeNone) moreover from \length Vs + max_vars (try Throw a catch(C V) e2) \ length (lcl (h, x))\ have "[Vs [\] x[length Vs := Addr a]] = [Vs [\] x]" by simp with \X \\<^sub>m [Vs [\] lcl (h, x)]\ have "X \\<^sub>m [Vs [\] x[length Vs := Addr a]]" by simp moreover note \e2 = compE1 (Vs @ [V']) E2'\ \E2 = try Throw a catch(C V') E2'\ \typeof_addr h a = \Class_type D\\ \compP1 P \ D \\<^sup>* C\ \V = length Vs\ ultimately show ?case by(auto intro!: exI) next case Red1TryFail thus ?case by(auto intro!: exI sim_move10_TryFail) next case (List1Red1 e s ta e' s' es Vs ES2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ \TA' e2' x'. sim_move10 P t ta e e' e2 (hp s) x TA' e2' (hp s') x' \ bisim vs e2' e' (lcl s') \ x' \\<^sub>m [vs [\] lcl s']\ from \bisims Vs ES2 (e # es) (lcl s)\ \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ obtain E ES where "ES2 = E # ES" "\ is_val E" "es = compEs1 Vs ES" "bisim Vs E e (lcl s)" and sync: "\ contains_insyncs ES" by(auto elim!: bisims_cases simp add: compEs1_conv_map) with IH[of Vs E X] \fvs ES2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_varss (e # es) \ length (lcl s)\ \\s ES2 \dom X\\ obtain TA' e2' x' where IH': "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by fastforce show ?case proof(cases "is_val e2'") case False with IH' \ES2 = E # ES\ \es = compEs1 Vs ES\ sync show ?thesis by(auto intro: sim_moves10_expr) next case True from \fvs ES2 \ set Vs\ \ES2 = E # ES\ \es = compEs1 Vs ES\ sync have "bisims Vs ES es (lcl s')" by(auto intro: compEs1_bisims) with IH' True \ES2 = E # ES\ \es = compEs1 Vs ES\ show ?thesis by(auto intro: sim_moves10_expr) qed next case (List1Red2 es s ta es' s' v Vs ES2 X) note IH = \\vs es2 x. \bisims vs es2 es (lcl s); fvs es2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_varss es \ length (lcl s); \s es2 \dom x\\ \ \TA' es2' x'. sim_moves10 P t ta es es' es2 (hp s) x TA' es2' (hp s') x' \ bisims vs es2' es' (lcl s') \ x' \\<^sub>m [vs [\] lcl s']\ from \bisims Vs ES2 (Val v # es) (lcl s)\ obtain ES where "ES2 = Val v # ES" "bisims Vs ES es (lcl s)" by(auto elim!: bisims_cases) with IH[of Vs ES X] \fvs ES2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_varss (Val v # es) \ length (lcl s)\ \\s ES2 \dom X\\ \ES2 = Val v # ES\ show ?case by(fastforce intro: sim_moves10_expr) next case Call1ThrowParams thus ?case by(fastforce intro: CallThrowParams elim!: bisim_cases simp add: bisims_map_Val_Throw2) next case (Synchronized1Throw2 xs V a' a ad h Vs E2 X) from \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) Throw ad) (lcl (h, xs))\ have "xs ! length Vs = Addr a" and "V = length Vs" by auto with \xs ! V = Addr a'\ have [simp]: "a' = a" by simp have "extTA2J0 P,P,t \ \insync(a) Throw ad, (h, X)\ -\Unlock\a, SyncUnlock a\\ \Throw ad, (h, X)\" by(rule SynchronizedThrow2) with \X \\<^sub>m [Vs [\] lcl (h, xs)]\ \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) Throw ad) (lcl (h, xs))\ show ?case by(auto intro: sim_move10_SyncLocks intro!: exI) next case (Synchronized1Throw2Null xs V a a' h Vs E2 X) from \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) Throw a') (lcl (h, xs))\ have "V = length Vs" "xs ! length Vs = Addr a" by(auto) with \xs ! V = Null\ have False by simp thus ?case .. next case (Synchronized1Throw2Fail xs V A' a' a h Vs E2 X) from \False\ show ?case .. next case InstanceOf1Red thus ?case by auto(blast) next case Red1InstanceOf thus ?case by hypsubst_thin auto next case InstanceOf1Throw thus ?case by auto next case CAS1Throw thus ?case by fastforce next case CAS1Throw2 thus ?case by fastforce next case CAS1Throw3 thus ?case by fastforce qed(simp_all del: fun_upd_apply, (fastforce intro: red_reds.intros simp del: fun_upd_apply simp add: finfun_upd_apply)+) lemma bisim_call_Some_call1: "\ bisim Vs e e' xs; call e = \aMvs\; length Vs + max_vars e' \ length xs \ \ \e'' xs'. \red1'r P t h (e', xs) (e'', xs') \ call1 e'' = \aMvs\ \ bisim Vs e e'' xs' \ take (length Vs) xs = take (length Vs) xs'" and bisims_calls_Some_calls1: "\ bisims Vs es es' xs; calls es = \aMvs\; length Vs + max_varss es' \ length xs \ \ \es'' xs'. \reds1'r P t h (es', xs) (es'', xs') \ calls1 es'' = \aMvs\ \ bisims Vs es es'' xs' \ take (length Vs) xs = take (length Vs) xs'" proof(induct rule: bisim_bisims.inducts) case bisimCallParams thus ?case by(fastforce simp add: is_vals_conv split: if_split_asm) next case bisimBlockNone thus ?case by(fastforce intro: take_eq_take_le_eq) next case (bisimBlockSome Vs V e e' xs v T) from \length Vs + max_vars {length Vs:T=\v\; e'} \ length xs\ have "\red1'r P t h ({length Vs:T=\v\; e'}, xs) ({length Vs:T=None; e'}, xs[length Vs := v])" by(auto intro!: \red1r_1step Block1Some) also from bisimBlockSome obtain e'' xs' where "\red1'r P t h (e', xs[length Vs := v]) (e'', xs')" and "call1 e'' = \aMvs\" "bisim (Vs @ [V]) e e'' xs'" and "take (length (Vs @ [V])) (xs[length Vs := v]) = take (length (Vs @ [V])) xs'" by auto hence "\red1'r P t h ({length Vs:T=None; e'}, xs[length Vs := v]) ({length Vs:T=None; e''}, xs')" by auto also from \call1 e'' = \aMvs\\ have "call1 {length Vs:T=None; e''} = \aMvs\" by simp moreover from \take (length (Vs @ [V])) (xs[length Vs := v]) = take (length (Vs @ [V])) xs'\ have "take (length Vs) xs = take (length Vs) xs'" - by(auto dest: take_eq_take_le_eq[where m="length Vs"] simp add: take_list_update_beyond) + by(auto dest: take_eq_take_le_eq[where m="length Vs"] simp add: take_update_cancel) moreover { have "xs' ! length Vs = take (length (Vs @ [V])) xs' ! length Vs" by simp also note \take (length (Vs @ [V])) (xs[length Vs := v]) = take (length (Vs @ [V])) xs'\[symmetric] also have "take (length (Vs @ [V])) (xs[length Vs := v]) ! length Vs = v" using \length Vs + max_vars {length Vs:T=\v\; e'} \ length xs\ by simp finally have "bisim Vs {V:T=\v\; e} {length Vs:T=None; e''} xs'" using \bisim (Vs @ [V]) e e'' xs'\ by auto } ultimately show ?case by blast next case (bisimBlockSomeNone Vs V e e' xs v T) then obtain e'' xs' where "\red1'r P t h (e', xs) (e'', xs')" "call1 e'' = \aMvs\" "bisim (Vs @ [V]) e e'' xs'" and "take (length (Vs @ [V])) xs = take (length (Vs @ [V])) xs'" by auto hence "\red1'r P t h ({length Vs:T=None; e'}, xs) ({length Vs:T=None; e''}, xs')" by auto moreover from \call1 e'' = \aMvs\\ have "call1 ({length Vs:T=None; e''}) = \aMvs\" by simp moreover from \take (length (Vs @ [V])) xs = take (length (Vs @ [V])) xs'\ have "take (length Vs) xs = take (length Vs) xs'" by(auto intro: take_eq_take_le_eq) moreover { have "xs' ! length Vs = take (length (Vs @ [V])) xs' ! length Vs" by simp also note \take (length (Vs @ [V])) xs = take (length (Vs @ [V])) xs'\[symmetric] also have "take (length (Vs @ [V])) xs ! length Vs = v" using \xs ! length Vs = v\ by simp finally have "bisim Vs {V:T=\v\; e} {length Vs:T=None; e''} xs'" using \bisim (Vs @ [V]) e e'' xs'\ by auto } ultimately show ?case by blast next case (bisimInSynchronized Vs e e' xs a) then obtain e'' xs' where "\red1'r P t h (e', xs) (e'', xs')" "call1 e'' = \aMvs\" "bisim (Vs @ [fresh_var Vs]) e e'' xs'" and "take (Suc (length Vs)) xs = take (Suc (length Vs)) xs'" by auto hence "\red1'r P t h (insync\<^bsub>length Vs\<^esub> (a) e', xs) (insync\<^bsub>length Vs\<^esub> (a) e'', xs')" by auto moreover from \call1 e'' = \aMvs\\ have "call1 (insync\<^bsub>length Vs\<^esub> (a) e'') = \aMvs\" by simp moreover from \take (Suc (length Vs)) xs = take (Suc (length Vs)) xs'\ have "take (length Vs) xs = take (length Vs) xs'" by(auto intro: take_eq_take_le_eq) moreover { have "xs' ! length Vs = take (Suc (length Vs)) xs' ! length Vs" by simp also note \take (Suc (length Vs)) xs = take (Suc (length Vs)) xs'\[symmetric] also have "take (Suc (length Vs)) xs ! length Vs = Addr a" using \xs ! length Vs = Addr a\ by simp finally have "bisim Vs (insync(a) e) (insync\<^bsub>length Vs\<^esub> (a) e'') xs'" using \bisim (Vs @ [fresh_var Vs]) e e'' xs'\ by auto } ultimately show ?case by blast next case bisimsCons1 thus ?case by(fastforce intro!: \red1r_inj_\reds1r) next case bisimsCons2 thus ?case by(fastforce intro!: \reds1r_cons_\reds1r) qed fastforce+ lemma sim_move01_into_Red1: "sim_move01 P t ta e E' h xs ta' e2' h' xs' \ if \Move0 P h (e, es1) then \Red1't P t h ((E', xs), exs2) ((e2', xs'), exs2) \ ta = \ \ h = h' else \ex2' exs2' ta'. \Red1'r P t h ((E', xs), exs2) (ex2', exs2') \ (call e = None \ call1 E' = None \ ex2' = (E', xs) \ exs2' = exs2) \ False,P,t \1 \ex2'/exs2',h\ -ta'\ \(e2', xs')/exs2,h'\ \ \ \Move1 P h (ex2', exs2') \ ta_bisim01 ta ta'" apply(auto simp add: sim_move01_def intro: \red1t_into_\Red1t \red1r_into_\Red1r red1Red split: if_split_asm) apply(fastforce intro: red1Red intro!: \red1r_into_\Red1r)+ done lemma sim_move01_max_vars_decr: "sim_move01 P t ta e0 e h xs ta' e' h' xs' \ max_vars e' \ max_vars e" by(fastforce simp add: sim_move01_def split: if_split_asm dest: \red1t_max_vars red1_max_vars_decr \red1r_max_vars) lemma Red1_simulates_red0: assumes wf: "wf_J_prog P" and red: "P,t \0 \e1/es1, h\ -ta\ \e1'/es1', h'\" and bisiml: "bisim_list1 (e1, es1) (ex2, exs2)" shows "\ex2'' exs2''. bisim_list1 (e1', es1') (ex2'', exs2'') \ (if \Move0 P h (e1, es1) then \Red1't (compP1 P) t h (ex2, exs2) (ex2'', exs2'') \ ta = \ \ h = h' else \ex2' exs2' ta'. \Red1'r (compP1 P) t h (ex2, exs2) (ex2', exs2') \ (call e1 = None \ call1 (fst ex2) = None \ ex2' = ex2 \ exs2' = exs2) \ False,compP1 P,t \1 \ex2'/exs2', h\ -ta'\ \ex2''/exs2'', h'\ \ \ \Move1 P h (ex2', exs2') \ ta_bisim01 ta ta')" (is "\ex2'' exs2'' . _ \ ?red ex2'' exs2''") using red proof(cases) case (red0Red XS') note [simp] = \es1' = es1\ and red = \extTA2J0 P,P,t \ \e1,(h, Map.empty)\ -ta\ \e1',(h', XS')\\ and notsynth = \\aMvs. call e1 = \aMvs\ \ synthesized_call P h aMvs\ from bisiml obtain E xs where ex2: "ex2 = (E, xs)" and bisim: "bisim [] e1 E xs" and fv: "fv e1 = {}" and length: "max_vars E \ length xs" and bsl: "bisim_list es1 exs2" and D: "\ e1 \{}\" by(auto elim!: bisim_list1_elim) from bisim_max_vars[OF bisim] length red1_simulates_red_aux[OF wf red bisim] fv notsynth obtain ta' e2' xs' where sim: "sim_move01 (compP1 P) t ta e1 E h xs ta' e2' h' xs'" and bisim': "bisim [] e1' e2' xs'" and XS': "XS' \\<^sub>m Map.empty" by auto from sim_move01_into_Red1[OF sim, of es1 exs2] have "?red (e2', xs') exs2" unfolding ex2 by auto moreover { note bsl bisim' moreover from fv red_fv_subset[OF wf_prog_wwf_prog[OF wf] red] have "fv e1' = {}" by simp moreover from red D have "\ e1' \dom XS'\" by(auto dest: red_preserves_defass[OF wf] split: if_split_asm) with red_dom_lcl[OF red] \fv e1 = {}\ have "\ e1' \{}\" by(auto elim!: D_mono' simp add: hyperset_defs) moreover from sim have "length xs = length xs'" "max_vars e2' \ max_vars E" by(auto dest: sim_move01_preserves_len sim_move01_max_vars_decr) with length have length': "max_vars e2' \ length xs'" by(auto) ultimately have "bisim_list1 (e1', es1) ((e2', xs'), exs2)" by(rule bisim_list1I) } ultimately show ?thesis using ex2 by(simp split del: if_split)(rule exI conjI|assumption)+ next case (red0Call a M vs U Ts T pns body D) note [simp] = \ta = \\ \h' = h\ and es1' = \es1' = e1 # es1\ and e1' = \e1' = blocks (this # pns) (Class D # Ts) (Addr a # vs) body\ and call = \call e1 = \(a, M, vs)\\ and ha = \typeof_addr h a = \U\\ and sees = \P \ class_type_of U sees M: Ts\T = \(pns, body)\ in D\ and len = \length vs = length pns\ \length Ts = length pns\ from bisiml obtain E xs where ex2: "ex2 = (E, xs)" and bisim: "bisim [] e1 E xs" and fv: "fv e1 = {}" and length: "max_vars E \ length xs" and bsl: "bisim_list es1 exs2" and D: "\ e1 \{}\" by(auto elim!: bisim_list1_elim) from bisim_call_Some_call1[OF bisim call, of "compP1 P" t h] length obtain e' xs' where red: "\red1'r (compP1 P) t h (E, xs) (e', xs')" and "call1 e' = \(a, M, vs)\" "bisim [] e1 e' xs'" and "take 0 xs = take 0 xs'" by auto let ?e1' = "blocks (this # pns) (Class D # Ts) (Addr a # vs) body" let ?e2' = "blocks1 0 (Class D#Ts) (compE1 (this # pns) body)" let ?xs2' = "Addr a # vs @ replicate (max_vars (compE1 (this # pns) body)) undefined_value" let ?exs2' = "(e', xs') # exs2" have "\Red1'r (compP1 P) t h ((E, xs), exs2) ((e', xs'), exs2)" using red by(rule \red1r_into_\Red1r) moreover { note \call1 e' = \(a, M, vs)\\ moreover note ha moreover have "compP1 P \ class_type_of U sees M:Ts \ T = map_option (\(pns, body). compE1 (this # pns) body) \(pns, body)\ in D" using sees unfolding compP1_def by(rule sees_method_compP) hence sees': "compP1 P \ class_type_of U sees M:Ts \ T = \compE1 (this # pns) body\ in D" by simp moreover from len have "length vs = length Ts" by simp ultimately have "False,compP1 P,t \1 \(e', xs')/exs2,h\ -\\ \(?e2', ?xs2')/?exs2', h\" by(rule red1Call) moreover have "\Move1 P h ((e', xs'), exs2)" using \call1 e' = \(a, M, vs)\\ ha sees by(auto simp add: synthesized_call_def \external'_def dest: sees_method_fun dest!: \move1_not_call1[where P=P and h=h]) ultimately have "\Red1' (compP1 P) t h ((e', xs'), exs2) ((?e2', ?xs2'), ?exs2')" by auto } ultimately have "\Red1't (compP1 P) t h ((E, xs), exs2) ((?e2', ?xs2'), ?exs2')" by(rule rtranclp_into_tranclp1) moreover { from red have "length xs' = length xs" by(rule \red1r_preserves_len) moreover from red have "max_vars e' \ max_vars E" by(rule \red1r_max_vars) ultimately have "max_vars e' \ length xs'" using length by simp } with bsl \bisim [] e1 e' xs'\ fv D have "bisim_list (e1 # es1) ?exs2'" using \call e1 = \(a, M, vs)\\ \call1 e' = \(a, M, vs)\\ by(rule bisim_listCons) hence "bisim_list1 (e1', es1') ((?e2', ?xs2'), ?exs2')" unfolding e1' es1' proof(rule bisim_list1I) from wf_prog_wwf_prog[OF wf] sees have "wf_mdecl wwf_J_mdecl P D (M,Ts,T,\(pns,body)\)" by(rule sees_wf_mdecl) hence fv': "fv body \ set pns \ {this}" by(auto simp add: wf_mdecl_def) moreover from \P \ class_type_of U sees M: Ts\T = \(pns, body)\ in D\ have "\ contains_insync body" by(auto dest!: sees_wf_mdecl[OF wf] WT_expr_locks simp add: wf_mdecl_def contains_insync_conv) ultimately have "bisim ([this] @ pns) body (compE1 ([this] @ pns) body) ?xs2'" by -(rule compE1_bisim, auto) with \length vs = length pns\ \length Ts = length pns\ have "bisim ([] @ [this]) (blocks pns Ts vs body) (blocks1 (Suc 0) Ts (compE1 (this # pns) body)) ?xs2'" by -(drule blocks_bisim,auto simp add: nth_append) from bisimBlockSomeNone[OF this, of "Addr a" "Class D"] show "bisim [] ?e1' ?e2' ?xs2'" by simp from fv' len show "fv ?e1' = {}" by auto from wf sees have "wf_mdecl wf_J_mdecl P D (M,Ts,T,\(pns,body)\)" by(rule sees_wf_mdecl) hence "\ body \set pns \ {this}\" by(auto simp add: wf_mdecl_def) with \length vs = length pns\ \length Ts = length pns\ have "\ (blocks pns Ts vs body) \dom [this \ Addr a]\" by(auto) thus "\ ?e1' \{}\" by auto from len show "max_vars ?e2' \ length ?xs2'" by(auto simp add: blocks1_max_vars) qed moreover have "\Move0 P h (e1, es1)" using call ha sees by(fastforce simp add: synthesized_call_def \external'_def dest: sees_method_fun \move0_callD[where P=P and h=h]) ultimately show ?thesis using ex2 \call e1 = \(a, M, vs)\\ by(simp del: \Move1.simps)(rule exI conjI|assumption)+ next case (red0Return e) note es1 = \es1 = e # es1'\ and e1' = \e1' = inline_call e1 e\ and [simp] = \ta = \\ \h' = h\ and fin = \final e1\ from bisiml es1 obtain E' xs' E xs exs' aMvs where ex2: "ex2 = (E', xs')" and exs2: "exs2 = (E, xs) # exs'" and bisim': "bisim [] e1 E' xs'" and bisim: "bisim [] e E xs" and fv: "fv e = {}" and length: "max_vars E \ length xs" and bisiml: "bisim_list es1' exs'" and D: "\ e \{}\" and call: "call e = \aMvs\" and call1: "call1 E = \aMvs\" by(fastforce elim: bisim_list1_elim) let ?e2' = "inline_call E' E" from \final e1\ bisim' have "final E'" by(auto) hence red': "False,compP1 P,t \1 \ex2/exs2, h\ -\\ \(?e2', xs)/exs', h\" unfolding ex2 exs2 by(rule red1Return) moreover have "\Move0 P h (e1, es1) = \Move1 P h ((E', xs'), exs2)" using \final e1\ \final E'\ by auto moreover { note bisiml moreover from bisim' fin bisim have "bisim [] (inline_call e1 e) (inline_call E' E) xs" using call by(rule bisim_inline_call)(simp add: fv) moreover from fv_inline_call[of e1 e] fv fin have "fv (inline_call e1 e) = {}" by auto moreover from fin have "\ (inline_call e1 e) \{}\" using call D by(rule defass_inline_call) moreover have "max_vars ?e2' \ max_vars E + max_vars E'" by(rule inline_call_max_vars1) with \final E'\ length have "max_vars ?e2' \ length xs" by(auto elim!: final.cases) ultimately have "bisim_list1 (inline_call e1 e, es1') ((?e2', xs), exs')" by(rule bisim_list1I) } ultimately show ?thesis using e1' \final e1\ \final E'\ ex2 apply(simp del: \Move0.simps \Move1.simps) apply(rule exI conjI impI|assumption)+ apply(rule tranclp.r_into_trancl, simp) apply blast done qed lemma sim_move10_into_red0: assumes wwf: "wwf_J_prog P" and sim: "sim_move10 P t ta e2 e2' e h Map.empty ta' e' h' x'" and fv: "fv e = {}" shows "if \move1 P h e2 then (\Red0t P t h (e, es) (e', es) \ countInitBlock e2' < countInitBlock e2 \ e' = e \ x' = Map.empty) \ ta = \ \ h = h' else \e'' ta'. \Red0r P t h (e, es) (e'', es) \ (call1 e2 = None \ call e = None \ e'' = e) \ P,t \0 \e''/es,h\ -ta'\ \e'/es,h'\ \ \ \Move0 P h (e'', es) \ ta_bisim01 ta' (extTA2J1 (compP1 P) ta)" proof(cases "\move1 P h e2") case True with sim have "\ final e2" and red: "\red0t (extTA2J0 P) P t h (e, Map.empty) (e', x') \ countInitBlock e2' < countInitBlock e2 \ e' = e \ x' = Map.empty" and [simp]: "h' = h" "ta = \" "ta' = \" by(simp_all add: sim_move10_def) from red have "\Red0t P t h (e, es) (e', es) \ countInitBlock e2' < countInitBlock e2 \ e' = e \ x' = Map.empty" proof assume red: "\red0t (extTA2J0 P) P t h (e, Map.empty) (e', x')" from \red0t_fv_subset[OF wwf red] \red0t_dom_lcl[OF wwf red] fv have "dom x' \ {}" by(auto split: if_split_asm) hence "x' = Map.empty" by auto with red have "\red0t (extTA2J0 P) P t h (e, Map.empty) (e', Map.empty)" by simp with wwf have "\Red0t P t h (e, es) (e', es)" using fv by(rule \red0t_into_\Red0t) thus ?thesis .. qed simp with True show ?thesis by simp next case False with sim obtain e'' xs'' where "\ final e2" and \red: "\red0r (extTA2J0 P) P t h (e, Map.empty) (e'', xs'')" and red: "extTA2J0 P,P,t \ \e'',(h, xs'')\ -ta'\ \e',(h', x')\" and call: "call1 e2 = None \ call e = None \ e'' = e" and "\ \move0 P h e''" "ta_bisim01 ta' (extTA2J1 (compP1 P) ta)" "no_call P h e''" by(auto simp add: sim_move10_def split: if_split_asm) from \red0r_fv_subset[OF wwf \red] \red0r_dom_lcl[OF wwf \red] fv have "dom xs'' \ {}" by(auto) hence "xs'' = Map.empty" by(auto) with \red have "\red0r (extTA2J0 P) P t h (e, Map.empty) (e'', Map.empty)" by simp with wwf have "\Red0r P t h (e, es) (e'', es)" using fv by(rule \red0r_into_\Red0r) moreover from red \xs'' = Map.empty\ have "extTA2J0 P,P,t \ \e'',(h, Map.empty)\ -ta'\ \e',(h', x')\" by simp from red0Red[OF this] \no_call P h e''\ have "P,t \0 \e''/es,h\ -ta'\ \e'/es,h'\" by(simp add: no_call_def) moreover from \\ \move0 P h e''\ red have "\ \Move0 P h (e'', es)" by auto ultimately show ?thesis using False \ta_bisim01 ta' (extTA2J1 (compP1 P) ta)\ call by(simp del: \Move0.simps) blast qed lemma red0_simulates_Red1: assumes wf: "wf_J_prog P" and red: "False,compP1 P,t \1 \ex2/exs2, h\ -ta\ \ex2'/exs2', h'\" and bisiml: "bisim_list1 (e, es) (ex2, exs2)" shows "\e' es'. bisim_list1 (e', es') (ex2', exs2') \ (if \Move1 P h (ex2, exs2) then (\Red0t P t h (e, es) (e', es') \ countInitBlock (fst ex2') < countInitBlock (fst ex2) \ exs2' = exs2 \ e' = e \ es' = es) \ ta = \ \ h = h' else \e'' es'' ta'. \Red0r P t h (e, es) (e'', es'') \ (call1 (fst ex2) = None \ call e = None \ e'' = e \ es'' = es) \ P,t \0 \e''/es'', h\ -ta'\ \e'/es', h'\ \ \ \Move0 P h (e'', es'') \ ta_bisim01 ta' ta)" (is "\e' es' . _ \ ?red e' es'") using red proof(cases) case (red1Red E xs TA E' xs') note red = \False,compP1 P,t \1 \E,(h, xs)\ -TA\ \E',(h', xs')\\ and ex2 = \ex2 = (E, xs)\ and ex2' = \ex2' = (E', xs')\ and [simp] = \ta = extTA2J1 (compP1 P) TA\ \exs2' = exs2\ from bisiml ex2 have bisim: "bisim [] e E xs" and fv: "fv e = {}" and length: "max_vars E \ length xs" and bsl: "bisim_list es exs2" and D: "\ e \{}\" by(auto elim: bisim_list1_elim) from red_simulates_red1_aux[OF wf red, simplified, OF bisim, of Map.empty] fv length D obtain TA' e2' x' where red': "sim_move10 P t TA E E' e h Map.empty TA' e2' h' x'" and bisim'': "bisim [] e2' E' xs'" and lcl': "x' \\<^sub>m Map.empty" by auto from red have "\ final E" by auto with sim_move10_into_red0[OF wf_prog_wwf_prog[OF wf] red', of es] fv ex2 ex2' have red'': "?red e2' es" by fastforce moreover { note bsl bisim'' moreover from red' fv have "fv e2' = {}" by(fastforce simp add: sim_move10_def split: if_split_asm dest: \red0r_fv_subset[OF wf_prog_wwf_prog[OF wf]] \red0t_fv_subset[OF wf_prog_wwf_prog[OF wf]] red_fv_subset[OF wf_prog_wwf_prog[OF wf]]) moreover from red' have "dom x' \ dom (Map.empty) \ fv e" unfolding sim_move10_def apply(auto split: if_split_asm del: subsetI dest: \red0r_dom_lcl[OF wf_prog_wwf_prog[OF wf]] \red0t_dom_lcl[OF wf_prog_wwf_prog[OF wf]]) apply(frule_tac [1-2] \red0r_fv_subset[OF wf_prog_wwf_prog[OF wf]]) apply(auto dest!: \red0r_dom_lcl[OF wf_prog_wwf_prog[OF wf]] red_dom_lcl del: subsetI, blast+) done with fv have "dom x' \ {}" by(auto) hence "x' = Map.empty" by(auto) with D red' have "\ e2' \{}\" by(auto dest!: sim_move10_preserves_defass[OF wf] split: if_split_asm) moreover from red have "length xs' = length xs" by(auto dest: red1_preserves_len) with red1_max_vars[OF red] length have "max_vars E' \ length xs'" by simp ultimately have "bisim_list1 (e2', es) ((E', xs'), exs2)" by(rule bisim_list1I) } ultimately show ?thesis using ex2' by(clarsimp split: if_split_asm)(rule exI conjI|assumption|simp)+ next case (red1Call E a M vs U Ts T body D xs) note [simp] = \ex2 = (E, xs)\ \h' = h\ \ta = \\ and ex2' = \ex2' = (blocks1 0 (Class D#Ts) body, Addr a # vs @ replicate (max_vars body) undefined_value)\ and exs' = \exs2' = (E, xs) # exs2\ and call = \call1 E = \(a, M, vs)\\ and ha = \typeof_addr h a = \U\\ and sees = \compP1 P \ class_type_of U sees M: Ts\T = \body\ in D\ and len = \length vs = length Ts\ let ?e2' = "blocks1 0 (Class D#Ts) body" let ?xs2' = "Addr a # vs @ replicate (max_vars body) undefined_value" from bisiml have bisim: "bisim [] e E xs" and fv: "fv e = {}" and length: "max_vars E \ length xs" and bsl: "bisim_list es exs2" and D: "\ e \{}\" by(auto elim: bisim_list1_elim) from bisim \call1 E = \(a, M, vs)\\ have "call e = \(a, M, vs)\" by(rule bisim_call1_Some_call) moreover note ha moreover from \compP1 P \ class_type_of U sees M: Ts\T = \body\ in D\ obtain pns Jbody where sees': "P \ class_type_of U sees M: Ts\T = \(pns, Jbody)\ in D" and body: "body = compE1 (this # pns) Jbody" by(auto dest: sees_method_compPD) let ?e' = "blocks (this # pns) (Class D # Ts) (Addr a # vs) Jbody" note sees' moreover from wf sees' have "length Ts = length pns" by(auto dest!: sees_wf_mdecl simp add: wf_mdecl_def) with len have "length vs = length pns" "length Ts = length pns" by simp_all ultimately have red': "P,t \0 \e/es, h\ -\\ \?e'/e#es, h\" by(rule red0Call) moreover from \call e = \(a, M, vs)\\ ha sees' have "\Move0 P h (e, es)" by(fastforce simp add: synthesized_call_def dest: \move0_callD[where P=P and h=h] sees_method_fun) ultimately have "\Red0t P t h (e, es) (?e', e#es)" by auto moreover from bsl bisim fv D length \call e = \(a, M, vs)\\ \call1 E = \(a, M, vs)\\ have "bisim_list (e # es) ((E, xs) # exs2)" by(rule bisim_list.intros) hence "bisim_list1 (?e', e # es) (ex2', (E, xs) # exs2)" unfolding ex2' proof(rule bisim_list1I) from wf_prog_wwf_prog[OF wf] sees' have "wf_mdecl wwf_J_mdecl P D (M,Ts,T,\(pns,Jbody)\)" by(rule sees_wf_mdecl) hence "fv Jbody \ set pns \ {this}" by(auto simp add: wf_mdecl_def) moreover from sees' have "\ contains_insync Jbody" by(auto dest!: sees_wf_mdecl[OF wf] WT_expr_locks simp add: wf_mdecl_def contains_insync_conv) ultimately have "bisim ([] @ this # pns) Jbody (compE1 ([] @ this # pns) Jbody) ?xs2'" by -(rule compE1_bisim, auto) with \length vs = length Ts\ \length Ts = length pns\ body have "bisim [] ?e' (blocks1 (length ([] :: vname list)) (Class D # Ts) body) ?xs2'" by -(rule blocks_bisim, auto simp add: nth_append nth_Cons') thus "bisim [] ?e' ?e2' ?xs2'" by simp from \length vs = length Ts\ \length Ts = length pns\ \fv Jbody \ set pns \ {this}\ show "fv ?e' = {}" by auto from wf sees' have "wf_mdecl wf_J_mdecl P D (M,Ts,T,\(pns,Jbody)\)" by(rule sees_wf_mdecl) hence "\ Jbody \set pns \ {this}\" by(auto simp add: wf_mdecl_def) with \length vs = length Ts\ \length Ts = length pns\ have "\ (blocks pns Ts vs Jbody) \dom [this \ Addr a]\" by(auto) thus "\ ?e' \{}\" by simp from len show "max_vars ?e2' \ length ?xs2'" by(simp add: blocks1_max_vars) qed moreover have "\Move1 P h (ex2, exs2)" using ha \call1 E = \(a, M, vs)\\ sees' by(auto simp add: synthesized_call_def \external'_def dest!: \move1_not_call1[where P=P and h=h] dest: sees_method_fun) ultimately show ?thesis using exs' by(simp del: \Move1.simps \Move0.simps)(rule exI conjI rtranclp.rtrancl_refl|assumption)+ next case (red1Return E' x' E x) note [simp] = \h' = h\ \ta = \\ and ex2 = \ex2 = (E', x')\ and exs2 = \exs2 = (E, x) # exs2'\ and ex2' = \ex2' = (inline_call E' E, x)\ and fin = \final E'\ from bisiml ex2 exs2 obtain e' es' aMvs where es: "es = e' # es'" and bsl: "bisim_list es' exs2'" and bisim: "bisim [] e E' x'" and bisim': "bisim [] e' E x" and fv: "fv e' = {}" and length: "max_vars E \ length x" and D: "\ e' \{}\" and call: "call e' = \aMvs\" and call1: "call1 E = \aMvs\" by(fastforce elim!: bisim_list1_elim) from \final E'\ bisim have fin': "final e" by(auto) hence "P,t \0 \e/e' # es',h\ -\\ \inline_call e e'/es',h\" by(rule red0Return) moreover from bisim fin' bisim' call have "bisim [] (inline_call e e') (inline_call E' E) x" by(rule bisim_inline_call)(simp add: fv) with bsl have "bisim_list1 (inline_call e e', es') (ex2', exs2')" unfolding ex2' proof(rule bisim_list1I) from fin' fv_inline_call[of e e'] fv show "fv (inline_call e e') = {}" by auto from fin' show "\ (inline_call e e') \{}\" using call D by(rule defass_inline_call) from call1_imp_call[OF call1] have "max_vars (inline_call E' E) \ max_vars E + max_vars E'" by(rule inline_call_max_vars) with fin length show "max_vars (inline_call E' E) \ length x" by(auto elim!: final.cases) qed moreover have "\Move1 P h (ex2, exs2) = \Move0 P h (e, es)" using ex2 call1 call fin fin' by(auto) ultimately show ?thesis using es by(simp del: \Move1.simps \Move0.simps) blast qed end sublocale J0_J1_heap_base < red0_Red1': FWdelay_bisimulation_base final_expr0 "mred0 P" final_expr1 "mred1' (compP1 P)" convert_RA "\t. bisim_red0_Red1" bisim_wait01 "\MOVE0 P" "\MOVE1 (compP1 P)" for P . context J0_J1_heap_base begin lemma delay_bisimulation_red0_Red1: assumes wf: "wf_J_prog P" shows "delay_bisimulation_measure (mred0 P t) (mred1' (compP1 P) t) bisim_red0_Red1 (ta_bisim (\t. bisim_red0_Red1)) (\MOVE0 P) (\MOVE1 (compP1 P)) (\es es'. False) (\(((e', xs'), exs'), h') (((e, xs), exs), h). countInitBlock e'< countInitBlock e)" (is "delay_bisimulation_measure _ _ _ _ _ _ ?\1 ?\2") proof(unfold_locales) fix s1 s2 s1' assume "bisim_red0_Red1 s1 s2" "red0_mthr.silent_move P t s1 s1'" moreover obtain ex1 exs1 h1 where s1: "s1 = ((ex1, exs1), h1)" by(cases s1, auto) moreover obtain ex1' exs1' h1' where s1': "s1' = ((ex1', exs1'), h1')" by(cases s1', auto) moreover obtain ex2 exs2 h2 where s2: "s2 = ((ex2, exs2), h2)" by(cases s2, auto) ultimately have bisim: "bisim_list1 (ex1, exs1) (ex2, exs2)" and heap: "h1 = h2" and red: "P,t \0 \ex1/exs1,h1\ -\\ \ex1'/exs1',h1'\" and \: "\Move0 P h1 (ex1, exs1)" by(auto simp add: bisim_red0_Red1_def red0_mthr.silent_move_iff) from Red1_simulates_red0[OF wf red bisim] \ obtain ex2'' exs2'' where bisim': "bisim_list1 (ex1', exs1') (ex2'', exs2'')" and red': "\Red1't (compP1 P) t h1 (ex2, exs2) (ex2'', exs2'')" and [simp]: "h1' = h1" by auto from \Red1't_into_Red1'_\mthr_silent_movet[OF red'] bisim' heap have "\s2'. Red1_mthr.silent_movet False (compP1 P) t s2 s2' \ bisim_red0_Red1 s1' s2'" unfolding s2 s1' by(auto simp add: bisim_red0_Red1_def) thus "bisim_red0_Red1 s1' s2 \ ?\1^++ s1' s1 \ (\s2'. Red1_mthr.silent_movet False (compP1 P) t s2 s2' \ bisim_red0_Red1 s1' s2')" .. next fix s1 s2 s2' assume "bisim_red0_Red1 s1 s2" "Red1_mthr.silent_move False (compP1 P) t s2 s2'" moreover obtain ex1 exs1 h1 where s1: "s1 = ((ex1, exs1), h1)" by(cases s1, auto) moreover obtain ex2 exs2 h2 where s2: "s2 = ((ex2, exs2), h2)" by(cases s2, auto) moreover obtain ex2' exs2' h2' where s2': "s2' = ((ex2', exs2'), h2')" by(cases s2', auto) ultimately have bisim: "bisim_list1 (ex1, exs1) (ex2, exs2)" and heap: "h1 = h2" and red: "False,compP1 P,t \1 \ex2/exs2,h2\ -\\ \ex2'/exs2',h2'\" and \: "\Move1 P h2 (ex2, exs2)" by(fastforce simp add: bisim_red0_Red1_def Red1_mthr.silent_move_iff)+ from red0_simulates_Red1[OF wf red bisim] \ obtain e' es' where bisim': "bisim_list1 (e', es') (ex2', exs2')" and red': "\Red0t P t h2 (ex1, exs1) (e', es') \ countInitBlock (fst ex2') < countInitBlock (fst ex2) \ exs2' = exs2 \ e' = ex1 \ es' = exs1" and [simp]: "h2' = h2" by auto from red' show "bisim_red0_Red1 s1 s2' \ ?\2^++ s2' s2 \ (\s1'. red0_mthr.silent_movet P t s1 s1' \ bisim_red0_Red1 s1' s2')" (is "?refl \ ?step") proof assume "\Red0t P t h2 (ex1, exs1) (e', es')" from \Red0t_into_red0_\mthr_silent_movet[OF this] bisim' heap have ?step unfolding s1 s2' by(auto simp add: bisim_red0_Red1_def) thus ?thesis .. next assume "countInitBlock (fst ex2') < countInitBlock (fst ex2) \ exs2' = exs2 \ e' = ex1 \ es' = exs1" hence ?refl using bisim' heap unfolding s1 s2' s2 by (auto simp add: bisim_red0_Red1_def split_beta) thus ?thesis .. qed next fix s1 s2 ta1 s1' assume "bisim_red0_Red1 s1 s2" and "mred0 P t s1 ta1 s1'" and \: "\ \MOVE0 P s1 ta1 s1'" moreover obtain ex1 exs1 h1 where s1: "s1 = ((ex1, exs1), h1)" by(cases s1, auto) moreover obtain ex1' exs1' h1' where s1': "s1' = ((ex1', exs1'), h1')" by(cases s1', auto) moreover obtain ex2 exs2 h2 where s2: "s2 = ((ex2, exs2), h2)" by(cases s2, auto) ultimately have heap: "h2 = h1" and bisim: "bisim_list1 (ex1, exs1) (ex2, exs2)" and red: "P,t \0 \ex1/exs1,h1\ -ta1\ \ex1'/exs1',h1'\" by(auto simp add: bisim_red0_Red1_def) from \ have "\ \Move0 P h1 (ex1, exs1)" unfolding s1 using red by(auto elim!: red0.cases dest: red_\_taD[where extTA="extTA2J0 P", OF extTA2J0_\]) with Red1_simulates_red0[OF wf red bisim] obtain ex2'' exs2'' ex2' exs2' ta' where bisim': "bisim_list1 (ex1', exs1') (ex2'', exs2'')" and red': "\Red1'r (compP1 P) t h1 (ex2, exs2) (ex2', exs2')" and red'': "False,compP1 P,t \1 \ex2'/exs2',h1\ -ta'\ \ex2''/exs2'',h1'\" and \': "\ \Move1 P h1 (ex2', exs2')" and ta: "ta_bisim01 ta1 ta'" by fastforce from red'' have "mred1' (compP1 P) t ((ex2', exs2'), h1) ta' ((ex2'', exs2''), h1')" by auto moreover from \' have "\ \MOVE1 (compP1 P) ((ex2', exs2'), h1) ta' ((ex2'', exs2''), h1')" by simp moreover from red' have "Red1_mthr.silent_moves False (compP1 P) t s2 ((ex2', exs2'), h1)" unfolding s2 heap by(rule \Red1'r_into_Red1'_\mthr_silent_moves) moreover from bisim' have "bisim_red0_Red1 ((ex1', exs1'), h1') ((ex2'', exs2''), h1')" by(auto simp add: bisim_red0_Red1_def) ultimately show "\s2' s2'' ta2. Red1_mthr.silent_moves False (compP1 P) t s2 s2' \ mred1' (compP1 P) t s2' ta2 s2'' \ \ \MOVE1 (compP1 P) s2' ta2 s2'' \ bisim_red0_Red1 s1' s2'' \ ta_bisim01 ta1 ta2" using ta unfolding s1' by blast next fix s1 s2 ta2 s2' assume "bisim_red0_Red1 s1 s2" and "mred1' (compP1 P) t s2 ta2 s2'" and \: "\ \MOVE1 (compP1 P) s2 ta2 s2'" moreover obtain ex1 exs1 h1 where s1: "s1 = ((ex1, exs1), h1)" by(cases s1, auto) moreover obtain ex2 exs2 h2 where s2: "s2 = ((ex2, exs2), h2)" by(cases s2, auto) moreover obtain ex2' exs2' h2' where s2': "s2' = ((ex2', exs2'), h2')" by(cases s2', auto) ultimately have heap: "h2 = h1" and bisim: "bisim_list1 (ex1, exs1) (ex2, exs2)" and red: "False,compP1 P,t \1 \ex2/exs2,h1\ -ta2\ \ex2'/exs2',h2'\" by(auto simp add: bisim_red0_Red1_def) from \ heap have "\ \Move1 P h2 (ex2, exs2)" unfolding s2 using red by(fastforce elim!: Red1.cases dest: red1_\_taD) with red0_simulates_Red1[OF wf red bisim] obtain e' es' e'' es'' ta' where bisim': "bisim_list1 (e', es') (ex2', exs2')" and red': "\Red0r P t h1 (ex1, exs1) (e'', es'')" and red'': "P,t \0 \e''/es'',h1\ -ta'\ \e'/es',h2'\" and \': "\ \Move0 P h1 (e'', es'')" and ta: "ta_bisim01 ta' ta2" using heap by fastforce from red'' have "mred0 P t ((e'', es''), h1) ta' ((e', es'), h2')" by auto moreover from red' have "red0_mthr.silent_moves P t ((ex1, exs1), h1) ((e'', es''), h1)" by(rule \Red0r_into_red0_\mthr_silent_moves) moreover from \' have "\ \MOVE0 P ((e'', es''), h1) ta' ((e', es'), h2')" by simp moreover from bisim' have "bisim_red0_Red1 ((e', es'), h2') ((ex2', exs2'), h2')" by(auto simp add: bisim_red0_Red1_def) ultimately show "\s1' s1'' ta1. red0_mthr.silent_moves P t s1 s1' \ mred0 P t s1' ta1 s1'' \ \ \MOVE0 P s1' ta1 s1'' \ bisim_red0_Red1 s1'' s2' \ ta_bisim01 ta1 ta2" using ta unfolding s1 s2' by blast next show "wfP ?\1" by auto next have "wf (measure countInitBlock)" .. hence "wf (inv_image (measure countInitBlock) (\(((e', xs'), exs'), h'). e'))" .. also have "inv_image (measure countInitBlock) (\(((e', xs'), exs'), h'). e') = {(s2', s2). ?\2 s2' s2}" by(simp add: inv_image_def split_beta) finally show "wfP ?\2" by(simp add: wfP_def) qed lemma delay_bisimulation_diverge_red0_Red1: assumes "wf_J_prog P" shows "delay_bisimulation_diverge (mred0 P t) (mred1' (compP1 P) t) bisim_red0_Red1 (ta_bisim (\t. bisim_red0_Red1)) (\MOVE0 P) (\MOVE1 (compP1 P))" proof - interpret delay_bisimulation_measure "mred0 P t" "mred1' (compP1 P) t" "bisim_red0_Red1" "ta_bisim (\t. bisim_red0_Red1)" "\MOVE0 P" "\MOVE1 (compP1 P)" "\es es'. False" "\(((e', xs'), exs'), h') (((e, xs), exs), h). countInitBlock e'< countInitBlock e" using assms by(rule delay_bisimulation_red0_Red1) show ?thesis by unfold_locales qed lemma red0_Red1'_FWweak_bisim: assumes wf: "wf_J_prog P" shows "FWdelay_bisimulation_diverge final_expr0 (mred0 P) final_expr1 (mred1' (compP1 P)) (\t. bisim_red0_Red1) bisim_wait01 (\MOVE0 P) (\MOVE1 (compP1 P))" proof - interpret delay_bisimulation_diverge "mred0 P t" "mred1' (compP1 P) t" bisim_red0_Red1 "ta_bisim (\t. bisim_red0_Red1)" "\MOVE0 P" "\MOVE1 (compP1 P)" for t using wf by(rule delay_bisimulation_diverge_red0_Red1) show ?thesis proof fix t and s1 and s2 :: "(('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list) \ 'heap" assume "bisim_red0_Red1 s1 s2" "(\(x1, m). final_expr0 x1) s1" moreover hence "(\(x2, m). final_expr1 x2) s2" by(cases s1)(cases s2,auto simp add: bisim_red0_Red1_def final_iff elim!: bisim_list1_elim elim: bisim_list.cases) ultimately show "\s2'. Red1_mthr.silent_moves False (compP1 P) t s2 s2' \ bisim_red0_Red1 s1 s2' \ (\(x2, m). final_expr1 x2) s2'" by blast next fix t s1 and s2 :: "(('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list) \ 'heap" assume "bisim_red0_Red1 s1 s2" "(\(x2, m). final_expr1 x2) s2" moreover hence "(\(x1, m). final_expr0 x1) s1" by(cases s1)(cases s2,auto simp add: bisim_red0_Red1_def final_iff elim!: bisim_list1_elim elim: bisim_list.cases) ultimately show "\s1'. red0_mthr.silent_moves P t s1 s1' \ bisim_red0_Red1 s1' s2 \ (\(x1, m). final_expr0 x1) s1'" by blast next fix t' x m1 x' m2 t x1 x2 x1' ta1 x1'' m1' x2' ta2 x2'' m2' assume b: "bisim_red0_Red1 (x, m1) (x', m2)" and bo: "bisim_red0_Red1 (x1, m1) (x2, m2)" and \red1: "red0_mthr.silent_moves P t (x1, m1) (x1', m1)" and red1: "mred0 P t (x1', m1) ta1 (x1'', m1')" and \1: "\ \MOVE0 P (x1', m1) ta1 (x1'', m1')" and \red2: "Red1_mthr.silent_moves False (compP1 P) t (x2, m2) (x2', m2)" and red2: "mred1' (compP1 P) t (x2', m2) ta2 (x2'', m2')" and bo': "bisim_red0_Red1 (x1'', m1') (x2'', m2')" and tb: "ta_bisim (\t. bisim_red0_Red1) ta1 ta2" from b have "m1 = m2" by(auto simp add: bisim_red0_Red1_def split_beta) moreover from bo' have "m1' = m2'" by(auto simp add: bisim_red0_Red1_def split_beta) ultimately show "bisim_red0_Red1 (x, m1') (x', m2')" using b by(auto simp add: bisim_red0_Red1_def split_beta) next fix t x1 m1 x2 m2 x1' ta1 x1'' m1' x2' ta2 x2'' m2' w assume "bisim_red0_Red1 (x1, m1) (x2, m2)" and "red0_mthr.silent_moves P t (x1, m1) (x1', m1)" and red0: "mred0 P t (x1', m1) ta1 (x1'', m1')" and "\ \MOVE0 P (x1', m1) ta1 (x1'', m1')" and "Red1_mthr.silent_moves False (compP1 P) t (x2, m2) (x2', m2)" and red1: "mred1' (compP1 P) t (x2', m2) ta2 (x2'', m2')" and "\ \MOVE1 (compP1 P) (x2', m2) ta2 (x2'', m2')" and "bisim_red0_Red1 (x1'', m1') (x2'', m2')" and "ta_bisim01 ta1 ta2" and Suspend: "Suspend w \ set \ta1\\<^bsub>w\<^esub>" "Suspend w \ set \ta2\\<^bsub>w\<^esub>" from red0 red1 Suspend show "bisim_wait01 x1'' x2''" by(cases x2')(cases x2'', auto dest: Red_Suspend_is_call Red1_Suspend_is_call simp add: split_beta bisim_wait01_def is_call_def) next fix t x1 m1 x2 m2 ta1 x1' m1' assume "bisim_red0_Red1 (x1, m1) (x2, m2)" and "bisim_wait01 x1 x2" and "mred0 P t (x1, m1) ta1 (x1', m1')" and wakeup: "Notified \ set \ta1\\<^bsub>w\<^esub> \ WokenUp \ set \ta1\\<^bsub>w\<^esub>" moreover obtain e0 es0 where [simp]: "x1 = (e0, es0)" by(cases x1) moreover obtain e0' es0' where [simp]: "x1' = (e0', es0')" by(cases x1') moreover obtain e1 xs1 exs1 where [simp]: "x2 = ((e1, xs1), exs1)" by(cases x2) auto ultimately have bisim: "bisim_list1 (e0, es0) ((e1, xs1), exs1)" and m1: "m2 = m1" and call: "call e0 \ None" "call1 e1 \ None" and red: "P,t \0 \e0/es0, m1\ -ta1\ \e0'/es0', m1'\" by(auto simp add: bisim_wait01_def bisim_red0_Red1_def) from red wakeup have "\ \Move0 P m1 (e0, es0)" by(auto elim!: red0.cases dest: red_\_taD[where extTA="extTA2J0 P", simplified]) with Red1_simulates_red0[OF wf red bisim] call m1 show "\ta2 x2' m2'. mred1' (compP1 P) t (x2, m2) ta2 (x2', m2') \ bisim_red0_Red1 (x1', m1') (x2', m2') \ ta_bisim01 ta1 ta2" by(auto simp add: bisim_red0_Red1_def) next fix t x1 m1 x2 m2 ta2 x2' m2' assume "bisim_red0_Red1 (x1, m1) (x2, m2)" and "bisim_wait01 x1 x2" and "mred1' (compP1 P) t (x2, m2) ta2 (x2', m2')" and wakeup: "Notified \ set \ta2\\<^bsub>w\<^esub> \ WokenUp \ set \ta2\\<^bsub>w\<^esub>" moreover obtain e0 es0 where [simp]: "x1 = (e0, es0)" by(cases x1) moreover obtain e1 xs1 exs1 where [simp]: "x2 = ((e1, xs1), exs1)" by(cases x2) auto moreover obtain e1' xs1' exs1' where [simp]: "x2' = ((e1', xs1'), exs1')" by(cases x2') auto ultimately have bisim: "bisim_list1 (e0, es0) ((e1, xs1), exs1)" and m1: "m2 = m1" and call: "call e0 \ None" "call1 e1 \ None" and red: "False,compP1 P,t \1 \(e1, xs1)/exs1, m1\ -ta2\ \(e1', xs1')/exs1', m2'\" by(auto simp add: bisim_wait01_def bisim_red0_Red1_def) from red wakeup have "\ \Move1 P m1 ((e1, xs1), exs1)" by(auto elim!: Red1.cases dest: red1_\_taD) with red0_simulates_Red1[OF wf red bisim] m1 call show "\ta1 x1' m1'. mred0 P t (x1, m1) ta1 (x1', m1') \ bisim_red0_Red1 (x1', m1') (x2', m2') \ ta_bisim01 ta1 ta2" by(auto simp add: bisim_red0_Red1_def) next show "(\x. final_expr0 x) \ (\x. final_expr1 x)" by(auto simp add: split_paired_Ex final_iff) qed qed lemma bisim_J0_J1_start: assumes wf: "wf_J_prog P" and start: "wf_start_state P C M vs" shows "red0_Red1'.mbisim (J0_start_state P C M vs) (J1_start_state (compP1 P) C M vs)" proof - from start obtain Ts T pns body D where sees: "P \ C sees M:Ts\T=\(pns,body)\ in D" and conf: "P,start_heap \ vs [:\] Ts" by cases auto from conf have vs: "length vs = length Ts" by(rule list_all2_lengthD) from sees_wf_mdecl[OF wf_prog_wwf_prog[OF wf] sees] have fvbody: "fv body \ {this} \ set pns" and len: "length pns = length Ts" by(auto simp add: wf_mdecl_def) with vs have fv: "fv (blocks pns Ts vs body) \ {this}" by auto have wfCM: "wf_J_mdecl P D (M, Ts, T, pns, body)" using sees_wf_mdecl[OF wf sees] by(auto simp add: wf_mdecl_def) then obtain T' where wtbody: "P,[this # pns [\] Class D # Ts] \ body :: T'" by auto hence elbody: "expr_locks body = (\l. 0)" by(rule WT_expr_locks) hence cisbody: "\ contains_insync body" by(auto simp add: contains_insync_conv) from wfCM len vs have dabody: "\ (blocks pns Ts vs body) \{this}\" by auto from sees have sees1: "compP1 P \ C sees M:Ts\T = \compE1 (this # pns) body\ in D" by(auto dest: sees_method_compP[where f="(\C M Ts T (pns, body). compE1 (this # pns) body)"]) let ?e = "blocks1 0 (Class C#Ts) (compE1 (this # pns) body)" let ?xs = "Null # vs @ replicate (max_vars body) undefined_value" from fvbody cisbody len vs have "bisim [] (blocks (this # pns) (Class D # Ts) (Null # vs) body) (blocks1 (length ([] :: vname list)) (Class D # Ts) (compE1 (this # pns) body)) ?xs" by-(rule blocks_bisim,(fastforce simp add: nth_Cons' nth_append)+) with fv dabody len vs elbody sees sees1 show ?thesis unfolding start_state_def by(auto intro!: red0_Red1'.mbisimI split: if_split_asm)(auto simp add: bisim_red0_Red1_def blocks1_max_vars intro!: bisim_list.intros bisim_list1I wset_thread_okI) qed end end diff --git a/thys/JinjaThreads/Compiler/J1JVMBisim.thy b/thys/JinjaThreads/Compiler/J1JVMBisim.thy --- a/thys/JinjaThreads/Compiler/J1JVMBisim.thy +++ b/thys/JinjaThreads/Compiler/J1JVMBisim.thy @@ -1,5582 +1,5581 @@ (* Title: JinjaThreads/Compiler/J1JVMBisim.thy Author: Andreas Lochbihler *) section \The delay bisimulation between intermediate language and JVM\ theory J1JVMBisim imports Execs "../BV/BVNoTypeError" J1 begin declare Listn.lesub_list_impl_same_size[simp del] lemma (in JVM_heap_conf_base') \exec_1_\exec_1_d: "\ wf_jvm_prog\<^bsub>\\<^esub> P; \exec_1 P t \ \'; \ |- t:\ [ok] \ \ \exec_1_d P t \ \'" by(auto simp add: \exec_1_def \exec_1_d_def welltyped_commute[symmetric] elim: jvmd_NormalE) context JVM_conf_read begin lemma \Exec_1r_preserves_correct_state: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and exec: "\Exec_1r P t \ \'" shows "\ |- t:\ [ok] \ \ |- t:\' [ok]" using exec by(induct)(blast intro: BV_correct_1[OF wf])+ lemma \Exec_1t_preserves_correct_state: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and exec: "\Exec_1t P t \ \'" shows "\ |- t:\ [ok] \ \ |- t:\' [ok]" using exec by(induct)(blast intro: BV_correct_1[OF wf])+ lemma \Exec_1r_\Exec_1_dr: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "\ \Exec_1r P t \ \'; \ |- t:\ [ok] \ \ \Exec_1_dr P t \ \'" apply(induct rule: rtranclp_induct) apply(blast intro: rtranclp.rtrancl_into_rtrancl \exec_1_\exec_1_d[OF wf] \Exec_1r_preserves_correct_state[OF wf])+ done lemma \Exec_1t_\Exec_1_dt: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "\ \Exec_1t P t \ \'; \ |- t:\ [ok] \ \ \Exec_1_dt P t \ \'" apply(induct rule: tranclp_induct) apply(blast intro: tranclp.trancl_into_trancl \exec_1_\exec_1_d[OF wf] \Exec_1t_preserves_correct_state[OF wf])+ done lemma \Exec_1_dr_preserves_correct_state: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and exec: "\Exec_1_dr P t \ \'" shows "\ |- t: \ [ok] \ \ |- t: \' [ok]" using exec by(induct)(blast intro: BV_correct_1[OF wf])+ lemma \Exec_1_dt_preserves_correct_state: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and exec: "\Exec_1_dt P t \ \'" shows "\ |- t:\ [ok] \ \ |- t:\' [ok]" using exec by(induct)(blast intro: BV_correct_1[OF wf])+ end locale J1_JVM_heap_base = J1_heap_base + JVM_heap_base + constrains addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" begin inductive bisim1 :: "'m prog \ 'heap \ 'addr expr1 \ ('addr expr1 \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" and bisims1 :: "'m prog \ 'heap \ 'addr expr1 list \ ('addr expr1 list \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" and bisim1_syntax :: "'m prog \ 'addr expr1 \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" ("_,_,_ \ _ \ _" [50, 0, 0, 0, 50] 100) and bisims1_syntax :: "'m prog \ 'addr expr1 list \ 'heap \ ('addr expr1 list \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" ("_,_,_ \ _ [\] _" [50, 0, 0, 0, 50] 100) for P :: "'m prog" and h :: 'heap where "P, e, h \ exs \ s \ bisim1 P h e exs s" | "P, es, h \ esxs [\] s \ bisims1 P h es esxs s" | bisim1Val2: "pc = length (compE2 e) \ P, e, h \ (Val v, xs) \ (v # [], xs, pc, None)" | bisim1New: "P, new C, h \ (new C, xs) \ ([], xs, 0, None)" | bisim1NewThrow: "P, new C, h \ (THROW OutOfMemory, xs) \ ([], xs, 0, \addr_of_sys_xcpt OutOfMemory\)" | bisim1NewArray: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, newA T\e\, h \ (newA T\e'\, xs) \ (stk, loc, pc, xcp)" | bisim1NewArrayThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, newA T\e\, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1NewArrayFail: "P, newA T\e\, h \ (Throw a, xs) \ ([v], xs, length (compE2 e), \a\)" | bisim1Cast: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, Cast T e, h \ (Cast T e', xs) \ (stk, loc, pc, xcp)" | bisim1CastThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, Cast T e, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1CastFail: "P, Cast T e, h \ (THROW ClassCast, xs) \ ([v], xs, length (compE2 e), \addr_of_sys_xcpt ClassCast\)" | bisim1InstanceOf: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e instanceof T, h \ (e' instanceof T, xs) \ (stk, loc, pc, xcp)" | bisim1InstanceOfThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, e instanceof T, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Val: "P, Val v, h \ (Val v, xs) \ ([], xs, 0, None)" | bisim1Var: "P, Var V, h \ (Var V, xs) \ ([], xs, 0, None)" | bisim1BinOp1: "P, e1, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e1\bop\e2, h \ (e'\bop\e2, xs) \ (stk, loc, pc, xcp)" | bisim1BinOp2: "P, e2, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e1\bop\e2, h \ (Val v1 \bop\ e', xs) \ (stk @ [v1], loc, length (compE2 e1) + pc, xcp)" | bisim1BinOpThrow1: "P, e1, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, e1\bop\e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1BinOpThrow2: "P, e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, e1\bop\e2, h \ (Throw a, xs) \ (stk @ [v1], loc, length (compE2 e1) + pc, \a\)" | bisim1BinOpThrow: "P, e1\bop\e2, h \ (Throw a, xs) \ ([v1, v2], xs, length (compE2 e1) + length (compE2 e2), \a\)" | bisim1LAss1: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, V:=e, h \ (V:=e', xs) \ (stk, loc, pc, xcp)" | bisim1LAss2: "P, V:=e, h \ (unit, xs) \ ([], xs, Suc (length (compE2 e)), None)" | bisim1LAssThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, V:=e, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1AAcc1: "P, a, h \ (a', xs) \ (stk, loc, pc, xcp) \ P, a\i\, h \ (a'\i\, xs) \ (stk, loc, pc, xcp)" | bisim1AAcc2: "P, i, h \ (i', xs) \ (stk, loc, pc, xcp) \ P, a\i\, h \ (Val v\i'\, xs) \ (stk @ [v], loc, length (compE2 a) + pc, xcp)" | bisim1AAccThrow1: "P, a, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\i\, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1AAccThrow2: "P, i, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\i\, h \ (Throw ad, xs) \ (stk @ [v], loc, length (compE2 a) + pc, \ad\)" | bisim1AAccFail: "P, a\i\, h \ (Throw ad, xs) \ ([v, v'], xs, length (compE2 a) + length (compE2 i), \ad\)" | bisim1AAss1: "P, a, h \ (a', xs) \ (stk, loc, pc, xcp) \ P, a\i\ := e, h \ (a'\i\ := e, xs) \ (stk, loc, pc, xcp)" | bisim1AAss2: "P, i, h \ (i', xs) \ (stk, loc, pc, xcp) \ P, a\i\ := e, h \ (Val v\i'\ := e, xs) \ (stk @ [v], loc, length (compE2 a) + pc, xcp)" | bisim1AAss3: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, a\i\ := e, h \ (Val v\Val v'\ := e', xs) \ (stk @ [v', v], loc, length (compE2 a) + length (compE2 i) + pc, xcp)" | bisim1AAssThrow1: "P, a, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\i\ := e, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1AAssThrow2: "P, i, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\i\ := e, h \ (Throw ad, xs) \ (stk @ [v], loc, length (compE2 a) + pc, \ad\)" | bisim1AAssThrow3: "P, e, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\i\ := e, h \ (Throw ad, xs) \ (stk @ [v', v], loc, length (compE2 a) + length (compE2 i) + pc, \ad\)" | bisim1AAssFail: "P, a\i\ := e, h \ (Throw ad, xs) \ ([v', v, v''], xs, length (compE2 a) + length (compE2 i) + length (compE2 e), \ad\)" | bisim1AAss4: "P, a\i\ := e, h \ (unit, xs) \ ([], xs, Suc (length (compE2 a) + length (compE2 i) + length (compE2 e)), None)" | bisim1ALength: "P, a, h \ (a', xs) \ (stk, loc, pc, xcp) \ P, a\length, h \ (a'\length, xs) \ (stk, loc, pc, xcp)" | bisim1ALengthThrow: "P, a, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\length, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1ALengthNull: "P, a\length, h \ (THROW NullPointer, xs) \ ([Null], xs, length (compE2 a), \addr_of_sys_xcpt NullPointer\)" | bisim1FAcc: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e\F{D}, h \ (e'\F{D}, xs) \ (stk, loc, pc, xcp)" | bisim1FAccThrow: "P, e, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e\F{D}, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1FAccNull: "P, e\F{D}, h \ (THROW NullPointer, xs) \ ([Null], xs, length (compE2 e), \addr_of_sys_xcpt NullPointer\)" | bisim1FAss1: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e\F{D} := e2, h \ (e'\F{D} := e2, xs) \ (stk, loc, pc, xcp)" | bisim1FAss2: "P, e2, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e\F{D} := e2, h \ (Val v\F{D} := e', xs) \ (stk @ [v], loc, length (compE2 e) + pc, xcp)" | bisim1FAssThrow1: "P, e, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e\F{D} := e2, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1FAssThrow2: "P, e2, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e\F{D} := e2, h \ (Throw ad, xs) \ (stk @ [v], loc, length (compE2 e) + pc, \ad\)" | bisim1FAssNull: "P, e\F{D} := e2, h \ (THROW NullPointer, xs) \ ([v, Null], xs, length (compE2 e) + length (compE2 e2), \addr_of_sys_xcpt NullPointer\)" | bisim1FAss3: "P, e\F{D} := e2, h \ (unit, xs) \ ([], xs, Suc (length (compE2 e) + length (compE2 e2)), None)" | bisim1CAS1: "P, e1, h \ (e1', xs) \ (stk, loc, pc, xcp) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (e1'\compareAndSwap(D\F, e2, e3), xs) \ (stk, loc, pc, xcp)" | bisim1CAS2: "P, e2, h \ (e2', xs) \ (stk, loc, pc, xcp) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (Val v\compareAndSwap(D\F, e2', e3), xs) \ (stk @ [v], loc, length (compE2 e1) + pc, xcp)" | bisim1CAS3: "P, e3, h \ (e3', xs) \ (stk, loc, pc, xcp) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (Val v\compareAndSwap(D\F, Val v', e3'), xs) \ (stk @ [v', v], loc, length (compE2 e1) + length (compE2 e2) + pc, xcp)" | bisim1CASThrow1: "P, e1, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1CASThrow2: "P, e2, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (Throw ad, xs) \ (stk @ [v], loc, length (compE2 e1) + pc, \ad\)" | bisim1CASThrow3: "P, e3, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (Throw ad, xs) \ (stk @ [v', v], loc, length (compE2 e1) + length (compE2 e2) + pc, \ad\)" | bisim1CASFail: "P, e1\compareAndSwap(D\F, e2, e3), h \ (Throw ad, xs) \ ([v', v, v''], xs, length (compE2 e1) + length (compE2 e2) + length (compE2 e3), \ad\)" | bisim1Call1: "P, obj, h \ (obj', xs) \ (stk, loc, pc, xcp) \ P, obj\M(ps), h \ (obj'\M(ps), xs) \ (stk, loc, pc, xcp)" | bisim1CallParams: "P, ps, h \ (ps', xs) [\] (stk, loc, pc, xcp) \ P, obj\M(ps), h \ (Val v\M(ps'), xs) \ (stk @ [v], loc, length (compE2 obj) + pc, xcp)" | bisim1CallThrowObj: "P, obj, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, obj\M(ps), h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1CallThrowParams: "P, ps, h \ (map Val vs @ Throw a # ps', xs) [\] (stk, loc, pc, \a\) \ P, obj\M(ps), h \ (Throw a, xs) \ (stk @ [v], loc, length (compE2 obj) + pc, \a\)" | bisim1CallThrow: "length ps = length vs \ P, obj\M(ps), h \ (Throw a, xs) \ (vs @ [v], xs, length (compE2 obj) + length (compEs2 ps), \a\)" | bisim1BlockSome1: "P, {V:T=\v\; e}, h \ ({V:T=\v\; e}, xs) \ ([], xs, 0, None)" | bisim1BlockSome2: "P, {V:T=\v\; e}, h \ ({V:T=\v\; e}, xs) \ ([v], xs, Suc 0, None)" | bisim1BlockSome4: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, {V:T=\v\; e}, h \ ({V:T=None; e'}, xs) \ (stk, loc, Suc (Suc pc), xcp)" | bisim1BlockThrowSome: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, {V:T=\v\; e}, h \ (Throw a, xs) \ (stk, loc, Suc (Suc pc), \a\)" | bisim1BlockNone: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, {V:T=None; e}, h \ ({V:T=None; e'}, xs) \ (stk, loc, pc, xcp)" | bisim1BlockThrowNone: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, {V:T=None; e}, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Sync1: "P, e1, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, sync\<^bsub>V\<^esub> (e1) e2, h \ (sync\<^bsub>V\<^esub> (e') e2, xs) \ (stk, loc, pc, xcp)" | bisim1Sync2: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (sync\<^bsub>V\<^esub> (Val v) e2, xs) \ ([v, v], xs, Suc (length (compE2 e1)), None)" | bisim1Sync3: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (sync\<^bsub>V\<^esub> (Val v) e2, xs) \ ([v], xs[V := v], Suc (Suc (length (compE2 e1))), None)" | bisim1Sync4: "P, e2, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, sync\<^bsub>V\<^esub> (e1) e2, h \ (insync\<^bsub>V\<^esub> (a) e', xs) \ (stk, loc, Suc (Suc (Suc (length (compE2 e1) + pc))), xcp)" | bisim1Sync5: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (insync\<^bsub>V\<^esub> (a) Val v, xs) \ ([xs ! V, v], xs, 4 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync6: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Val v, xs) \ ([v], xs, 5 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync7: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (insync\<^bsub>V\<^esub> (a) Throw a', xs) \ ([Addr a'], xs, 6 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync8: "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)" | bisim1Sync9: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Throw a, xs) \ ([Addr a], xs, 8 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync10: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Throw a, xs) \ ([Addr a], xs, 8 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1Sync11: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (THROW NullPointer, xs) \ ([Null], xs, Suc (Suc (length (compE2 e1))), \addr_of_sys_xcpt NullPointer\)" | bisim1Sync12: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Throw a, xs) \ ([v, v'], xs, 4 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1Sync14: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Throw a, xs) \ ([v, Addr a'], xs, 7 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1SyncThrow: "P, e1, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1InSync: \ \This rule only exists such that @{text "P,e,h \ (e, xs) \ ([], xs, 0, None)"} holds for all @{text "e"}\ "P, insync\<^bsub>V\<^esub> (a) e, h \ (insync\<^bsub>V\<^esub> (a) e, xs) \ ([], xs, 0, None)" | bisim1Seq1: "P, e1, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e1;;e2, h \ (e';;e2, xs) \ (stk, loc, pc, xcp)" | bisim1SeqThrow1: "P, e1, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, e1;;e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Seq2: "P, e2, h \ exs \ (stk, loc, pc, xcp) \ P, e1;;e2, h \ exs \ (stk, loc, Suc (length (compE2 e1) + pc), xcp)" | bisim1Cond1: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, if (e) e1 else e2, h \ (if (e') e1 else e2, xs) \ (stk, loc, pc, xcp)" | bisim1CondThen: "P, e1, h \ exs \ (stk, loc, pc, xcp) \ P, if (e) e1 else e2, h \ exs \ (stk, loc, Suc (length (compE2 e) + pc), xcp)" | bisim1CondElse: "P, e2, h \ exs \ (stk, loc, pc, xcp) \ P, if (e) e1 else e2, h \ exs \ (stk, loc, Suc (Suc (length (compE2 e) + length (compE2 e1) + pc)), xcp)" | bisim1CondThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, if (e) e1 else e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1While1: "P, while (c) e, h \ (while (c) e, xs) \ ([], xs, 0, None)" | bisim1While3: "P, c, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, while (c) e, h \ (if (e') (e;; while (c) e) else unit, xs) \ (stk, loc, pc, xcp)" | bisim1While4: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, while (c) e, h \ (e';; while (c) e, xs) \ (stk, loc, Suc (length (compE2 c) + pc), xcp)" | bisim1While6: "P, while (c) e, h \ (while (c) e, xs) \ ([], xs, Suc (Suc (length (compE2 c) + length (compE2 e))), None)" | bisim1While7: "P, while (c) e, h \ (unit, xs) \ ([], xs, Suc (Suc (Suc (length (compE2 c) + length (compE2 e)))), None)" | bisim1WhileThrow1: "P, c, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, while (c) e, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1WhileThrow2: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, while (c) e, h \ (Throw a, xs) \ (stk, loc, Suc (length (compE2 c) + pc), \a\)" | bisim1Throw1: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, throw e, h \ (throw e', xs) \ (stk, loc, pc, xcp)" | bisim1Throw2: "P, throw e, h \ (Throw a, xs) \ ([Addr a], xs, length (compE2 e), \a\)" | bisim1ThrowNull: "P, throw e, h \ (THROW NullPointer, xs) \ ([Null], xs, length (compE2 e), \addr_of_sys_xcpt NullPointer\)" | bisim1ThrowThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, throw e, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Try: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, try e catch(C V) e2, h \ (try e' catch(C V) e2, xs) \ (stk, loc, pc, xcp)" | bisim1TryCatch1: "\ P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\); typeof_addr h a = \Class_type C'\; P \ C' \\<^sup>* C \ \ P, try e catch(C V) e2, h \ ({V:Class C=None; e2}, xs[V := Addr a]) \ ([Addr a], loc, Suc (length (compE2 e)), None)" | bisim1TryCatch2: "P, e2, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, try e catch(C V) e2, h \ ({V:Class C=None; e'}, xs) \ (stk, loc, Suc (Suc (length (compE2 e) + pc)), xcp)" | bisim1TryFail: "\ P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\); typeof_addr h a = \Class_type C'\; \ P \ C' \\<^sup>* C \ \ P, try e catch(C V) e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1TryCatchThrow: "P, e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, try e catch(C V) e2, h \ (Throw a, xs) \ (stk, loc, Suc (Suc (length (compE2 e) + pc)), \a\)" | bisims1Nil: "P, [], h \ ([], xs) [\] ([], xs, 0, None)" | bisims1List1: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e#es, h \ (e'#es, xs) [\] (stk, loc, pc, xcp)" | bisims1List2: "P, es, h \ (es', xs) [\] (stk, loc, pc, xcp) \ P, e#es, h \ (Val v # es', xs) [\] (stk @ [v], loc, length (compE2 e) + pc, xcp)" inductive_cases bisim1_cases: "P,e,h \ (Val v, xs) \ (stk, loc, pc, xcp)" lemma bisim1_refl: "P,e,h \ (e, xs) \ ([], xs, 0, None)" and bisims1_refl: "P,es,h \ (es, xs) [\] ([], xs, 0, None)" apply(induct e and es rule: call.induct calls.induct) apply(auto intro: bisim1_bisims1.intros simp add: nat_fun_sum_eq_conv) apply(rename_tac option a) apply(case_tac option) apply(auto intro: bisim1_bisims1.intros split: if_split_asm) done lemma bisims1_lengthD: "P, es, h \ (es', xs) [\] s \ length es = length es'" apply(induct es arbitrary: es' s) apply(auto elim: bisims1.cases) done text \ Derive an alternative induction rule for @{term bisim1} such that (i) induction hypothesis are generated for all subexpressions and (ii) the number of surrounding blocks is passed through. \ inductive bisim1' :: "'m prog \ 'heap \ 'addr expr1 \ nat \ ('addr expr1 \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" and bisims1' :: "'m prog \ 'heap \ 'addr expr1 list \ nat \ ('addr expr1 list \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" and bisim1'_syntax :: "'m prog \ 'addr expr1 \ nat \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" ("_,_,_,_ \'' _ \ _" [50, 0, 0, 0, 0, 50] 100) and bisims1'_syntax :: "'m prog \ 'addr expr1 list \ nat \ 'heap \ ('addr expr1 list \ 'addr val list) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" ("_,_,_,_ \'' _ [\] _" [50, 0, 0, 0, 0, 50] 100) for P :: "'m prog" and h :: 'heap where "P, e, n, h \' exs \ s \ bisim1' P h e n exs s" | "P, es, n, h \' esxs [\] s \ bisims1' P h es n esxs s" | bisim1Val2': "P, e, n, h \' (Val v, xs) \ (v # [], xs, length (compE2 e), None)" | bisim1New': "P, new C, n, h \' (new C, xs) \ ([], xs, 0, None)" | bisim1NewThrow': "P, new C, n, h \' (THROW OutOfMemory, xs) \ ([], xs, 0, \addr_of_sys_xcpt OutOfMemory\)" | bisim1NewArray': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, newA T\e\, n, h \' (newA T\e'\, xs) \ (stk, loc, pc, xcp)" | bisim1NewArrayThrow': "P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, newA T\e\, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1NewArrayFail': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, newA T\e\, n, h \' (Throw a, xs) \ ([v], xs, length (compE2 e), \a\)" | bisim1Cast': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, Cast T e, n, h \' (Cast T e', xs) \ (stk, loc, pc, xcp)" | bisim1CastThrow': "P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, Cast T e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1CastFail': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, Cast T e, n, h \' (THROW ClassCast, xs) \ ([v], xs, length (compE2 e), \addr_of_sys_xcpt ClassCast\)" | bisim1InstanceOf': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, e instanceof T, n, h \' (e' instanceof T, xs) \ (stk, loc, pc, xcp)" | bisim1InstanceOfThrow': "P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, e instanceof T, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Val': "P, Val v, n, h \' (Val v, xs) \ ([], xs, 0, None)" | bisim1Var': "P, Var V, n, h \' (Var V, xs) \ ([], xs, 0, None)" | bisim1BinOp1': "\ P, e1, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1\bop\e2, n, h \' (e'\bop\e2, xs) \ (stk, loc, pc, xcp)" | bisim1BinOp2': "\ P, e2, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None) \ \ P, e1\bop\e2, n, h \' (Val v1 \bop\ e', xs) \ (stk @ [v1], loc, length (compE2 e1) + pc, xcp)" | bisim1BinOpThrow1': "\ P, e1, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1\bop\e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1BinOpThrow2': "\ P, e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None) \ \ P, e1\bop\e2, n, h \' (Throw a, xs) \ (stk @ [v1], loc, length (compE2 e1) + pc, \a\)" | bisim1BinOpThrow': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1\bop\e2, n, h \' (Throw a, xs) \ ([v1, v2], xs, length (compE2 e1) + length (compE2 e2), \a\)" | bisim1LAss1': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, V:=e, n, h \' (V:=e', xs) \ (stk, loc, pc, xcp)" | bisim1LAss2': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, V:=e, n, h \' (unit, xs) \ ([], xs, Suc (length (compE2 e)), None)" | bisim1LAssThrow': "P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, V:=e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1AAcc1': "\ P, a, n, h \' (a', xs) \ (stk, loc, pc, xcp); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None) \ \ P, a\i\, n, h \' (a'\i\, xs) \ (stk, loc, pc, xcp)" | bisim1AAcc2': "\ P, i, n, h \' (i', xs) \ (stk, loc, pc, xcp); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None) \ \ P, a\i\, n, h \' (Val v\i'\, xs) \ (stk @ [v], loc, length (compE2 a) + pc, xcp)" | bisim1AAccThrow1': "\ P, a, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None) \ \ P, a\i\, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1AAccThrow2': "\ P, i, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None) \ \ P, a\i\, n, h \' (Throw ad, xs) \ (stk @ [v], loc, length (compE2 a) + pc, \ad\)" | bisim1AAccFail': "\ \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None) \ \ P, a\i\, n, h \' (Throw ad, xs) \ ([v, v'], xs, length (compE2 a) + length (compE2 i), \ad\)" | bisim1AAss1': "\ P, a, n, h \' (a', xs) \ (stk, loc, pc, xcp); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (a'\i\ := e, xs) \ (stk, loc, pc, xcp)" | bisim1AAss2': "\ P, i, n, h \' (i', xs) \ (stk, loc, pc, xcp); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Val v\i'\ := e, xs) \ (stk @ [v], loc, length (compE2 a) + pc, xcp)" | bisim1AAss3': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Val v\Val v'\ := e', xs) \ (stk @ [v', v], loc, length (compE2 a) + length (compE2 i) + pc, xcp)" | bisim1AAssThrow1': "\ P, a, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1AAssThrow2': "\ P, i, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Throw ad, xs) \ (stk @ [v], loc, length (compE2 a) + pc, \ad\)" | bisim1AAssThrow3': "\ P, e, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Throw ad, xs) \ (stk @ [v', v], loc, length (compE2 a) + length (compE2 i) + pc, \ad\)" | bisim1AAssFail': "\ \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Throw ad, xs) \ ([v', v, v''], xs, length (compE2 a) + length (compE2 i) + length (compE2 e), \ad\)" | bisim1AAss4': "\ \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (unit, xs) \ ([], xs, Suc (length (compE2 a) + length (compE2 i) + length (compE2 e)), None)" | bisim1ALength': "P, a, n, h \' (a', xs) \ (stk, loc, pc, xcp) \ P, a\length, n, h \' (a'\length, xs) \ (stk, loc, pc, xcp)" | bisim1ALengthThrow': "P, a, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\length, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1ALengthNull': "(\xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None)) \ P, a\length, n, h \' (THROW NullPointer, xs) \ ([Null], xs, length (compE2 a), \addr_of_sys_xcpt NullPointer\)" | bisim1FAcc': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, e\F{D}, n, h \' (e'\F{D}, xs) \ (stk, loc, pc, xcp)" | bisim1FAccThrow': "P, e, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e\F{D}, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1FAccNull': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, e\F{D}, n, h \' (THROW NullPointer, xs) \ ([Null], xs, length (compE2 e), \addr_of_sys_xcpt NullPointer\)" | bisim1FAss1': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (e'\F{D} := e2, xs) \ (stk, loc, pc, xcp)" | bisim1FAss2': "\ P, e2, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (Val v\F{D} := e', xs) \ (stk @ [v], loc, length (compE2 e) + pc, xcp)" | bisim1FAssThrow1': "\ P, e, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1FAssThrow2': "\ P, e2, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (Throw ad, xs) \ (stk @ [v], loc, length (compE2 e) + pc, \ad\)" | bisim1FAssNull': "\ \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (THROW NullPointer, xs) \ ([v, Null], xs, length (compE2 e) + length (compE2 e2), \addr_of_sys_xcpt NullPointer\)" | bisim1FAss3': "\ \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (unit, xs) \ ([], xs, Suc (length (compE2 e) + length (compE2 e2)), None)" | bisim1CAS1': "\ P, e1, n, h \' (e1', xs) \ (stk, loc, pc, xcp); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None); \xs. P, e3, n, h \' (e3, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (e1'\compareAndSwap(D\F, e2, e3), xs) \ (stk, loc, pc, xcp)" | bisim1CAS2': "\ P, e2, n, h \' (e2', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e3, n, h \' (e3, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Val v\compareAndSwap(D\F, e2', e3), xs) \ (stk @ [v], loc, length (compE2 e1) + pc, xcp)" | bisim1CAS3': "\ P, e3, n, h \' (e3', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Val v\compareAndSwap(D\F, Val v', e3'), xs) \ (stk @ [v', v], loc, length (compE2 e1) + length (compE2 e2) + pc, xcp)" | bisim1CASThrow1': "\ P, e1, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None); \xs. P, e3, n, h \' (e3, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1CASThrow2': "\ P, e2, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e3, n, h \' (e3, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Throw ad, xs) \ (stk @ [v], loc, length (compE2 e1) + pc, \ad\)" | bisim1CASThrow3': "\ P, e3, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Throw ad, xs) \ (stk @ [v', v], loc, length (compE2 e1) + length (compE2 e2) + pc, \ad\)" | bisim1CASFail': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None); \xs. P, e3, n, h \' (e3, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Throw ad, xs) \ ([v', v, v''], xs, length (compE2 e1) + length (compE2 e2) + length (compE2 e3), \ad\)" | bisim1Call1': "\ P, obj, n, h \' (obj', xs) \ (stk, loc, pc, xcp); \xs. P, ps, n, h \' (ps, xs) [\] ([], xs, 0, None) \ \ P, obj\M(ps), n, h \' (obj'\M(ps), xs) \ (stk, loc, pc, xcp)" | bisim1CallParams': "\ P, ps, n, h \' (ps', xs) [\] (stk, loc, pc, xcp); ps \ []; \xs. P, obj, n, h \' (obj, xs) \ ([], xs, 0, None) \ \ P, obj\M(ps), n, h \' (Val v\M(ps'), xs) \ (stk @ [v], loc, length (compE2 obj) + pc, xcp)" | bisim1CallThrowObj': "\ P, obj, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, ps, n, h \' (ps, xs) [\] ([], xs, 0, None)\ \ P, obj\M(ps), n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1CallThrowParams': "\ P, ps, n, h \' (map Val vs @ Throw a # ps', xs) [\] (stk, loc, pc, \a\); \xs. P, obj, n, h \' (obj, xs) \ ([], xs, 0, None) \ \ P, obj\M(ps), n, h \' (Throw a, xs) \ (stk @ [v], loc, length (compE2 obj) + pc, \a\)" | bisim1CallThrow': "\ length ps = length vs; \xs. P, obj, n, h \' (obj, xs) \ ([], xs, 0, None); \xs. P, ps, n, h \' (ps, xs) [\] ([], xs, 0, None) \ \ P, obj\M(ps), n, h \' (Throw a, xs) \ (vs @ [v], xs, length (compE2 obj) + length (compEs2 ps), \a\)" | bisim1BlockSome1': "(\xs. P, e, Suc n, h \' (e, xs) \ ([], xs, 0, None)) \ P, {V:T=\v\; e}, n, h \' ({V:T=\v\; e}, xs) \ ([], xs, 0, None)" | bisim1BlockSome2': "(\xs. P, e, Suc n, h \' (e, xs) \ ([], xs, 0, None)) \ P, {V:T=\v\; e}, n, h \' ({V:T=\v\; e}, xs) \ ([v], xs, Suc 0, None)" | bisim1BlockSome4': "P, e, Suc n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, {V:T=\v\; e}, n, h \' ({V:T=None; e'}, xs) \ (stk, loc, Suc (Suc pc), xcp)" | bisim1BlockThrowSome': "P, e, Suc n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, {V:T=\v\; e}, n, h \' (Throw a, xs) \ (stk, loc, Suc (Suc pc), \a\)" | bisim1BlockNone': "P, e, Suc n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, {V:T=None; e}, n, h \' ({V:T=None; e'}, xs) \ (stk, loc, pc, xcp)" | bisim1BlockThrowNone': "P, e, Suc n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, {V:T=None; e}, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Sync1': "\ P, e1, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (sync\<^bsub>V\<^esub> (e') e2, xs) \ (stk, loc, pc, xcp)" | bisim1Sync2': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (sync\<^bsub>V\<^esub> (Val v) e2, xs) \ ([v, v], xs, Suc (length (compE2 e1)), None)" | bisim1Sync3': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (sync\<^bsub>V\<^esub> (Val v) e2, xs) \ ([v], xs[V := v], Suc (Suc (length (compE2 e1))), None)" | bisim1Sync4': "\ P, e2, Suc n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (insync\<^bsub>V\<^esub> (a) e', xs) \ (stk, loc, Suc (Suc (Suc (length (compE2 e1) + pc))), xcp)" | bisim1Sync5': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (insync\<^bsub>V\<^esub> (a) Val v, xs) \ ([xs ! V, v], xs, 4 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync6': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Val v, xs) \ ([v], xs, 5 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync7': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (insync\<^bsub>V\<^esub> (a) Throw a', xs) \ ([Addr a'], xs, 6 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync8': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (insync\<^bsub>V\<^esub> (a) Throw a', xs) \ ([xs ! V, Addr a'], xs, 7 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync9': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Throw a, xs) \ ([Addr a], xs, 8 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync10': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Throw a, xs) \ ([Addr a], xs, 8 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1Sync11': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (THROW NullPointer, xs) \ ([Null], xs, Suc (Suc (length (compE2 e1))), \addr_of_sys_xcpt NullPointer\)" | bisim1Sync12': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Throw a, xs) \ ([v, v'], xs, 4 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1Sync14': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Throw a, xs) \ ([v, Addr a'], xs, 7 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1SyncThrow': "\ P, e1, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1InSync': "P, insync\<^bsub>V\<^esub> (a) e, n, h \' (insync\<^bsub>V\<^esub> (a) e, xs) \ ([], xs, 0, None)" | bisim1Seq1': "\ P, e1, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1;;e2, n, h \' (e';;e2, xs) \ (stk, loc, pc, xcp)" | bisim1SeqThrow1': "\ P, e1, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1;;e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Seq2': "\ P, e2, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None) \ \ P, e1;;e2, n, h \' (e', xs) \ (stk, loc, Suc (length (compE2 e1) + pc), xcp)" | bisim1Cond1': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, if (e) e1 else e2, n, h \' (if (e') e1 else e2, xs) \ (stk, loc, pc, xcp)" | bisim1CondThen': "\ P, e1, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, if (e) e1 else e2, n, h \' (e', xs) \ (stk, loc, Suc (length (compE2 e) + pc), xcp)" | bisim1CondElse': "\ P, e2, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None) \ \ P, if (e) e1 else e2, n, h \' (e', xs) \ (stk, loc, Suc (Suc (length (compE2 e) + length (compE2 e1) + pc)), xcp)" | bisim1CondThrow': "\ P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, if (e) e1 else e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1While1': "\ \xs. P, c, n, h \' (c, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (while (c) e, xs) \ ([], xs, 0, None)" | bisim1While3': "\ P, c, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (if (e') (e;; while (c) e) else unit, xs) \ (stk, loc, pc, xcp)" | bisim1While4': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, c, n, h \' (c, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (e';; while (c) e, xs) \ (stk, loc, Suc (length (compE2 c) + pc), xcp)" | bisim1While6': "\ \xs. P, c, n, h \' (c, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (while (c) e, xs) \ ([], xs, Suc (Suc (length (compE2 c) + length (compE2 e))), None)" | bisim1While7': "\ \xs. P, c, n, h \' (c, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (unit, xs) \ ([], xs, Suc (Suc (Suc (length (compE2 c) + length (compE2 e)))), None)" | bisim1WhileThrow1': "\ P, c, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1WhileThrow2': "\ P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, c, n, h \' (c, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (Throw a, xs) \ (stk, loc, Suc (length (compE2 c) + pc), \a\)" | bisim1Throw1': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, throw e, n, h \' (throw e', xs) \ (stk, loc, pc, xcp)" | bisim1Throw2': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, throw e, n, h \' (Throw a, xs) \ ([Addr a], xs, length (compE2 e), \a\)" | bisim1ThrowNull': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, throw e, n, h \' (THROW NullPointer, xs) \ ([Null], xs, length (compE2 e), \addr_of_sys_xcpt NullPointer\)" | bisim1ThrowThrow': "P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, throw e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Try': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, try e catch(C V) e2, n, h \' (try e' catch(C V) e2, xs) \ (stk, loc, pc, xcp)" | bisim1TryCatch1': "\ P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); typeof_addr h a = \Class_type C'\; P \ C' \\<^sup>* C; \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, try e catch(C V) e2, n, h \' ({V:Class C=None; e2}, xs[V := Addr a]) \ ([Addr a], loc, Suc (length (compE2 e)), None)" | bisim1TryCatch2': "\ P, e2, Suc n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, try e catch(C V) e2, n, h \' ({V:Class C=None; e'}, xs) \ (stk, loc, Suc (Suc (length (compE2 e) + pc)), xcp)" | bisim1TryFail': "\ P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); typeof_addr h a = \Class_type C'\; \ P \ C' \\<^sup>* C; \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, try e catch(C V) e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1TryCatchThrow': "\ P, e2, Suc n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, try e catch(C V) e2, n, h \' (Throw a, xs) \ (stk, loc, Suc (Suc (length (compE2 e) + pc)), \a\)" | bisims1Nil': "P, [], n, h \' ([], xs) [\] ([], xs, 0, None)" | bisims1List1': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, es, n, h \' (es, xs) [\] ([], xs, 0, None) \ \ P, e#es, n, h \' (e'#es, xs) [\] (stk, loc, pc, xcp)" | bisims1List2': "\ P, es, n, h \' (es', xs) [\] (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, e#es, n, h \' (Val v # es', xs) [\] (stk @ [v], loc, length (compE2 e) + pc, xcp)" lemma bisim1'_refl: "P,e,n,h \' (e,xs) \ ([],xs,0,None)" and bisims1'_refl: "P,es,n,h \' (es,xs) [\] ([],xs,0,None)" apply(induct e and es arbitrary: n xs and n xs rule: call.induct calls.induct) apply(auto intro: bisim1'_bisims1'.intros simp add: nat_fun_sum_eq_conv) apply(rename_tac option a b c) apply(case_tac option) apply(auto intro: bisim1'_bisims1'.intros simp add: fun_eq_iff split: if_split_asm) done lemma bisim1_imp_bisim1': "P, e, h \ exs \ s \ P, e, n, h \' exs \ s" and bisims1_imp_bisims1': "P, es, h \ esxs [\] s \ P, es, n, h \' esxs [\] s" proof(induct arbitrary: n and n rule: bisim1_bisims1.inducts) case (bisim1CallParams ps ps' xs stk loc pc xcp obj M v) show ?case proof(cases "ps = []") case True with \P,ps,h \ (ps', xs) [\] (stk, loc, pc, xcp)\ have "ps' = []" "pc = 0" "stk = []" "loc = xs" "xcp = None" by(auto elim: bisims1.cases) moreover have "P,obj,n,h \' (Val v,xs) \ ([v],xs,length (compE2 obj),None)" by(blast intro: bisim1Val2' bisim1'_refl) hence "P,obj\M([]),n,h \' (Val v\M([]),xs) \ ([v],xs,length (compE2 obj),None)" by-(rule bisim1Call1', auto intro!: bisims1Nil' simp add: bsoks_def) ultimately show ?thesis using True by simp next case False with bisim1CallParams show ?thesis by(auto intro: bisim1CallParams' bisims1'_refl bisim1'_refl) qed qed(auto intro: bisim1'_bisims1'.intros bisim1'_refl bisims1'_refl) lemma bisim1'_imp_bisim1: "P, e, n, h \' exs \ s \ P, e, h \ exs \ s" and bisims1'_imp_bisims1: "P, es, n, h \' esxs [\] s \ P, es, h \ esxs [\] s" apply(induct rule: bisim1'_bisims1'.inducts) apply(blast intro: bisim1_bisims1.intros)+ done lemma bisim1'_eq_bisim1: "bisim1' P h e n = bisim1 P h e" and bisims1'_eq_bisims1: "bisims1' P h es n = bisims1 P h es" by(blast intro!: ext bisim1_imp_bisim1' bisims1_imp_bisims1' bisim1'_imp_bisim1 bisims1'_imp_bisims1)+ end (* FIXME: Take lemmas out of locale to speed up opening the context *) lemmas bisim1_bisims1_inducts = J1_JVM_heap_base.bisim1'_bisims1'.inducts [simplified J1_JVM_heap_base.bisim1'_eq_bisim1 J1_JVM_heap_base.bisims1'_eq_bisims1, consumes 1, case_names bisim1Val2 bisim1New bisim1NewThrow bisim1NewArray bisim1NewArrayThrow bisim1NewArrayFail bisim1Cast bisim1CastThrow bisim1CastFail bisim1InstanceOf bisim1InstanceOfThrow bisim1Val bisim1Var bisim1BinOp1 bisim1BinOp2 bisim1BinOpThrow1 bisim1BinOpThrow2 bisim1BinOpThrow bisim1LAss1 bisim1LAss2 bisim1LAssThrow bisim1AAcc1 bisim1AAcc2 bisim1AAccThrow1 bisim1AAccThrow2 bisim1AAccFail bisim1AAss1 bisim1AAss2 bisim1AAss3 bisim1AAssThrow1 bisim1AAssThrow2 bisim1AAssThrow3 bisim1AAssFail bisim1AAss4 bisim1ALength bisim1ALengthThrow bisim1ALengthNull bisim1FAcc bisim1FAccThrow bisim1FAccNull bisim1FAss1 bisim1FAss2 bisim1FAssThrow1 bisim1FAssThrow2 bisim1FAssNull bisim1FAss3 bisim1CAS1 bisim1CAS2 bisim1CAS3 bisim1CASThrow1 bisim1CASThrow2 bisim1CASThrow3 bisim1CASFail bisim1Call1 bisim1CallParams bisim1CallThrowObj bisim1CallThrowParams bisim1CallThrow bisim1BlockSome1 bisim1BlockSome2 bisim1BlockSome4 bisim1BlockThrowSome bisim1BlockNone bisim1BlockThrowNone bisim1Sync1 bisim1Sync2 bisim1Sync3 bisim1Sync4 bisim1Sync5 bisim1Sync6 bisim1Sync7 bisim1Sync8 bisim1Sync9 bisim1Sync10 bisim1Sync11 bisim1Sync12 bisim1Sync14 bisim1SyncThrow bisim1InSync bisim1Seq1 bisim1SeqThrow1 bisim1Seq2 bisim1Cond1 bisim1CondThen bisim1CondElse bisim1CondThrow bisim1While1 bisim1While3 bisim1While4 bisim1While6 bisim1While7 bisim1WhileThrow1 bisim1WhileThrow2 bisim1Throw1 bisim1Throw2 bisim1ThrowNull bisim1ThrowThrow bisim1Try bisim1TryCatch1 bisim1TryCatch2 bisim1TryFail bisim1TryCatchThrow bisims1Nil bisims1List1 bisims1List2] lemmas bisim1_bisims1_inducts_split = bisim1_bisims1_inducts[split_format (complete)] context J1_JVM_heap_base begin lemma bisim1_pc_length_compE2: "P,E,h \ (e, xs) \ (stk, loc, pc, xcp) \ pc \ length (compE2 E)" and bisims1_pc_length_compEs2: "P,Es,h \ (es, xs) [\] (stk, loc, pc, xcp) \ pc \ length (compEs2 Es)" apply(induct "(stk, loc, pc, xcp)" and "(stk, loc, pc, xcp)" arbitrary: stk loc pc xcp and stk loc pc xcp rule: bisim1_bisims1.inducts) apply(auto) done lemma bisim1_pc_length_compE2D: "P,e,h \ (e', xs) \ (stk,loc,length (compE2 e),xcp) \ xcp = None \ call1 e' = None \ (\v. stk = [v] \ (is_val e' \ e' = Val v \ xs = loc))" and bisims1_pc_length_compEs2D: "P,es,h \ (es', xs) [\] (stk,loc,length (compEs2 es),xcp) \ xcp = None \ calls1 es' = None \ (\vs. stk = rev vs \ length vs = length es \ (is_vals es' \ es' = map Val vs \ xs = loc))" proof(induct "(e', xs)" "(stk, loc, length (compE2 e), xcp)" and "(es', xs)" "(stk, loc, length (compEs2 es), xcp)" arbitrary: e' xs stk loc xcp and es' xs stk loc xcp rule: bisim1_bisims1.inducts) case (bisims1List2 es es' xs stk loc pc xcp e v) then obtain vs where "xcp = None" "calls1 es' = None" "stk = rev vs" "length vs = length es" "is_vals es' \ es' = map Val vs \ xs = loc" by auto thus ?case by(clarsimp)(rule_tac x="v#vs" in exI, auto) qed(simp_all (no_asm_use), (fastforce dest: bisim1_pc_length_compE2 bisims1_pc_length_compEs2 split: bop.split_asm if_split_asm)+) corollary bisim1_call_pcD: "\ P,e,h \ (e', xs) \ (stk, loc, pc, xcp); call1 e' = \aMvs\ \ \ pc < length (compE2 e)" and bisims1_calls_pcD: "\ P,es,h \ (es', xs) [\] (stk, loc, pc, xcp); calls1 es' = \aMvs\ \ \ pc < length (compEs2 es)" proof - assume bisim: "P,e,h \ (e', xs) \ (stk, loc, pc, xcp)" and call: "call1 e' = \aMvs\" { assume "pc = length (compE2 e)" with bisim call have False by(auto dest: bisim1_pc_length_compE2D) } moreover from bisim have "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) ultimately show "pc < length (compE2 e)" by(cases "pc < length (compE2 e)")(auto) next assume bisim: "P,es,h \ (es', xs) [\] (stk, loc, pc, xcp)" and call: "calls1 es' = \aMvs\" { assume "pc = length (compEs2 es)" with bisim call have False by(auto dest: bisims1_pc_length_compEs2D) } moreover from bisim have "pc \ length (compEs2 es)" by(rule bisims1_pc_length_compEs2) ultimately show "pc < length (compEs2 es)" by(cases "pc < length (compEs2 es)")(auto) qed lemma bisim1_length_xs: "P,e,h \ (e',xs) \ (stk, loc, pc, xcp) \ length xs = length loc" and bisims1_length_xs: "P,es,h \ (es',xs) [\] (stk, loc, pc, xcp) \ length xs = length loc" by(induct "(e',xs)" "(stk, loc, pc, xcp)" and "(es',xs)" "(stk, loc, pc, xcp)" arbitrary: e' xs stk loc pc xcp and es' xs stk loc pc xcp rule: bisim1_bisims1.inducts) auto lemma bisim1_Val_length_compE2D: "P,e,h \ (Val v,xs) \ (stk, loc, length (compE2 e), xcp) \ stk = [v] \ xs = loc \ xcp = None" and bisims1_Val_length_compEs2D: "P,es,h \ (map Val vs,xs) [\] (stk, loc, length (compEs2 es), xcp) \ stk = rev vs \ xs = loc \ xcp = None" by(auto dest: bisim1_pc_length_compE2D bisims1_pc_length_compEs2D) lemma bisim_Val_loc_eq_xcp_None: "P, e, h \ (Val v, xs) \ (stk, loc, pc, xcp) \ xs = loc \ xcp = None" and bisims_Val_loc_eq_xcp_None: "P, es, h \ (map Val vs, xs) [\] (stk, loc, pc, xcp) \ xs = loc \ xcp = None" apply(induct "(Val v :: 'addr expr1, xs)" "(stk, loc, pc, xcp)" and "(map Val vs :: 'addr expr1 list, xs)" "(stk, loc, pc, xcp)" arbitrary: v xs stk loc pc xcp and vs xs stk loc pc xcp rule: bisim1_bisims1.inducts) apply(auto) done lemma bisim_Val_pc_not_Invoke: "\ P,e,h \ (Val v,xs) \ (stk,loc,pc,xcp); pc < length (compE2 e) \ \ compE2 e ! pc \ Invoke M n'" and bisims_Val_pc_not_Invoke: "\ P,es,h \ (map Val vs,xs) [\] (stk,loc,pc,xcp); pc < length (compEs2 es) \ \ compEs2 es ! pc \ Invoke M n'" apply(induct "(Val v :: 'addr expr1, xs)" "(stk, loc, pc, xcp)" and "(map Val vs :: 'addr expr1 list, xs)" "(stk, loc, pc, xcp)" arbitrary: v xs stk loc pc xcp and vs xs stk loc pc xcp rule: bisim1_bisims1.inducts) apply(auto simp add: nth_append compEs2_map_Val dest: bisim1_pc_length_compE2) done lemma bisim1_VarD: "P, E, h \ (Var V,xs) \ (stk,loc,pc,xcp) \ xs = loc" and "P, es, h \ (es', xs) [\] (stk, loc, pc, xcp) \ True" by(induct "(Var V :: 'addr expr1, xs)" "(stk, loc, pc, xcp)" and arbitrary: V xs stk loc pc xcp and rule: bisim1_bisims1.inducts) auto lemma bisim1_ThrowD: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, xcp) \ pc < length (compE2 e) \ (xcp = \a\ \ xcp = None) \ xs = loc" and bisims1_ThrowD: "P, es, h \ (map Val vs @ Throw a # es', xs) [\] (stk, loc, pc, xcp) \ pc < length (compEs2 es) \ (xcp = \a\ \ xcp = None) \ xs = loc" apply(induct "(Throw a :: 'addr expr1, xs)" "(stk, loc, pc, xcp)" and "(map Val vs @ Throw a # es', xs)" "(stk, loc, pc, xcp)" arbitrary: xs stk loc pc xcp and vs es' xs stk loc pc xcp rule: bisim1_bisims1.inducts) apply(fastforce dest: bisim1_pc_length_compE2 bisim_Val_loc_eq_xcp_None simp add: Cons_eq_append_conv)+ done lemma fixes P :: "'addr J1_prog" shows bisim1_Invoke_stkD: "\ P,e,h \ exs \ (stk,loc,pc,None); pc < length (compE2 e); compE2 e ! pc = Invoke M n' \ \ \vs v stk'. stk = vs @ v # stk' \ length vs = n'" and bisims1_Invoke_stkD: "\ P,es,h \ esxs [\] (stk,loc,pc,None); pc < length (compEs2 es); compEs2 es ! pc = Invoke M n' \ \ \vs v stk'. stk = vs @ v # stk' \ length vs = n'" proof(induct "(stk, loc, pc, None :: 'addr option)" and "(stk, loc, pc, None :: 'addr option)" arbitrary: stk loc pc and stk loc pc rule: bisim1_bisims1.inducts) case bisim1Call1 thus ?case apply(clarsimp simp add: nth_append append_eq_append_conv2 neq_Nil_conv split: if_split_asm) apply(drule bisim1_pc_length_compE2, clarsimp simp add: neq_Nil_conv nth_append) apply(frule bisim1_pc_length_compE2, clarsimp) apply(drule bisim1_pc_length_compE2D, fastforce) done next case bisim1CallParams thus ?case apply(clarsimp simp add: nth_append append_eq_Cons_conv split: if_split_asm) apply(fastforce simp add: append_eq_append_conv2 Cons_eq_append_conv) apply(frule bisims1_pc_length_compEs2, clarsimp) apply(drule bisims1_pc_length_compEs2D, fastforce simp add: append_eq_append_conv2) done qed(fastforce simp add: nth_append append_eq_append_conv2 neq_Nil_conv split: if_split_asm bop.split_asm dest: bisim1_pc_length_compE2 bisims1_pc_length_compEs2)+ lemma fixes P :: "'addr J1_prog" shows bisim1_call_xcpNone: "P,e,h \ (e',xs) \ (stk,loc,pc,\a\) \ call1 e' = None" and bisims1_calls_xcpNone: "P,es,h \ (es',xs) [\] (stk,loc,pc,\a\) \ calls1 es' = None" apply(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) apply(auto dest: bisim_Val_loc_eq_xcp_None bisims_Val_loc_eq_xcp_None simp add: is_vals_conv) done lemma bisims1_map_Val_append: assumes bisim: "P, es', h \ (es'', xs) [\] (stk, loc, pc, xcp)" shows "length es = length vs \ P, es @ es', h \ (map Val vs @ es'', xs) [\] (stk @ rev vs, loc, length (compEs2 es) + pc, xcp)" proof(induction vs arbitrary: es) case Nil thus ?case using bisim by simp next case (Cons v vs) from \length es = length (v # vs)\ obtain e es''' where [simp]: "es = e # es'''" by(cases es, auto) with \length es = length (v # vs)\ have len: "length es''' = length vs" by simp from Cons.IH[OF len] show ?case by(simp add: add.assoc append_assoc[symmetric] del: append_assoc)(rule bisims1List2, auto) qed lemma bisim1_hext_mono: "\ P,e,h \ exs \ s; hext h h' \ \ P,e,h' \ exs \ s" (is "PROP ?thesis1") and bisims1_hext_mono: "\ P,es,h \ esxs [\] s; hext h h' \ \ P,es,h' \ esxs [\] s" (is "PROP ?thesis2") proof - assume hext: "hext h h'" have "P,e,h \ exs \ s \ P,e,h' \ exs \ s" and "P,es,h \ esxs [\] s \ P,es,h' \ esxs [\] s" apply(induct rule: bisim1_bisims1.inducts) apply(insert hext) apply(auto intro: bisim1_bisims1.intros dest: hext_objD) done thus "PROP ?thesis1" and "PROP ?thesis2" by auto qed declare match_ex_table_append_not_pcs [simp] match_ex_table_eq_NoneI[simp] outside_pcs_compxE2_not_matches_entry[simp] outside_pcs_compxEs2_not_matches_entry[simp] lemma bisim1_xcp_Some_not_caught: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e pc' d) = None" and bisims1_xcp_Some_not_caught: "P, es, h \ (map Val vs @ Throw a # es', xs) [\] (stk, loc, pc, \a\) \ match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxEs2 es pc' d) = None" proof(induct "(Throw a :: 'addr expr1, xs)" "(stk, loc, pc, \a :: 'addr\)" and "(map Val vs @ Throw a # es' :: 'addr expr1 list, xs)" "(stk, loc, pc, \a :: 'addr\)" arbitrary: xs stk loc pc pc' d and xs stk loc pc vs es' pc' d rule: bisim1_bisims1.inducts) case bisim1Sync10 thus ?case by(simp add: matches_ex_entry_def) next case bisim1Sync11 thus ?case by(simp add: matches_ex_entry_def) next case (bisim1SyncThrow e1 xs stk loc pc e2) note IH = \match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e1 pc' d) = None\ from \P,e1,h \ (Throw a,xs) \ (stk,loc,pc,\a\)\ have "pc < length (compE2 e1)" by(auto dest: bisim1_ThrowD) with IH show ?case by(auto simp add: match_ex_table_append matches_ex_entry_def dest: match_ex_table_pc_length_compE2 intro: match_ex_table_not_pcs_None) next case bisims1List1 thus ?case by(auto simp add: Cons_eq_append_conv dest: bisim1_ThrowD bisim_Val_loc_eq_xcp_None) next case (bisims1List2 es es'' xs stk loc pc e v) hence "\pc' d. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxEs2 es pc' d) = None" by(auto simp add: Cons_eq_append_conv) from this[of "pc' + length (compE2 e)" "Suc d"] show ?case by(auto simp add: add.assoc) next case (bisim1BlockThrowSome e xs stk loc pc T v) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e pc' d) = None" by auto from this[of "2+pc'"] show ?case by(auto) next case (bisim1Seq2 e2 stk loc pc e1 xs) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e2 pc' d) = None" by auto from this[of "Suc (pc' + length (compE2 e1))"] show ?case by(simp add: add.assoc) next case (bisim1CondThen e1 stk loc pc e e2 xs) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e1 pc' d) = None" by auto note this[of "Suc (pc' + length (compE2 e))"] moreover from \P,e1,h \ (Throw a,xs) \ (stk,loc,pc,\a\)\ have "pc < length (compE2 e1)" by(auto dest: bisim1_ThrowD) ultimately show ?case by(simp add: add.assoc match_ex_table_eq_NoneI outside_pcs_compxE2_not_matches_entry) next case (bisim1CondElse e2 stk loc pc e e1 xs) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e2 pc' d) = None" by auto note this[of "Suc (Suc (pc' + (length (compE2 e) + length (compE2 e1))))"] thus ?case by(simp add: add.assoc) next case (bisim1WhileThrow2 e xs stk loc pc c) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e pc' d) = None" by auto from this[of "Suc (pc' + (length (compE2 c)))"] show ?case by(simp add: add.assoc) next case (bisim1Throw1 e xs stk loc pc) thus ?case by(auto dest: bisim_Val_loc_eq_xcp_None) next case (bisim1TryFail e xs stk loc pc C' C e2) hence "match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e pc' d) = None" by auto moreover from \P,e,h \ (Throw a,xs) \ (stk,loc,pc,\a\)\ have "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) ultimately show ?case using \typeof_addr h a = \Class_type C'\\ \\ P \ C' \\<^sup>* C\ by(simp add: matches_ex_entry_def cname_of_def) next case (bisim1TryCatchThrow e2 xs stk loc pc e C) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e2 pc' d) = None" by auto from this[of "Suc (Suc (pc' + (length (compE2 e))))"] show ?case by(simp add: add.assoc matches_ex_entry_def) next case bisim1Sync12 thus ?case by(auto dest: bisim1_ThrowD simp add: match_ex_table_append eval_nat_numeral, simp add: matches_ex_entry_def) next case bisim1Sync14 thus ?case by(auto dest: bisim1_ThrowD simp add: match_ex_table_append eval_nat_numeral, simp add: matches_ex_entry_def) qed(fastforce dest: bisim1_ThrowD simp add: add.assoc[symmetric])+ declare match_ex_table_append_not_pcs [simp del] match_ex_table_eq_NoneI[simp del] outside_pcs_compxE2_not_matches_entry[simp del] outside_pcs_compxEs2_not_matches_entry[simp del] lemma bisim1_xcp_pcD: "P,e,h \ (e', xs) \ (stk, loc, pc, \a\) \ pc < length (compE2 e)" and bisims1_xcp_pcD: "P,es,h \ (es', xs) [\] (stk, loc, pc, \a\) \ pc < length (compEs2 es)" by(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) auto -declare nth_Cons_subtract[simp] declare nth_append [simp] lemma bisim1_Val_\Exec_move: "\ P, E, h \ (Val v, xs) \ (stk, loc, pc, xcp); pc < length (compE2 E) \ \ xs = loc \ xcp = None \ \Exec_mover_a P t E h (stk, xs, pc, None) ([v], xs, length (compE2 E), None)" and bisims1_Val_\Exec_moves: "\ P, Es, h \ (map Val vs, xs) [\] (stk, loc, pc, xcp); pc < length (compEs2 Es) \ \ xs = loc \ xcp = None \ \Exec_movesr_a P t Es h (stk, xs, pc, None) (rev vs, xs, length (compEs2 Es), None)" proof(induct "(Val v :: 'addr expr1, xs)" "(stk, loc, pc, xcp)" and "(map Val vs :: 'addr expr1 list, xs)" "(stk, loc, pc, xcp)" arbitrary: v xs stk loc pc xcp and vs xs stk loc pc xcp rule: bisim1_bisims1.inducts) case bisim1Val thus ?case by(auto intro!: \Execr1step exec_instr \move2Val simp add: exec_move_def) next case (bisim1LAss2 V e xs) have "\Exec_mover_a P t (V:=e) h ([], xs, Suc (length (compE2 e)), None) ([Unit], xs, Suc (Suc (length (compE2 e))), None)" - by(auto intro!: \Execr1step exec_instr \move2LAssRed2 simp add: nth_append exec_move_def) + by(auto intro!: \Execr1step exec_instr \move2LAssRed2 simp add: exec_move_def) with bisim1LAss2 show ?case by simp next case (bisim1AAss4 a i e xs) have "\Exec_mover_a P t (a\i\ := e) h ([], xs, Suc (length (compE2 a) + length (compE2 i) + length (compE2 e)), None) ([Unit], xs, Suc (Suc (length (compE2 a) + length (compE2 i) + length (compE2 e))), None)" - by(auto intro!: \Execr1step exec_instr \move2AAssRed simp add: nth_append exec_move_def) + by(auto intro!: \Execr1step exec_instr \move2AAssRed simp add: exec_move_def) with bisim1AAss4 show ?case by(simp add: ac_simps) next case (bisim1FAss3 e F D e2 xs) have "\Exec_mover_a P t (e\F{D} := e2) h ([], xs, Suc (length (compE2 e) + length (compE2 e2)), None) ([Unit], xs, Suc (Suc (length (compE2 e) + length (compE2 e2))), None)" - by(auto intro!: \Execr1step exec_instr \move2FAssRed simp add: nth_append exec_move_def) + by(auto intro!: \Execr1step exec_instr \move2FAssRed simp add: exec_move_def) with bisim1FAss3 show ?case by simp next case (bisim1Sync6 V e1 e2 v xs) have "\Exec_mover_a P t (sync\<^bsub>V\<^esub> (e1) e2) h ([v], xs, 5 + length (compE2 e1) + length (compE2 e2), None) ([v], xs, 9 + length (compE2 e1) + length (compE2 e2), None)" by(rule \Execr1step)(auto intro: exec_instr \move2Sync6 simp add: exec_move_def) with bisim1Sync6 show ?case by(auto simp add: eval_nat_numeral) next case (bisim1Seq2 e2 stk loc pc xcp e1 v xs) from \Suc (length (compE2 e1) + pc) < length (compE2 (e1;; e2))\ have pc: "pc < length (compE2 e2)" by simp with \pc < length (compE2 e2) \ xs = loc \ xcp = None \ \Exec_mover_a P t e2 h (stk, xs, pc, None) ([v], xs, length (compE2 e2), None)\ have "xs = loc" "xcp = None" "\Exec_mover_a P t e2 h (stk, xs, pc, None) ([v], xs, length (compE2 e2), None)" by auto moreover hence "\Exec_mover_a P t (e1;;e2) h (stk, xs, Suc (length (compE2 e1) + pc), None) ([v], xs, Suc (length (compE2 e1) + length (compE2 e2)), None)" by -(rule Seq_\ExecrI2) ultimately show ?case by(simp) next case (bisim1CondThen e1 stk loc pc xcp e e2 v xs) from \P, e1, h \ (Val v,xs) \ (stk,loc,pc,xcp)\ have "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) have e: "\Exec_mover_a P t (if (e) e1 else e2) h ([v], xs, Suc (length (compE2 e) + (length (compE2 e1))), None) ([v], xs, length (compE2 (if (e) e1 else e2)), None)" - by(rule \Execr1step)(auto simp add: nth_append exec_move_def intro!: exec_instr \move2CondThenExit) + by(rule \Execr1step)(auto simp add: exec_move_def intro!: exec_instr \move2CondThenExit) show ?case proof(cases "pc < length (compE2 e1)") case True with \pc < length (compE2 e1) \ xs = loc \ xcp = None \ \Exec_mover_a P t e1 h (stk, xs, pc, None) ([v], xs, length (compE2 e1), None)\ have s: "xs = loc" "xcp = None" and "\Exec_mover_a P t e1 h (stk, xs, pc, None) ([v], xs, length (compE2 e1), None)" by auto hence "\Exec_mover_a P t (if (e) e1 else e2) h (stk, xs, Suc (length (compE2 e) + pc), None) ([v], xs, Suc (length (compE2 e) + length (compE2 e1)), None)" by -(rule Cond_\ExecrI2) with e True s show ?thesis by(simp) next case False with \pc \ length (compE2 e1)\ have pc: "pc = length (compE2 e1)" by auto with \P, e1, h \ (Val v,xs) \ (stk,loc,pc,xcp)\ have "stk = [v]" "xs = loc" "xcp = None" by(auto dest: bisim1_Val_length_compE2D) with pc e show ?thesis by(simp) qed next case (bisim1CondElse e2 stk loc pc xcp e e1 v xs) from \P, e2, h \ (Val v,xs) \ (stk,loc,pc,xcp)\ have "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e2)") case True with \pc < length (compE2 e2) \ xs = loc \ xcp = None \ \Exec_mover_a P t e2 h (stk, xs, pc, None) ([v], xs, length (compE2 e2), None)\ have s: "xs = loc" "xcp = None" and e: "\Exec_mover_a P t e2 h (stk, xs, pc, None) ([v], xs, length (compE2 e2), None)" by auto from e have "\Exec_mover_a P t (if (e) e1 else e2) h (stk, xs, Suc (Suc (length (compE2 e) + length (compE2 e1) + pc)), None) ([v], xs, Suc (Suc (length (compE2 e) + length (compE2 e1) + length (compE2 e2))), None)" by(rule Cond_\ExecrI3) with True s show ?thesis by(simp add: add.assoc) next case False with \pc \ length (compE2 e2)\ have pc: "pc = length (compE2 e2)" by auto with \P, e2, h \ (Val v,xs) \ (stk,loc,pc,xcp)\ have "stk = [v]" "xs = loc" "xcp = None" by(auto dest: bisim1_Val_length_compE2D) with pc show ?thesis by(simp add: add.assoc) qed next case (bisim1While7 c e xs) have "\Exec_mover_a P t (while (c) e) h ([], xs, Suc (Suc (Suc (length (compE2 c) + length (compE2 e)))), None) ([Unit], xs, length (compE2 (while (c) e)), None)" - by(auto intro!: \Execr1step exec_instr \move2While4 simp add: nth_append exec_move_def) + by(auto intro!: \Execr1step exec_instr \move2While4 simp add: exec_move_def) thus ?case by(simp) next case (bisims1List1 e e' xs stk loc pc xcp es) from \e' # es = map Val vs\ obtain v vs' where [simp]: "vs = v # vs'" "e' = Val v" "es = map Val vs'" by auto from \P,e,h \ (e',xs) \ (stk,loc,pc,xcp)\ have length: "pc \ length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) hence "xs = loc \ xcp = None \ \Exec_mover_a P t e h (stk, xs, pc, None) ([v], xs, length (compE2 e), None)" proof(cases "pc < length (compE2 e)") case True with \\e' = Val v; pc < length (compE2 e)\ \ xs = loc \ xcp = None \ \Exec_mover_a P t e h (stk, xs, pc, None) ([v], xs, length (compE2 e), None)\ show ?thesis by auto next case False with length have pc: "pc = length (compE2 e)" by auto with \P,e,h \ (e',xs) \ (stk,loc,pc,xcp)\ have "stk = [v]" "xs = loc" "xcp = None" by(auto dest: bisim1_Val_length_compE2D) with pc show ?thesis by(auto) qed hence s: "xs = loc" "xcp = None" and exec1: "\Exec_mover_a P t e h (stk, xs, pc, None) ([v], xs, length (compE2 e), None)" by auto from exec1 have "\Exec_movesr_a P t (e # es) h (stk, xs, pc, None) ([v], xs, length (compE2 e), None)" by(auto intro: \Exec_mover_\Exec_movesr) moreover have "\Exec_movesr_a P t (map Val vs') h ([], xs, 0, None) (rev vs', xs, length (compEs2 (map Val vs')), None)" by(rule \Exec_movesr_map_Val) hence "\Exec_movesr_a P t ([e] @ map Val vs') h ([] @ [v], xs, length (compEs2 [e]) + 0, None) (rev vs' @ [v], xs, length (compEs2 [e]) + length (compEs2 (map Val vs')), None)" by -(rule append_\Exec_movesr, auto) ultimately show ?case using s by(auto) next case (bisims1List2 es es' xs stk loc pc xcp e v) from \Val v # es' = map Val vs\ obtain vs' where [simp]: "vs = v # vs'" "es' = map Val vs'" by auto from \P,es,h \ (es',xs) [\] (stk,loc,pc,xcp)\ have length: "pc \ length (compEs2 es)" by(auto dest: bisims1_pc_length_compEs2) hence "xs = loc \ xcp = None \ \Exec_movesr_a P t es h (stk, xs, pc, None) (rev vs', xs, length (compEs2 es), None)" proof(cases "pc < length (compEs2 es)") case True with \\es' = map Val vs'; pc < length (compEs2 es)\ \ xs = loc \ xcp = None \ \Exec_movesr_a P t es h (stk, xs, pc, None) (rev vs', xs, length (compEs2 es), None)\ show ?thesis by auto next case False with length have pc: "pc = length (compEs2 es)" by auto with \P,es,h \ (es',xs) [\] (stk,loc,pc,xcp)\ have "stk = rev vs'" "xs = loc" "xcp = None" by(auto dest: bisims1_Val_length_compEs2D) with pc show ?thesis by(auto) qed hence s: "xs = loc" "xcp = None" and exec1: "\Exec_movesr_a P t es h (stk, xs, pc, None) (rev vs', xs, length (compEs2 es), None)" by auto from exec1 have "\Exec_movesr_a P t ([e] @ es) h (stk @ [v], xs, length (compEs2 [e]) + pc, None) (rev vs' @ [v], xs, length (compEs2 [e]) + length (compEs2 es), None)" by -(rule append_\Exec_movesr, auto) thus ?case using s by(auto) qed(auto) lemma bisim1Val2D1: assumes bisim: "P, e, h \ (Val v,xs) \ (stk,loc,pc,xcp)" shows "xcp = None \ xs = loc \ \Exec_mover_a P t e h (stk, loc, pc, xcp) ([v], loc, length (compE2 e), None)" proof - from bisim have "xcp = None" "xs = loc" by(auto dest: bisim_Val_loc_eq_xcp_None) moreover have "\Exec_mover_a P t e h (stk, loc, pc, xcp) ([v], loc, length (compE2 e), None)" proof(cases "pc < length (compE2 e)") case True from bisim1_Val_\Exec_move[OF bisim True] show ?thesis by auto next case False from bisim have "pc \ length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) with False have "pc = length (compE2 e)" by auto with bisim have "stk = [v]" "loc = xs" "xcp=None" by(auto dest: bisim1_Val_length_compE2D) with \pc = length (compE2 e)\ show ?thesis by(auto) qed ultimately show ?thesis by simp qed lemma bisim1_Throw_\Exec_movet: "\ P, e, h \ (Throw a,xs) \ (stk,loc,pc,None) \ \ \pc'. \Exec_movet_a P t e h (stk, loc, pc, None) ([Addr a], loc, pc', \a\) \ P, e, h \ (Throw a,xs) \ ([Addr a], loc, pc', \a\) \ xs = loc" and bisims1_Throw_\Exec_movest: "\ P, es, h \ (map Val vs @ Throw a # es',xs) [\] (stk,loc,pc,None) \ \ \pc'. \Exec_movest_a P t es h (stk, loc, pc, None) (Addr a # rev vs, loc, pc', \a\) \ P, es, h \ (map Val vs @ Throw a # es',xs) [\] (Addr a # rev vs, loc, pc', \a\) \ xs = loc" proof(induct e "n :: nat" "Throw a :: 'addr expr1" xs stk loc pc "None :: 'addr option" and es "n :: nat" "map Val vs @ Throw a # es' :: 'addr expr1 list" xs stk loc pc "None :: 'addr option" arbitrary: and vs rule: bisim1_bisims1_inducts_split) case (bisim1Sync9 e1 n e2 V xs) let ?pc = "8 + length (compE2 e1) + length (compE2 e2)" have "\Exec_movet_a P t (sync\<^bsub>V\<^esub> (e1) e2) h ([Addr a], xs, ?pc, None) ([Addr a], xs, ?pc, \a\)" by(rule \Exect1step)(auto intro: exec_instr \move2_\moves2.intros simp add: is_Ref_def exec_move_def) moreover have "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (Throw a,xs) \ ([Addr a],xs,?pc,\a\)" by(rule bisim1Sync10) ultimately show ?case by auto next case (bisim1Seq2 e2 n xs stk loc pc e1) then obtain pc' where "\Exec_movet_a P t e2 h (stk, loc, pc, None) ([Addr a], loc, pc', \a\)" "P, e2, h \ (Throw a,xs) \ ([Addr a],loc,pc',\a\)" "xs = loc" by auto thus ?case by(auto intro: Seq_\ExectI2 bisim1_bisims1.bisim1Seq2) next case (bisim1CondThen e1 n xs stk loc pc e e2) then obtain pc' where exec: "\Exec_movet_a P t e1 h (stk, loc, pc, None) ([Addr a], loc, pc', \a\)" and bisim: "P, e1, h \ (Throw a,xs) \ ([Addr a],loc,pc',\a\)" and s: "xs = loc" by auto from exec have "\Exec_movet_a P t (if (e) e1 else e2) h (stk, loc, Suc (length (compE2 e) + pc), None) ([Addr a], loc, Suc (length (compE2 e) + pc'), \a\)" by(rule Cond_\ExectI2) moreover from bisim have "P, if (e) e1 else e2, h \ (Throw a, xs) \ ([Addr a], loc, Suc (length (compE2 e) + pc'), \a\)" by(rule bisim1_bisims1.bisim1CondThen) ultimately show ?case using s by(auto) next case (bisim1CondElse e2 n xs stk loc pc e e1) then obtain pc' where exec: "\Exec_movet_a P t e2 h (stk, loc, pc, None) ([Addr a], loc, pc', \a\)" and bisim: "P, e2, h \ (Throw a, xs ) \ ([Addr a], loc, pc', \a\)" and s: "xs = loc" by auto let "?pc pc" = "Suc (Suc (length (compE2 e) + length (compE2 e1) + pc))" from exec have "\Exec_movet_a P t (if (e) e1 else e2) h (stk, loc, (?pc pc), None) ([Addr a], loc, ?pc pc', \a\)" by(rule Cond_\ExectI3) moreover from bisim have "P, if (e) e1 else e2, h \ (Throw a, xs ) \ ([Addr a], loc, ?pc pc', \a\)" by(rule bisim1_bisims1.bisim1CondElse) ultimately show ?case using s by auto next case (bisim1Throw1 e n xs stk loc pc) note bisim = \P, e, h \ (addr a, xs) \ (stk, loc, pc, None)\ hence s: "xs = loc" and exec: "\Exec_mover_a P t e h (stk, loc, pc, None) ([Addr a], loc, length (compE2 e), None)" by(auto dest: bisim1Val2D1) from exec have "\Exec_mover_a P t (throw e) h (stk, loc, pc, None) ([Addr a], loc, length (compE2 e), None)" by(rule Throw_\ExecrI) also have "\Exec_movet_a P t (throw e) h ([Addr a], loc, length (compE2 e), None) ([Addr a], loc, length (compE2 e), \a\)" by(rule \Exect1step, auto intro: exec_instr \move2Throw2 simp add: is_Ref_def exec_move_def) also have "P, throw e, h \ (Throw a, loc ) \ ([Addr a], loc, length (compE2 e), \a\)" by(rule bisim1Throw2) ultimately show ?case using s by auto next case (bisims1List1 e n e' xs stk loc pc es vs) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ show ?case proof(cases "is_val e'") case True with \e' # es = map Val vs @ Throw a # es'\ obtain v vs' where "vs = v # vs'" "e' = Val v" and es: "es = map Val vs' @ Throw a # es'" by(auto simp add: Cons_eq_append_conv) with bisim have "P,e,h \ (Val v, xs) \ (stk, loc, pc, None)" by simp from bisim1Val2D1[OF this] have [simp]: "xs = loc" and exec: "\Exec_mover_a P t e h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" by auto from exec have "\Exec_movesr_a P t (e # es) h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" by(rule \Exec_mover_\Exec_movesr) also from es \es = map Val vs' @ Throw a # es' \ \pc'. \Exec_movest_a P t es h ([], loc, 0, None) (Addr a # rev vs', loc, pc', \a\) \ P,es,h \ (map Val vs' @ Throw a # es', loc) [\] (Addr a # rev vs', loc, pc', \a\) \ loc = loc\ obtain pc' where execes: "\Exec_movest_a P t es h ([], loc, 0, None) (Addr a # rev vs', loc, pc', \a\)" and bisim': "P,es,h \ (map Val vs' @ Throw a # es', loc) [\] (Addr a # rev vs', loc, pc', \a\)" by auto from append_\Exec_movest[OF _ execes, of "[v]" "[e]"] have "\Exec_movest_a P t (e # es) h ([v], loc, length (compE2 e), None) (Addr a # rev vs' @ [v], loc, length (compE2 e) + pc', \a\)" by simp also from bisims1List2[OF bisim', of e v] es \e' = Val v\ \vs = v # vs'\ have "P,e # es,h \ (e' # es, xs) [\] ((Addr a # rev vs), loc, length (compE2 e) + pc', \a\)" by simp ultimately show ?thesis using \vs = v # vs'\ es \e' = Val v\ by auto next case False with \e' # es = map Val vs @ Throw a # es'\ have [simp]: "e' = Throw a" "es = es'" "vs = []" by(auto simp add: Cons_eq_append_conv) from \e' = Throw a \ \pc'. \Exec_movet_a P t e h (stk, loc, pc, None) ([Addr a], loc, pc', \a\) \ P,e,h \ (Throw a, xs ) \ ([Addr a], loc, pc', \a\) \ xs = loc\ obtain pc' where "\Exec_movet_a P t e h (stk, loc, pc, None) ([Addr a], loc, pc', \a\)" and bisim: "P,e,h \ (Throw a, xs ) \ ([Addr a], loc, pc', \a\)" and s: "xs = loc" by auto hence "\Exec_movest_a P t (e # es) h (stk, loc, pc, None) ([Addr a], loc, pc', \a\)" by-(rule \Exec_movet_\Exec_movest) moreover from bisim have "P,e#es,h \ (Throw a#es,xs) [\] ([Addr a],loc,pc',\a\)" by(rule bisim1_bisims1.bisims1List1) ultimately show ?thesis using s by auto qed next case (bisims1List2 es n es'' xs stk loc pc e v) note IH = \\vs. es'' = map Val vs @ Throw a # es' \ \pc'. \Exec_movest_a P t es h (stk, loc, pc, None) (Addr a # rev vs, loc, pc', \a\) \ P,es,h \ (map Val vs @ Throw a # es',xs) [\] (Addr a # rev vs,loc,pc',\a\) \ xs = loc\ from \Val v # es'' = map Val vs @ Throw a # es'\ obtain vs' where [simp]: "vs = v # vs'" "es'' = map Val vs' @ Throw a # es'" by(auto simp add: Cons_eq_append_conv) from IH[OF \es'' = map Val vs' @ Throw a # es'\] obtain pc' where exec: "\Exec_movest_a P t es h (stk, loc, pc, None) (Addr a # rev vs', loc, pc', \a\)" and bisim: "P,es,h \ (map Val vs' @ Throw a # es',xs) [\] (Addr a # rev vs',loc,pc',\a\)" and [simp]: "xs = loc" by auto from append_\Exec_movest[OF _ exec, of "[v]" "[e]"] have "\Exec_movest_a P t (e # es) h (stk @ [v], loc, length (compE2 e) + pc, None) (Addr a # rev vs, loc, length (compE2 e) + pc', \a\)" by simp moreover from bisim have "P,e#es,h \ (Val v # map Val vs' @ Throw a # es',xs) [\] ((Addr a # rev vs')@[v],loc,length (compE2 e) + pc',\a\)" by(rule bisim1_bisims1.bisims1List2) ultimately show ?case by(auto) qed(auto) lemma bisim1_Throw_\Exec_mover: "\ P, e, h \ (Throw a,xs) \ (stk,loc,pc,None) \ \ \pc'. \Exec_mover_a P t e h (stk, loc, pc, None) ([Addr a], loc, pc', \a\) \ P, e, h \ (Throw a,xs) \ ([Addr a], loc, pc', \a\) \ xs = loc" by(drule bisim1_Throw_\Exec_movet)(blast intro: tranclp_into_rtranclp) lemma bisims1_Throw_\Exec_movesr: "\ P, es, h \ (map Val vs @ Throw a # es',xs) [\] (stk,loc,pc,None) \ \ \pc'. \Exec_movesr_a P t es h (stk, loc, pc, None) (Addr a # rev vs, loc, pc', \a\) \ P, es, h \ (map Val vs @ Throw a # es',xs) [\] (Addr a # rev vs, loc, pc', \a\) \ xs = loc" by(drule bisims1_Throw_\Exec_movest)(blast intro: tranclp_into_rtranclp) declare split_beta [simp] lemma bisim1_inline_call_Throw: "\ P,e,h \ (e', xs) \ (stk, loc, pc, None); call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0; pc < length (compE2 e) \ \ n0 = length vs \ P,e,h \ (inline_call (Throw A) e', xs) \ (stk, loc, pc, \A\)" (is "\ _; _; _; _ \ \ ?concl e n e' xs pc stk loc") and bisims1_inline_calls_Throw: "\ P,es,h \ (es', xs) [\] (stk, loc, pc, None); calls1 es' = \(a, M, vs)\; compEs2 es ! pc = Invoke M n0; pc < length (compEs2 es) \ \ n0 = length vs \ P,es,h \ (inline_calls (Throw A) es', xs) [\] (stk, loc, pc, \A\)" (is "\ _; _; _; _ \ \ ?concls es n es' xs pc stk loc") proof(induct e "n :: nat" e' xs stk loc pc "None :: 'addr option" and es "n :: nat" es' xs stk loc pc "None :: 'addr option" rule: bisim1_bisims1_inducts_split) case (bisim1BinOp1 e1 n e' xs stk loc pc e2 bop) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e1 ! pc = Invoke M n0; pc < length (compE2 e1) \ \ ?concl e1 n e' xs pc stk loc\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note ins = \compE2 (e1 \bop\ e2) ! pc = Invoke M n0\ note call = \call1 (e' \bop\ e2) = \(a, M, vs)\\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e1)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis - by(auto simp add: nth_append intro: bisim1_bisims1.bisim1BinOp1) + by(auto intro: bisim1_bisims1.bisim1BinOp1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e1)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e1)" with bisim1 ins have False - by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 e1)" by(cases "pc < length (compE2 e1)") auto from call ins show ?thesis by simp qed next case bisim1BinOp2 thus ?case by(auto split: if_split_asm bop.split_asm dest: bisim1_bisims1.bisim1BinOp2) next case (bisim1AAcc1 A n a' xs stk loc pc i) note IH1 = \\call1 a' = \(a, M, vs)\; compE2 A ! pc = Invoke M n0; pc < length (compE2 A) \ \ ?concl A n a' xs pc stk loc\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note ins = \compE2 (A\i\) ! pc = Invoke M n0\ note call = \call1 (a'\i\) = \(a, M, vs)\\ show ?case proof(cases "is_val a'") case False with bisim1 call have "pc < length (compE2 A)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis - by(auto simp add: nth_append intro: bisim1_bisims1.bisim1AAcc1) + by(auto intro: bisim1_bisims1.bisim1AAcc1) next case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "pc \ length (compE2 A)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 A)" with bisim1 ins have False - by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 A)" by(cases "pc < length (compE2 A)") auto from call ins show ?thesis by simp qed next case bisim1AAcc2 thus ?case by(auto split: if_split_asm dest: bisim1_bisims1.bisim1AAcc2) next case (bisim1AAss1 A n a' xs stk loc pc i e) note IH1 = \\call1 a' = \(a, M, vs)\; compE2 A ! pc = Invoke M n0; pc < length (compE2 A) \ \ ?concl A n a' xs pc stk loc\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note ins = \compE2 (A\i\ := e) ! pc = Invoke M n0\ note call = \call1 (a'\i\ := e) = \(a, M, vs)\\ show ?case proof(cases "is_val a'") case False with bisim1 call have "pc < length (compE2 A)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis - by(auto simp add: nth_append intro: bisim1_bisims1.bisim1AAss1) + by(auto intro: bisim1_bisims1.bisim1AAss1) next case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "pc \ length (compE2 A)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 A)" with bisim1 ins have False - by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 A)" by(cases "pc < length (compE2 A)") auto from call ins show ?thesis by simp qed next case (bisim1AAss2 i n i' xs stk loc pc A e v) note IH1 = \\call1 i' = \(a, M, vs)\; compE2 i ! pc = Invoke M n0; pc < length (compE2 i) \ \ ?concl i n i' xs pc stk loc\ note bisim1 = \P,i,h \ (i', xs) \ (stk, loc, pc, None)\ note ins = \compE2 (A\i\ := e) ! (length (compE2 A) + pc) = Invoke M n0\ note call = \call1 (Val v\i'\ := e) = \(a, M, vs)\\ show ?case proof(cases "is_val i'") case False with bisim1 call have "pc < length (compE2 i)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis - by(auto simp add: nth_append intro: bisim1_bisims1.bisim1AAss2) + by(auto intro: bisim1_bisims1.bisim1AAss2) next case True then obtain v where [simp]: "i' = Val v" by auto from bisim1 have "pc \ length (compE2 i)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 i)" with bisim1 ins have False - by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 i)" by(cases "pc < length (compE2 i)") auto from call ins show ?thesis by simp qed next case bisim1AAss3 thus ?case by(auto split: if_split_asm nat.split_asm simp add: nth_Cons dest: bisim1_bisims1.bisim1AAss3) next case (bisim1FAss1 e n e' xs stk loc pc e2 F D) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0; pc < length (compE2 e) \ \ ?concl e n e' xs pc stk loc\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note ins = \compE2 (e\F{D} := e2) ! pc = Invoke M n0\ note call = \call1 (e'\F{D} := e2) = \(a, M, vs)\\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis - by(auto simp add: nth_append intro: bisim1_bisims1.bisim1FAss1) + by(auto intro: bisim1_bisims1.bisim1FAss1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e)" with bisim1 ins have False - by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 e)" by(cases "pc < length (compE2 e)") auto from call ins show ?thesis by simp qed next case bisim1FAss2 thus ?case by(auto split: if_split_asm nat.split_asm simp add: nth_Cons dest: bisim1_bisims1.bisim1FAss2) next case (bisim1CAS1 E n e' xs stk loc pc e2 e3 D F) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 E ! pc = Invoke M n0; pc < length (compE2 E) \ \ ?concl E n e' xs pc stk loc\ note bisim1 = \P,E,h \ (e', xs) \ (stk, loc, pc, None)\ note ins = \compE2 _ ! pc = Invoke M n0\ note call = \call1 _ = \(a, M, vs)\\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 E)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis - by(auto simp add: nth_append intro: bisim1_bisims1.bisim1CAS1) + by(auto intro: bisim1_bisims1.bisim1CAS1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 E)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 E)" with bisim1 ins have False - by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 E)" by(cases "pc < length (compE2 E)") auto from call ins show ?thesis by simp qed next case (bisim1CAS2 e2 n e2' xs stk loc pc e1 e3 D F v) note IH1 = \\call1 e2' = \(a, M, vs)\; compE2 e2 ! pc = Invoke M n0; pc < length (compE2 e2) \ \ ?concl e2 n e2' xs pc stk loc\ note bisim1 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, None)\ note ins = \compE2 _ ! (length (compE2 e1) + pc) = Invoke M n0\ note call = \call1 _ = \(a, M, vs)\\ show ?case proof(cases "is_val e2'") case False with bisim1 call have "pc < length (compE2 e2)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis - by(auto simp add: nth_append intro: bisim1_bisims1.bisim1CAS2) + by(auto intro: bisim1_bisims1.bisim1CAS2) next case True then obtain v where [simp]: "e2' = Val v" by auto from bisim1 have "pc \ length (compE2 e2)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e2)" with bisim1 ins have False - by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 e2)" by(cases "pc < length (compE2 e2)") auto from call ins show ?thesis by simp qed next case (bisim1Call1 obj n obj' xs stk loc pc ps M') note IH1 = \\call1 obj' = \(a, M, vs)\; compE2 obj ! pc = Invoke M n0; pc < length (compE2 obj) \ \ ?concl obj n obj' xs pc stk loc\ note IH2 = \\xs. \calls1 ps = \(a, M, vs)\; compEs2 ps ! 0 = Invoke M n0; 0 < length (compEs2 ps) \ \ ?concls ps n ps xs 0 [] xs\ note ins = \compE2 (obj\M'(ps)) ! pc = Invoke M n0\ note bisim1 = \P,obj,h \ (obj', xs) \ (stk, loc, pc, None)\ note call = \call1 (obj'\M'(ps)) = \(a, M, vs)\\ thus ?case proof(cases rule: call1_callE) case CallObj with bisim1 call have "pc < length (compE2 obj)" by(auto intro: bisim1_call_pcD) with call ins CallObj IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1Call1) next case (CallParams v) from bisim1 have "pc \ length (compE2 obj)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 obj)" with bisim1 ins CallParams have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 obj)" by(cases "pc < length (compE2 obj)") auto with bisim1 CallParams have [simp]: "stk = [v]" "loc = xs" by(auto dest: bisim1_Val_length_compE2D) from IH2[of loc] CallParams ins show ?thesis apply(clarsimp simp add: compEs2_map_Val is_vals_conv split: if_split_asm) apply(drule bisim1_bisims1.bisim1CallParams) apply(auto simp add: neq_Nil_conv) done next case [simp]: Call from bisim1 have "pc \ length (compE2 obj)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 obj)" - with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 obj)" by(cases "pc < length (compE2 obj)") auto - with ins have [simp]: "vs = []" by(auto simp add: nth_append compEs2_map_Val split: if_split_asm) + with ins have [simp]: "vs = []" by(auto simp add: compEs2_map_Val split: if_split_asm) from bisim1 have [simp]: "stk = [Addr a]" "xs = loc" by(auto dest: bisim1_Val_length_compE2D) from ins show ?thesis by(auto intro: bisim1CallThrow[of "[]" "[]", simplified]) qed next case (bisim1CallParams ps n ps' xs stk loc pc obj M' v) note IH2 = \\calls1 ps' = \(a, M, vs)\; compEs2 ps ! pc = Invoke M n0; pc < length (compEs2 ps) \ \ ?concls ps n ps' xs pc stk loc\ note ins = \compE2 (obj\M'(ps)) ! (length (compE2 obj) + pc) = Invoke M n0\ note bisim2 = \P,ps,h \ (ps', xs) [\] (stk, loc, pc, None)\ note call = \call1 (Val v\M'(ps')) = \(a, M, vs)\\ thus ?case proof(cases rule: call1_callE) case CallObj thus ?thesis by simp next case (CallParams v') hence [simp]: "v' = v" and call': "calls1 ps' = \(a, M, vs)\" by auto from bisim2 call' have "pc < length (compEs2 ps)" by(auto intro: bisims1_calls_pcD) with IH2 CallParams ins show ?thesis by(auto simp add: is_vals_conv split: if_split_asm intro: bisim1_bisims1.bisim1CallParams) next case Call hence [simp]: "v = Addr a" "M' = M" "ps' = map Val vs" by auto from bisim2 have "pc \ length (compEs2 ps)" by(auto dest: bisims1_pc_length_compEs2) moreover { assume pc: "pc < length (compEs2 ps)" - with bisim2 ins have False by(auto dest: bisims_Val_pc_not_Invoke simp add: nth_append) } + with bisim2 ins have False by(auto dest: bisims_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compEs2 ps)" by(cases "pc < length (compEs2 ps)") auto from bisim2 have [simp]: "stk = rev vs" "xs = loc" by(auto dest: bisims1_Val_length_compEs2D) from bisim2 have "length ps = length vs" by(auto dest: bisims1_lengthD) with ins show ?thesis by(auto intro: bisim1CallThrow) qed next case (bisims1List1 e n e' xs stk loc pc es) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0; pc < length (compE2 e) \ \ ?concl e n e' xs pc stk loc\ note IH2 = \\xs. \calls1 es = \(a, M, vs)\; compEs2 es ! 0 = Invoke M n0; 0 < length (compEs2 es) \ \ ?concls es n es xs 0 [] xs\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \calls1 (e' # es) = \(a, M, vs)\\ note ins = \compEs2 (e # es) ! pc = Invoke M n0\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis - by(auto simp add: nth_append split_beta intro: bisim1_bisims1.bisims1List1) + by(auto intro: bisim1_bisims1.bisims1List1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e)" - with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 e)" by(cases "pc < length (compE2 e)") auto with bisim1 have [simp]: "stk = [v]" "loc = xs" by(auto dest: bisim1_Val_length_compE2D) from call have "es \ []" by(cases es) simp_all with IH2[of loc] call ins show ?thesis by(auto split: if_split_asm dest: bisims1List2) qed qed(auto split: if_split_asm bop.split_asm intro: bisim1_bisims1.intros dest: bisim1_pc_length_compE2) lemma bisim1_max_stack: "P,e,h \ (e', xs) \ (stk, loc, pc, xcp) \ length stk \ max_stack e" and bisims1_max_stacks: "P,es,h \ (es', xs) [\] (stk, loc, pc, xcp) \ length stk \ max_stacks es" apply(induct "(e', xs)" "(stk, loc, pc, xcp)" and "(es', xs)" "(stk, loc, pc, xcp)" arbitrary: e' xs stk loc pc xcp and es' xs stk loc pc xcp rule: bisim1_bisims1.inducts) apply(auto simp add: max_stack1[simplified] max_def max_stacks_ge_length) apply(drule sym, simp add: max_stacks_ge_length, drule sym, simp, rule le_trans[OF max_stacks_ge_length], simp) done inductive bisim1_fr :: "'addr J1_prog \ 'heap \ 'addr expr1 \ 'addr locals1 \ 'addr frame \ bool" for P :: "'addr J1_prog" and h :: 'heap where "\ P \ C sees M:Ts\T = \body\ in D; P,blocks1 0 (Class D#Ts) body, h \ (e, xs) \ (stk, loc, pc, None); call1 e = \(a, M', vs)\; max_vars e \ length xs \ \ bisim1_fr P h (e, xs) (stk, loc, C, M, pc)" declare bisim1_fr.intros [intro] declare bisim1_fr.cases [elim] lemma bisim1_fr_hext_mono: "\ bisim1_fr P h exs fr; hext h h' \ \ bisim1_fr P h' exs fr" by(auto intro: bisim1_hext_mono) lemma bisim1_max_vars: "P,E,h \ (e, xs) \ (stk, loc, pc, xcp) \ max_vars E \ max_vars e" and bisims1_max_varss: "P,Es,h \ (es,xs) [\] (stk,loc,pc,xcp) \ max_varss Es \ max_varss es" apply(induct E "(e, xs)" "(stk, loc, pc, xcp)" and Es "(es, xs)" "(stk, loc, pc, xcp)" arbitrary: e xs stk loc pc xcp and es xs stk loc pc xcp rule: bisim1_bisims1.inducts) apply(auto) done lemma bisim1_call_\Exec_move: "\ P,e,h \ (e', xs) \ (stk, loc, pc, None); call1 e' = \(a, M', vs)\; n + max_vars e' \ length xs; \ contains_insync e \ \ \pc' loc' stk'. \Exec_mover_a P t e h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None) \ pc' < length (compE2 e) \ compE2 e ! pc' = Invoke M' (length vs) \ P,e,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" (is "\ _; _; _; _ \ \ ?concl e n e' xs pc stk loc") and bisims1_calls_\Exec_moves: "\ P,es,h \ (es',xs) [\] (stk, loc, pc, None); calls1 es' = \(a, M', vs)\; n + max_varss es' \ length xs; \ contains_insyncs es \ \ \pc' stk' loc'. \Exec_movesr_a P t es h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None) \ pc' < length (compEs2 es) \ compEs2 es ! pc' = Invoke M' (length vs) \ P,es,h \ (es', xs) [\] (rev vs @ Addr a # stk', loc', pc', None)" (is "\_; _; _; _ \ \ ?concls es n es' xs pc stk loc") proof(induct e "n :: nat" e' xs stk loc pc xcp\"None :: 'addr option" and es "n :: nat" es' xs stk loc pc xcp\"None :: 'addr option" rule: bisim1_bisims1_inducts_split) case bisim1Val2 thus ?case by auto next case bisim1New thus ?case by auto next case bisim1NewArray thus ?case by auto (fastforce intro: bisim1_bisims1.bisim1NewArray elim!: NewArray_\ExecrI intro!: exI) next case bisim1Cast thus ?case by(auto)(fastforce intro: bisim1_bisims1.bisim1Cast elim!: Cast_\ExecrI intro!: exI)+ next case bisim1InstanceOf thus ?case by(auto)(fastforce intro: bisim1_bisims1.bisim1InstanceOf elim!: InstanceOf_\ExecrI intro!: exI)+ next case bisim1Val thus ?case by auto next case bisim1Var thus ?case by auto next case (bisim1BinOp1 e1 n e' xs stk loc pc e2 bop) note IH1 = \\call1 e' = \(a, M', vs)\; n + max_vars e' \ length xs; \ contains_insync e1 \ \ ?concl e1 n e' xs pc stk loc\ note IH2 = \\xs. \call1 e2 = \(a, M', vs)\; n + max_vars e2 \ length xs; \ contains_insync e2 \ \ ?concl e2 n e2 xs 0 [] xs\ note call = \call1 (e' \bop\ e2) = \(a, M', vs)\\ note len = \n + max_vars (e' \bop\ e2) \ length xs\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note cs = \\ contains_insync (e1 \bop\ e2)\ show ?case proof(cases "is_val e'") case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "\Exec_mover_a P t e1 h (stk, loc, pc, None) ([v], loc, length (compE2 e1), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_mover_a P t (e1\bop\e2) h (stk, loc, pc, None) ([v], loc, length (compE2 e1), None)" by-(rule BinOp_\ExecrI1) also from call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e2 h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e2 ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e2)" and bisim': "P,e2,h \ (e2, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from BinOp_\ExecrI2[OF exec, of e1 bop v] have "\Exec_mover_a P t (e1\bop\e2) h ([v], loc, length (compE2 e1), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 e1) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,e1\bop\e2,h \ (Val v \bop\ e2, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e1) + pc', None)" by(rule bisim1BinOp2) ultimately show ?thesis using ins by fastforce next case False with IH1 len False call cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1BinOp1 elim!: BinOp_\ExecrI1 intro!: exI) qed next case (bisim1BinOp2 e2 n e' xs stk loc pc e1 bop v1) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e2,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from exec have "\Exec_mover_a P t (e1 \bop\ e2) h (stk @ [v1], loc, length (compE2 e1) + pc, None) ((rev vs @ Addr a # stk') @ [v1], loc', length (compE2 e1) + pc', None)" by(rule BinOp_\ExecrI2) moreover 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) ultimately show ?case using pc' by(fastforce) next case bisim1LAss1 thus ?case by(auto)(fastforce intro: bisim1_bisims1.bisim1LAss1 elim!: LAss_\ExecrI intro!: exI) next case bisim1LAss2 thus ?case by simp next case (bisim1AAcc1 A n a' xs stk loc pc i) note IH1 = \\call1 a' = \(a, M', vs)\; n + max_vars a' \ length xs; \ contains_insync A\ \ ?concl A n a' xs pc stk loc\ note IH2 = \\xs. \call1 i = \(a, M', vs)\; n + max_vars i \ length xs; \ contains_insync i\ \ ?concl i n i xs 0 [] xs\ note call = \call1 (a'\i\) = \(a, M', vs)\\ note len = \n + max_vars (a'\i\) \ length xs\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note cs = \\ contains_insync (A\i\)\ show ?case proof(cases "is_val a'") case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "\Exec_mover_a P t A h (stk, loc, pc, None) ([v], loc, length (compE2 A), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_mover_a P t (A\i\) h (stk, loc, pc, None) ([v], loc, length (compE2 A), None)" by-(rule AAcc_\ExecrI1) also from call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t i h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 i ! pc' = Invoke M' (length vs)" "pc' < length (compE2 i)" and bisim': "P,i,h \ (i, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from AAcc_\ExecrI2[OF exec, of A v] have "\Exec_mover_a P t (A\i\) h ([v], loc, length (compE2 A), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 A) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,A\i\,h \ (Val v\i\, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 A) + pc', None)" by(rule bisim1AAcc2) ultimately show ?thesis using ins by fastforce next case False with IH1 len False call cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1AAcc1 elim!: AAcc_\ExecrI1 intro!: exI) qed next case (bisim1AAcc2 i n i' xs stk loc pc A v) then obtain pc' loc' stk' where pc': "pc' < length (compE2 i)" "compE2 i ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t i h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,i,h \ (i', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from exec have "\Exec_mover_a P t (A\i\) h (stk @ [v], loc, length (compE2 A) + pc, None) ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 A) + pc', None)" by(rule AAcc_\ExecrI2) moreover from bisim' have "P,A\i\,h \ (Val v\i'\, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 A) + pc', None)" by(rule bisim1_bisims1.bisim1AAcc2) ultimately show ?case using pc' by(fastforce) next case (bisim1AAss1 A n a' xs stk loc pc i e) note IH1 = \\call1 a' = \(a, M', vs)\; n + max_vars a' \ length xs; \ contains_insync A \ \ ?concl A n a' xs pc stk loc\ note IH2 = \\xs. \call1 i = \(a, M', vs)\; n + max_vars i \ length xs; \ contains_insync i\ \ ?concl i n i xs 0 [] xs\ note IH3 = \\xs. \call1 e = \(a, M', vs)\; n + max_vars e \ length xs; \ contains_insync e\ \ ?concl e n e xs 0 [] xs\ note call = \call1 (a'\i\ := e) = \(a, M', vs)\\ note len = \n + max_vars (a'\i\ := e) \ length xs\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note bisim2 = \P,i,h \ (i, loc) \ ([], loc, 0, None)\ note cs = \\ contains_insync (A\i\ := e)\ show ?case proof(cases "is_val a'") case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "\Exec_mover_a P t A h (stk, loc, pc, None) ([v], loc, length (compE2 A), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence exec: "\Exec_mover_a P t (A\i\ := e) h (stk, loc, pc, None) ([v], loc, length (compE2 A), None)" by-(rule AAss_\ExecrI1) show ?thesis proof(cases "is_val i") case True then obtain v' where [simp]: "i = Val v'" by auto note exec also from bisim2 have "\Exec_mover_a P t i h ([], loc, 0, None) ([v'], loc, length (compE2 i), None)" by(auto dest!: bisim1Val2D1) from AAss_\ExecrI2[OF this, of A e v] have "\Exec_mover_a P t (A\i\ := e) h ([v], loc, length (compE2 A), None) ([v', v], loc, length (compE2 A) + length (compE2 i), None)" by simp also (rtranclp_trans) from call IH3[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e h ([], loc, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e)" and bisim': "P,e,h \ (e, loc) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from AAss_\ExecrI3[OF exec, of A i v' v] have "\Exec_mover_a P t (A\i\ := e) h ([v', v], loc, length (compE2 A) + length (compE2 i), None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,A\i\ := e,h \ (Val v\Val v'\ := e, xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by - (rule bisim1AAss3, simp) ultimately show ?thesis using ins by fastforce next case False note exec also from False call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t i h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 i ! pc' = Invoke M' (length vs)" "pc' < length (compE2 i)" and bisim': "P,i,h \ (i, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from AAss_\ExecrI2[OF exec, of A e v] have "\Exec_mover_a P t (A\i\ := e) h ([v], loc, length (compE2 A), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 A) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,A\i\ := e,h \ (Val v\i\ := e, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 A) + pc', None)" by(rule bisim1AAss2) ultimately show ?thesis using ins False by(fastforce intro!: exI) qed next case False with IH1 len False call cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1AAss1 elim!: AAss_\ExecrI1 intro!: exI) qed next case (bisim1AAss2 i n i' xs stk loc pc A e v) note IH2 = \\call1 i' = \(a, M', vs)\; n + max_vars i' \ length xs; \ contains_insync i\ \ ?concl i n i' xs pc stk loc\ note IH3 = \\xs. \call1 e = \(a, M', vs)\; n + max_vars e \ length xs; \ contains_insync e\ \ ?concl e n e xs 0 [] xs\ note call = \call1 (Val v\i'\ := e) = \(a, M', vs)\\ note len = \n + max_vars (Val v\i'\ := e) \ length xs\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, None)\ note cs = \\ contains_insync (A\i\ := e)\ show ?case proof(cases "is_val i'") case True then obtain v' where [simp]: "i' = Val v'" by auto from bisim2 have exec: "\Exec_mover_a P t i h (stk, loc, pc, None) ([v'], loc, length (compE2 i), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) from AAss_\ExecrI2[OF exec, of A e v] have "\Exec_mover_a P t (A\i\ := e) h (stk @ [v], loc, length (compE2 A) + pc, None) ([v', v], loc, length (compE2 A) + length (compE2 i), None)" by simp also from call IH3[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e)" and bisim': "P,e,h \ (e, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from AAss_\ExecrI3[OF exec, of A i v' v] have "\Exec_mover_a P t (A\i\ := e) h ([v', v], loc, length (compE2 A) + length (compE2 i), None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,A\i\ := e,h \ (Val v\Val v'\ := e, xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by(rule bisim1AAss3) ultimately show ?thesis using ins by(fastforce intro!: exI) next case False with IH2 len call cs obtain pc' loc' stk' where ins: "pc' < length (compE2 i)" "compE2 i ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t i h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,i,h \ (i', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from bisim' have "P,A\i\ := e,h \ (Val v\i'\ := e, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 A) + pc', None)" by(rule bisim1_bisims1.bisim1AAss2) with AAss_\ExecrI2[OF exec, of A e v] ins False show ?thesis by(auto intro!: exI) qed next case (bisim1AAss3 e n e' xs stk loc pc A i v v') then obtain pc' loc' stk' where pc': "pc' < length (compE2 e)" "compE2 e ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from exec have "\Exec_mover_a P t (A\i\:=e) h (stk @ [v', v], loc, length (compE2 A) + length (compE2 i) + pc, None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by(rule AAss_\ExecrI3) moreover from bisim' have "P,A\i\ := e,h \ (Val v\Val v'\ := e', xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by(rule bisim1_bisims1.bisim1AAss3) ultimately show ?case using pc' by(fastforce intro!: exI) next case bisim1AAss4 thus ?case by simp next case bisim1ALength thus ?case by(auto)(fastforce intro: bisim1_bisims1.bisim1ALength elim!: ALength_\ExecrI intro!: exI) next case bisim1FAcc thus ?case by(auto)(fastforce intro: bisim1_bisims1.bisim1FAcc elim!: FAcc_\ExecrI intro!: exI) next case (bisim1FAss1 e n e' xs stk loc pc e2 F D) note IH1 = \\call1 e' = \(a, M', vs)\; n + max_vars e' \ length xs; \ contains_insync e\ \ ?concl e n e' xs pc stk loc\ note IH2 = \\xs. \call1 e2 = \(a, M', vs)\; n + max_vars e2 \ length xs; \ contains_insync e2\ \ ?concl e2 n e2 xs 0 [] xs\ note call = \call1 (e'\F{D} := e2) = \(a, M', vs)\\ note len = \n + max_vars (e'\F{D} := e2) \ length xs\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note cs = \\ contains_insync (e\F{D} := e2)\ show ?case proof(cases "is_val e'") case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "\Exec_mover_a P t e h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_mover_a P t (e\F{D} := e2) h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" by-(rule FAss_\ExecrI1) also from call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e2 h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e2 ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e2)" and bisim': "P,e2,h \ (e2, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from FAss_\ExecrI2[OF exec, of e F D v] have "\Exec_mover_a P t (e\F{D} := e2) h ([v], loc, length (compE2 e), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 e) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,e\F{D} := e2,h \ (Val v\F{D} := e2, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e) + pc', None)" by(rule bisim1FAss2) ultimately show ?thesis using ins by fastforce next case False with IH1 len False call cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1FAss1 elim!: FAss_\ExecrI1 intro!: exI) qed next case (bisim1FAss2 e2 n e' xs stk loc pc e F D v) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e2,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from exec have "\Exec_mover_a P t (e\F{D} := e2) h (stk @ [v], loc, length (compE2 e) + pc, None) ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e) + pc', None)" by(rule FAss_\ExecrI2) moreover from bisim' have "P,e\F{D} := e2,h \ (Val v\F{D} := e', xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e) + pc', None)" by(rule bisim1_bisims1.bisim1FAss2) ultimately show ?case using pc' by(fastforce) next case bisim1FAss3 thus ?case by simp next case (bisim1CAS1 e1 n e' xs stk loc pc e2 e3 D F) note IH1 = \\call1 e' = \(a, M', vs)\; n + max_vars e' \ length xs; \ contains_insync e1 \ \ ?concl e1 n e' xs pc stk loc\ note IH2 = \\xs. \call1 e2 = \(a, M', vs)\; n + max_vars e2 \ length xs; \ contains_insync e2\ \ ?concl e2 n e2 xs 0 [] xs\ note IH3 = \\xs. \call1 e3 = \(a, M', vs)\; n + max_vars e3 \ length xs; \ contains_insync e3\ \ ?concl e3 n e3 xs 0 [] xs\ note call = \call1 _ = \(a, M', vs)\\ note len = \n + max_vars _ \ length xs\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note bisim2 = \P,e2,h \ (e2, loc) \ ([], loc, 0, None)\ note cs = \\ contains_insync _\ show ?case proof(cases "is_val e'") case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "\Exec_mover_a P t e1 h (stk, loc, pc, None) ([v], loc, length (compE2 e1), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence exec: "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h (stk, loc, pc, None) ([v], loc, length (compE2 e1), None)" by-(rule CAS_\ExecrI1) show ?thesis proof(cases "is_val e2") case True then obtain v' where [simp]: "e2 = Val v'" by auto note exec also from bisim2 have "\Exec_mover_a P t e2 h ([], loc, 0, None) ([v'], loc, length (compE2 e2), None)" by(auto dest!: bisim1Val2D1) from CAS_\ExecrI2[OF this, of e1 D F e3] have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h ([v], loc, length (compE2 e1), None) ([v', v], loc, length (compE2 e1) + length (compE2 e2), None)" by simp also (rtranclp_trans) from call IH3[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e3 h ([], loc, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e3 ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e3)" and bisim': "P,e3,h \ (e3, loc) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from CAS_\ExecrI3[OF exec, of e1 D F e2 v' v] have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h ([v', v], loc, length (compE2 e1) + length (compE2 e2), None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,e1\compareAndSwap(D\F, e2, e3),h \ (Val v\compareAndSwap(D\F, Val v', e3), xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by - (rule bisim1CAS3, simp) ultimately show ?thesis using ins by fastforce next case False note exec also from False call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e2 h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e2 ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e2)" and bisim': "P,e2,h \ (e2, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from CAS_\ExecrI2[OF exec, of e1 D F e3 v] have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h ([v], loc, length (compE2 e1), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 e1) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,e1\compareAndSwap(D\F, e2, e3),h \ (Val v\compareAndSwap(D\F, e2, e3), xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e1) + pc', None)" by(rule bisim1CAS2) ultimately show ?thesis using ins False by(fastforce intro!: exI) qed next case False with IH1 len False call cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1CAS1 elim!: CAS_\ExecrI1 intro!: exI) qed next case (bisim1CAS2 e2 n e2' xs stk loc pc e1 e3 D F v) note IH2 = \\call1 e2' = \(a, M', vs)\; n + max_vars e2' \ length xs; \ contains_insync e2\ \ ?concl e2 n e2' xs pc stk loc\ note IH3 = \\xs. \call1 e3 = \(a, M', vs)\; n + max_vars e3 \ length xs; \ contains_insync e3\ \ ?concl e3 n e3 xs 0 [] xs\ note call = \call1 _ = \(a, M', vs)\\ note len = \n + max_vars _ \ length xs\ note bisim2 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, None)\ note cs = \\ contains_insync _\ show ?case proof(cases "is_val e2'") case True then obtain v' where [simp]: "e2' = Val v'" by auto from bisim2 have exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) ([v'], loc, length (compE2 e2), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) from CAS_\ExecrI2[OF exec, of e1 D F e3 v] have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h (stk @ [v], loc, length (compE2 e1) + pc, None) ([v', v], loc, length (compE2 e1) + length (compE2 e2), None)" by simp also from call IH3[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e3 h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e3 ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e3)" and bisim': "P,e3,h \ (e3, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from CAS_\ExecrI3[OF exec, of e1 D F e2 v' v] have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h ([v', v], loc, length (compE2 e1) + length (compE2 e2), None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,e1\compareAndSwap(D\F, e2, e3),h \ (Val v\compareAndSwap(D\F, Val v', e3), xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by(rule bisim1CAS3) ultimately show ?thesis using ins by(fastforce intro!: exI) next case False with IH2 len call cs obtain pc' loc' stk' where ins: "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e2,h \ (e2', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from bisim' have "P,e1\compareAndSwap(D\F, e2, e3),h \ (Val v\compareAndSwap(D\F, e2', e3), xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e1) + pc', None)" by(rule bisim1_bisims1.bisim1CAS2) with CAS_\ExecrI2[OF exec, of e1 D F e3 v] ins False show ?thesis by(auto intro!: exI) qed next case (bisim1CAS3 e3 n e3' xs stk loc pc e1 e2 D F v v') then obtain pc' loc' stk' where pc': "pc' < length (compE2 e3)" "compE2 e3 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e3 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e3,h \ (e3', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from exec have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h (stk @ [v', v], loc, length (compE2 e1) + length (compE2 e2) + pc, None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by(rule CAS_\ExecrI3) moreover from bisim' have "P,e1\compareAndSwap(D\F, e2, e3),h \ (Val v\compareAndSwap(D\F, Val v', e3'), xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by(rule bisim1_bisims1.bisim1CAS3) ultimately show ?case using pc' by(fastforce intro!: exI) next case (bisim1Call1 obj n obj' xs stk loc pc ps M) note IH1 = \\call1 obj' = \(a, M', vs)\; n + max_vars obj' \ length xs; \ contains_insync obj\ \ ?concl obj n obj' xs pc stk loc\ note IH2 = \\xs. \calls1 ps = \(a, M', vs)\; n + max_varss ps \ length xs; \ contains_insyncs ps\ \ ?concls ps n ps xs 0 [] xs\ note len = \n + max_vars (obj'\M(ps)) \ length xs\ note bisim1 = \P,obj,h \ (obj', xs) \ (stk, loc, pc, None)\ note call = \call1 (obj'\M(ps)) = \(a, M', vs)\\ note cs = \\ contains_insync (obj\M(ps))\ from call show ?case proof(cases rule: call1_callE) case CallObj hence "\ is_val obj'" by auto with CallObj IH1 len cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1Call1 elim!: Call_\ExecrI1 intro!: exI) next case (CallParams v) with bisim1 have "\Exec_mover_a P t obj h (stk, loc, pc, None) ([v], loc, length (compE2 obj), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_mover_a P t (obj\M(ps)) h (stk, loc, pc, None) ([v], loc, length (compE2 obj), None)" by-(rule Call_\ExecrI1) also from IH2[of loc] CallParams len cs obtain pc' stk' loc' where exec: "\Exec_movesr_a P t ps h ([], loc, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compEs2 ps ! pc' = Invoke M' (length vs)" "pc' < length (compEs2 ps)" and bisim': "P,ps,h \ (ps, xs) [\] (rev vs @ Addr a # stk',loc',pc',None)" by auto from Call_\ExecrI2[OF exec, of obj M v] have "\Exec_mover_a P t (obj\M(ps)) h ([v], loc, length (compE2 obj), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 obj) + pc', None)" by simp also (rtranclp_trans) have "P,obj\M(ps),h \ (Val v\M(ps), xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 obj) + pc', None)" using bisim' by(rule bisim1CallParams) ultimately show ?thesis using ins CallParams by fastforce next case [simp]: Call from bisim1 have "\Exec_mover_a P t obj h (stk, loc, pc, None) ([Addr a], loc, length (compE2 obj), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_mover_a P t (obj\M(ps)) h (stk, loc, pc, None) ([Addr a], loc, length (compE2 obj), None)" by-(rule Call_\ExecrI1) also have "\Exec_movesr_a P t ps h ([], xs, 0, None) (rev vs, xs, length (compEs2 ps), None)" proof(cases vs) case Nil with Call show ?thesis by(auto) next case Cons with Call bisims1_Val_\Exec_moves[OF bisims1_refl[of P h "map Val vs" loc]] show ?thesis by(auto simp add: bsoks_def) qed from Call_\ExecrI2[OF this, of obj M "Addr a"] have "\Exec_mover_a P t (obj\M(ps)) h ([Addr a], loc, length (compE2 obj), None) (rev vs @ [Addr a], xs, length (compE2 obj) + length (compEs2 ps), None)" by simp also (rtranclp_trans) have "P,ps,h \ (map Val vs,xs) [\] (rev vs,xs,length (compEs2 ps),None)" by(rule bisims1_map_Val_append[OF bisims1Nil, simplified])(simp_all add: bsoks_def) hence "P,obj\M(ps),h \ (addr a\M(map Val vs), xs) \ (rev vs @ [Addr a], xs, length (compE2 obj) + length (compEs2 ps), None)" by(rule bisim1CallParams) ultimately show ?thesis by fastforce qed next case (bisim1CallParams ps n ps' xs stk loc pc obj M v) note IH2 = \\calls1 ps' = \(a, M', vs)\; n + max_varss ps' \ length xs; \ contains_insyncs ps\ \ ?concls ps n ps' xs pc stk loc\ note bisim2 = \P,ps,h \ (ps', xs) [\] (stk, loc, pc, None)\ note call = \call1 (Val v\M(ps')) = \(a, M', vs)\\ note len = \n + max_vars (Val v\M(ps')) \ length xs\ note cs = \\ contains_insync (obj\M(ps))\ from call show ?case proof(cases rule: call1_callE) case CallObj thus ?thesis by simp next case (CallParams v') with IH2 len cs obtain pc' stk' loc' where exec: "\Exec_movesr_a P t ps h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "pc' < length (compEs2 ps)" "compEs2 ps ! pc' = Invoke M' (length vs)" and bisim': "P,ps,h \ (ps', xs) [\] (rev vs @ Addr a # stk',loc',pc',None)" by auto from exec have "\Exec_mover_a P t (obj\M(ps)) h (stk @ [v], loc, length (compE2 obj) + pc, None) ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 obj) + pc', None)" by(rule Call_\ExecrI2) moreover have "P,obj\M(ps),h \ (Val v\M(ps'), xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 obj) + pc', None)" using bisim' by(rule bisim1_bisims1.bisim1CallParams) ultimately show ?thesis using ins by fastforce next case Call hence [simp]: "v = Addr a" "ps' = map Val vs" "M' = M" by simp_all have "xs = loc \ \Exec_movesr_a P t ps h (stk, loc, pc, None) (rev vs, loc, length (compEs2 ps), None)" proof(cases "pc < length (compEs2 ps)") case True with bisim2 show ?thesis by(auto dest: bisims1_Val_\Exec_moves) next case False from bisim2 have "pc \ length (compEs2 ps)" by(rule bisims1_pc_length_compEs2) with False have "pc = length (compEs2 ps)" by simp with bisim2 show ?thesis by(auto dest: bisims1_Val_length_compEs2D) qed then obtain [simp]: "xs = loc" and exec: "\Exec_movesr_a P t ps h (stk, loc, pc, None) (rev vs, loc, length (compEs2 ps), None)" .. from exec have "\Exec_mover_a P t (obj\M(ps)) h (stk @ [v], loc, length (compE2 obj) + pc, None) (rev vs @ [v], loc, length (compE2 obj) + length (compEs2 ps), None)" by(rule Call_\ExecrI2) moreover from bisim2 have len: "length ps = length ps'" by(auto dest: bisims1_lengthD) moreover have "P,ps,h \ (map Val vs,xs) [\] (rev vs,xs,length (compEs2 ps),None)" using len by-(rule bisims1_map_Val_append[OF bisims1Nil, simplified], simp_all) hence "P,obj\M(ps),h \ (addr a\M(map Val vs), xs) \ (rev vs @ [Addr a], xs, length (compE2 obj) + length (compEs2 ps), None)" by(rule bisim1_bisims1.bisim1CallParams) ultimately show ?thesis by fastforce qed next case bisim1BlockSome1 thus ?case by simp next case bisim1BlockSome2 thus ?case by simp next case (bisim1BlockSome4 e n e' xs stk loc pc V T v) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e)" "compE2 e ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto note Block_\ExecrI_Some[OF exec, of V T v] moreover from bisim' have "P,{V:T=\v\; e},h \ ({V:T=None; e'}, xs) \ (rev vs @ Addr a # stk', loc', Suc (Suc pc'), None)" by(rule bisim1_bisims1.bisim1BlockSome4) ultimately show ?case using pc' by fastforce next case (bisim1BlockNone e n e' xs stk loc pc V T) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e)" "compE2 e ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto note Block_\ExecrI_None[OF exec, of V T] moreover from bisim' have "P,{V:T=None; e},h \ ({V:T=None; e'}, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by(rule bisim1_bisims1.bisim1BlockNone) ultimately show ?case using pc' by fastforce next case bisim1Sync1 thus ?case by (auto)(fastforce intro: bisim1_bisims1.bisim1Sync1 elim!: Sync_\ExecrI intro!: exI) next case bisim1Sync2 thus ?case by simp next case bisim1Sync3 thus ?case by simp next case bisim1Sync4 thus ?case by (auto)(fastforce intro: bisim1_bisims1.bisim1Sync4 elim!: Insync_\ExecrI intro!: exI) 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 by (auto)(fastforce intro: bisim1_bisims1.bisim1Seq1 elim!: Seq_\ExecrI1 intro!: exI) next case (bisim1Seq2 e2 n e' xs stk loc pc e1) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e2,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from Seq_\ExecrI2[OF exec, of e1] pc' bisim' show ?case by(fastforce intro: bisim1_bisims1.bisim1Seq2 intro!: exI) next case bisim1Cond1 thus ?case by (auto)(fastforce intro: bisim1_bisims1.bisim1Cond1 elim!: Cond_\ExecrI1 intro!: exI)+ next case (bisim1CondThen e1 n e' xs stk loc pc e e2) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e1)" "compE2 e1 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e1 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e1,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from Cond_\ExecrI2[OF exec] pc' bisim' show ?case by(fastforce intro: bisim1_bisims1.bisim1CondThen intro!: exI) next case (bisim1CondElse e2 n e' xs stk loc pc e e1) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e2,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from Cond_\ExecrI3[OF exec] pc' bisim' show ?case by (fastforce intro: bisim1_bisims1.bisim1CondElse intro!: exI) next case bisim1While1 thus ?case by simp next case bisim1While3 thus ?case by (auto)(fastforce intro: bisim1_bisims1.bisim1While3 elim!: While_\ExecrI1 intro!: exI)+ next case bisim1While4 thus ?case by (auto)(fastforce intro!: While_\ExecrI2 bisim1_bisims1.bisim1While4 exI)+ next case bisim1While6 thus ?case by simp next case bisim1While7 thus ?case by simp next case bisim1Throw1 thus ?case by (auto)(fastforce intro!: exI bisim1_bisims1.bisim1Throw1 elim!: Throw_\ExecrI)+ next case bisim1Try thus ?case by (auto)(fastforce intro: bisim1_bisims1.bisim1Try elim!: Try_\ExecrI1 intro!: exI)+ next case (bisim1TryCatch1 e n a' xs stk loc pc C' C e2 V) note IH2 = \\xs. \call1 e2 = \(a, M', vs)\; Suc n + max_vars e2 \ length xs; \ contains_insync e2 \ \ ?concl e2 (Suc V) e2 xs 0 [] xs\ note bisim1 = \P,e,h \ (Throw a', xs) \ (stk, loc, pc, \a'\)\ note bisim2 = \\xs. P,e2,h \ (e2, xs) \ ([], xs, 0, None)\ note len = \n + max_vars {V:Class C=None; e2} \ length (xs[V := Addr a'])\ note cs = \\ contains_insync (try e catch(C V) e2)\ from bisim1 have [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from len have "\Exec_mover_a P t (try e catch(C V) e2) h ([Addr a'], loc, Suc (length (compE2 e)), None) ([], loc[V := Addr a'], Suc (Suc (length (compE2 e))), None)" by -(rule \Execr1step,auto simp add: exec_move_def intro: \move2_\moves2.intros exec_instr) also from IH2[of "loc[V := Addr a']"] len \call1 {V:Class C=None; e2} = \(a, M', vs)\\ cs obtain pc' loc' stk' where exec: "\Exec_mover_a P t e2 h ([], loc[V := Addr a'], 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and bisim': "P,e2,h \ (e2, loc[V := Addr a']) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from Try_\ExecrI2[OF exec, of e C V] have "\Exec_mover_a P t (try e catch(C V) e2) h ([], loc[V := Addr a'], Suc (Suc (length (compE2 e))), None) (rev vs @ Addr a # stk', loc', Suc (Suc (length (compE2 e) + pc')), None)" by simp also from bisim' have "P,try e catch(C V) e2,h \ ({V:Class C=None; e2}, loc[V := Addr a']) \ (rev vs @ Addr a # stk', loc', (Suc (Suc (length (compE2 e) + pc'))), None)" by(rule bisim1TryCatch2) ultimately show ?case using ins by fastforce next case bisim1TryCatch2 thus ?case by (auto)(fastforce intro!: Try_\ExecrI2 bisim1_bisims1.bisim1TryCatch2 exI)+ next case bisims1Nil thus ?case by simp next case (bisims1List1 e n e' xs stk loc pc es) note IH1 = \\call1 e' = \(a, M', vs)\; n + max_vars e' \ length xs; \ contains_insync e\ \ ?concl e n e' xs pc stk loc\ note IH2 = \\xs. \calls1 es = \(a, M', vs)\; n + max_varss es \ length xs; \ contains_insyncs es\ \ ?concls es n es xs 0 [] xs\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \calls1 (e' # es) = \(a, M', vs)\\ note len = \n + max_varss (e' # es) \ length xs\ note cs = \\ contains_insyncs (e # es)\ show ?case proof(cases "is_val e'") case True then obtain v where [simp]: "e' = Val v" by auto with bisim1 have "\Exec_mover_a P t e h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_movesr_a P t (e # es) h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" by-(rule \Exec_mover_\Exec_movesr) also from call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_movesr_a P t es h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compEs2 es ! pc' = Invoke M' (length vs)" "pc' < length (compEs2 es)" and bisim': "P,es,h \ (es, xs) [\] (rev vs @ Addr a # stk',loc',pc',None)" by auto from append_\Exec_movesr[OF _ exec, of "[v]" "[e]"] have "\Exec_movesr_a P t (e # es) h ([v], loc, length (compE2 e), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 e) + pc', None)" by simp also (rtranclp_trans) 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) ultimately show ?thesis using ins by fastforce next case False with call IH1 len cs show ?thesis by (auto)(fastforce intro!: \Exec_mover_\Exec_movesr bisim1_bisims1.bisims1List1 exI)+ qed next case (bisims1List2 es n es' xs stk loc pc e v) then obtain pc' stk' loc' where pc': "pc' < length (compEs2 es)" "compEs2 es ! pc' = Invoke M' (length vs)" and exec: "\Exec_movesr_a P t es h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,es,h \ (es', xs) [\] (rev vs @ Addr a # stk', loc', pc', None)" by auto note append_\Exec_movesr[OF _ exec, of "[v]" "[e]"] moreover 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) ultimately show ?case using pc' by fastforce qed lemma fixes P :: "'addr J1_prog" shows bisim1_inline_call_Val: "\ P,e,h \ (e', xs) \ (stk, loc, pc, None); call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0 \ \ length stk \ Suc (length vs) \ n0 = length vs \ P,e,h \ (inline_call (Val v) e', xs) \ (v # drop (Suc (length vs)) stk, loc, Suc pc, None)" (is "\ _; _; _ \ \ ?concl e n e' xs pc stk loc") and bisims1_inline_calls_Val: "\ P,es,h \ (es',xs) [\] (stk,loc,pc,None); calls1 es' = \(a, M, vs)\; compEs2 es ! pc = Invoke M n0 \ \ length stk \ Suc (length vs) \ n0 = length vs \ P,es,h \ (inline_calls (Val v) es', xs) [\] (v # drop (Suc (length vs)) stk,loc,Suc pc,None)" (is "\ _; _; _ \ \ ?concls es n es' xs pc stk loc") proof(induct "(e', xs)" "(stk, loc, pc, None :: 'addr option)" and "(es', xs)" "(stk, loc, pc, None :: 'addr option)" arbitrary: e' xs stk loc pc and es' xs stk loc pc rule: bisim1_bisims1.inducts) case bisim1Val2 thus ?case by simp next case bisim1New thus ?case by simp next case bisim1NewArray thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1NewArray) next case bisim1Cast thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Cast) next case bisim1InstanceOf thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1InstanceOf) next case bisim1Val thus ?case by simp next case bisim1Var thus ?case by simp next case (bisim1BinOp1 e1 e' xs stk loc pc bop e2) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e1 ! pc = Invoke M n0 \ \ ?concl e1 n e' xs pc stk loc\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 (e' \bop\ e2) = \(a, M, vs)\\ note ins = \compE2 (e1 \bop\ e2) ! pc = Invoke M n0\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e1)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1BinOp1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e1)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e1)" - with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 e1)" by(cases "pc < length (compE2 e1)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1BinOp2 e2 e' xs stk loc pc e1 bop v1) note IH2 = \\call1 e' = \(a, M, vs)\; compE2 e2 ! pc = Invoke M n0\ \ ?concl e2 n e' xs pc stk loc\ note bisim2 = \P,e2,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 (Val v1 \bop\ e') = \(a, M, vs)\\ note ins = \compE2 (e1 \bop\ e2) ! (length (compE2 e1) + pc) = Invoke M n0\ from call bisim2 have pc: "pc < length (compE2 e2)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 e2 ! pc = Invoke M n0" by(simp) from IH2 ins' pc call show ?case by(auto dest: bisim1_bisims1.bisim1BinOp2) next case bisim1LAss1 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1LAss1) next case bisim1LAss2 thus ?case by simp next case (bisim1AAcc1 A a' xs stk loc pc i) note IH1 = \\call1 a' = \(a, M, vs)\; compE2 A ! pc = Invoke M n0\ \ ?concl A n a' xs pc stk loc\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note call = \call1 (a'\i\) = \(a, M, vs)\\ note ins = \compE2 (A\i\) ! pc = Invoke M n0\ show ?case proof(cases "is_val a'") case False with bisim1 call have "pc < length (compE2 A)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1AAcc1) next case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "pc \ length (compE2 A)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 A)" - with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 A)" by(cases "pc < length (compE2 A)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1AAcc2 i i' xs stk loc pc A v) note IH2 = \\call1 i' = \(a, M, vs)\; compE2 i ! pc = Invoke M n0\ \ ?concl i n i' xs pc stk loc\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, None)\ note call = \call1 (Val v\i'\) = \(a, M, vs)\\ note ins = \compE2 (A\i\) ! (length (compE2 A) + pc) = Invoke M n0\ from call bisim2 have pc: "pc < length (compE2 i)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 i ! pc = Invoke M n0" by(simp) from IH2 ins' pc call show ?case by(auto dest: bisim1_bisims1.bisim1AAcc2) next case (bisim1AAss1 A a' xs stk loc pc i e) note IH1 = \\call1 a' = \(a, M, vs)\; compE2 A ! pc = Invoke M n0\ \ ?concl A n a' xs pc stk loc\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note call = \call1 (a'\i\ := e) = \(a, M, vs)\\ note ins = \compE2 (A\i\ := e) ! pc = Invoke M n0\ show ?case proof(cases "is_val a'") case False with bisim1 call have "pc < length (compE2 A)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1AAss1) next case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "pc \ length (compE2 A)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 A)" - with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 A)" by(cases "pc < length (compE2 A)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1AAss2 i i' xs stk loc pc A e v) note IH2 = \\call1 i' = \(a, M, vs)\; compE2 i ! pc = Invoke M n0\ \ ?concl i n i' xs pc stk loc\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, None)\ note call = \call1 (Val v\i'\ := e) = \(a, M, vs)\\ note ins = \compE2 (A\i\ := e) ! (length (compE2 A) + pc) = Invoke M n0\ show ?case proof(cases "is_val i'") case False with bisim2 call have pc: "pc < length (compE2 i)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 i ! pc = Invoke M n0" by(simp) from IH2 ins' pc False call show ?thesis by(auto dest: bisim1_bisims1.bisim1AAss2) next case True then obtain v where [simp]: "i' = Val v" by auto from bisim2 have "pc \ length (compE2 i)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 i)" - with bisim2 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + with bisim2 ins have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 i)" by(cases "pc < length (compE2 i)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1AAss3 e e' xs stk loc pc i A v v') note IH2 = \\call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0\ \ ?concl e n e' xs pc stk loc\ note bisim3 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 (Val v\Val v'\ := e') = \(a, M, vs)\\ note ins = \compE2 (i\A\ := e) ! (length (compE2 i) + length (compE2 A) + pc) = Invoke M n0\ from call bisim3 have pc: "pc < length (compE2 e)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 e ! pc = Invoke M n0" by(simp) from IH2 ins' pc call show ?case by(auto dest: bisim1_bisims1.bisim1AAss3) next case bisim1AAss4 thus ?case by simp next case bisim1ALength thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1ALength) next case bisim1FAcc thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1FAcc) next case (bisim1FAss1 e1 e' xs stk loc pc F D e2) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e1 ! pc = Invoke M n0\ \ ?concl e1 n e' xs pc stk loc\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 (e'\F{D} := e2) = \(a, M, vs)\\ note ins = \compE2 (e1\F{D} := e2) ! pc = Invoke M n0\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e1)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1FAss1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e1)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e1)" - with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 e1)" by(cases "pc < length (compE2 e1)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1FAss2 e2 e' xs stk loc pc e1 F D v1) note IH2 = \\call1 e' = \(a, M, vs)\; compE2 e2 ! pc = Invoke M n0\ \ ?concl e2 n e' xs pc stk loc\ note bisim2 = \P,e2,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 (Val v1\F{D} := e') = \(a, M, vs)\\ note ins = \compE2 (e1\F{D} := e2) ! (length (compE2 e1) + pc) = Invoke M n0\ from call bisim2 have pc: "pc < length (compE2 e2)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 e2 ! pc = Invoke M n0" by(simp) from IH2 ins' pc call show ?case by(auto dest: bisim1_bisims1.bisim1FAss2) next case bisim1FAss3 thus ?case by simp next case (bisim1CAS1 e1 e' xs stk loc pc D F e2 E3) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e1 ! pc = Invoke M n0\ \ ?concl e1 n e' xs pc stk loc\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 _ = \(a, M, vs)\\ note ins = \compE2 _ ! pc = Invoke M n0\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e1)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1CAS1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e1)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e1)" - with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 e1)" by(cases "pc < length (compE2 e1)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1CAS2 e2 e2' xs stk loc pc e1 D F e3 v) note IH2 = \\call1 e2' = \(a, M, vs)\; compE2 e2 ! pc = Invoke M n0\ \ ?concl e2 n e2' xs pc stk loc\ note bisim2 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, None)\ note call = \call1 _ = \(a, M, vs)\\ note ins = \compE2 _ ! (length (compE2 e1) + pc) = Invoke M n0\ show ?case proof(cases "is_val e2'") case False with bisim2 call have pc: "pc < length (compE2 e2)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 e2 ! pc = Invoke M n0" by(simp) from IH2 ins' pc False call show ?thesis by(auto dest: bisim1_bisims1.bisim1CAS2) next case True then obtain v where [simp]: "e2' = Val v" by auto from bisim2 have "pc \ length (compE2 e2)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e2)" - with bisim2 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + with bisim2 ins have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 e2)" by(cases "pc < length (compE2 e2)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1CAS3 e3 e3' xs stk loc pc e1 D F e2 v v') note IH2 = \\call1 e3' = \(a, M, vs)\; compE2 e3 ! pc = Invoke M n0\ \ ?concl e3 n e3' xs pc stk loc\ note bisim3 = \P,e3,h \ (e3', xs) \ (stk, loc, pc, None)\ note call = \call1 _ = \(a, M, vs)\\ note ins = \compE2 _ ! (length (compE2 e1) + length (compE2 e2) + pc) = Invoke M n0\ from call bisim3 have pc: "pc < length (compE2 e3)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 e3 ! pc = Invoke M n0" by(simp) from IH2 ins' pc call show ?case by(auto dest: bisim1_bisims1.bisim1CAS3) next case (bisim1Call1 obj obj' xs stk loc pc M' ps) note IH1 = \\call1 obj' = \(a, M, vs)\; compE2 obj ! pc = Invoke M n0\ \ ?concl obj n obj' xs pc stk loc\ note bisim1 = \P,obj,h \ (obj', xs) \ (stk, loc, pc, None)\ note call = \call1 (obj'\M'(ps)) = \(a, M, vs)\\ note ins = \compE2 (obj\M'(ps)) ! pc = Invoke M n0\ show ?case proof(cases "is_val obj'") case False with call bisim1 have "pc < length (compE2 obj)" by(auto intro: bisim1_call_pcD) with call False ins IH1 False show ?thesis by(auto intro: bisim1_bisims1.bisim1Call1) next case True then obtain v' where [simp]: "obj' = Val v'" by auto from bisim1 have "pc \ length (compE2 obj)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 obj)" - with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 obj)" by(cases "pc < length (compE2 obj)") auto with ins have [simp]: "ps = []" "M' = M" - by(auto simp add: nth_append split: if_split_asm)(auto simp add: neq_Nil_conv nth_append) + by(auto split: if_split_asm)(auto simp add: neq_Nil_conv) from ins call have [simp]: "vs = []" by(auto split: if_split_asm) with bisim1 have [simp]: "stk = [v']" "xs = loc" by(auto dest: bisim1_pc_length_compE2D) from bisim1Val2[of "length (compE2 (obj\M([])))" "obj\M([])" P h v loc] call ins show ?thesis by(auto simp add: is_val_iff) qed next case (bisim1CallParams ps ps' xs stk loc pc obj M' v') note IH2 = \\calls1 ps' = \(a, M, vs)\; compEs2 ps ! pc = Invoke M n0\ \ ?concls ps n ps' xs pc stk loc\ note bisim = \P,ps,h \ (ps', xs) [\] (stk, loc, pc, None)\ note call = \call1 (Val v'\M'(ps')) = \(a, M, vs)\\ note ins = \compE2 (obj\M'(ps)) ! (length (compE2 obj) + pc) = Invoke M n0\ from call show ?case proof(cases rule: call1_callE) case CallObj thus ?thesis by simp next case (CallParams v'') hence [simp]: "v'' = v'" and call': "calls1 ps' = \(a, M, vs)\" by simp_all from bisim call' have pc: "pc < length (compEs2 ps)" by(rule bisims1_calls_pcD) - with ins have ins': "compEs2 ps ! pc = Invoke M n0" by(simp add: nth_append) + with ins have ins': "compEs2 ps ! pc = Invoke M n0" by(simp) with IH2 call' ins pc have "P,ps,h \ (inline_calls (Val v) ps', xs) [\] (v # drop (Suc (length vs)) stk, loc, Suc pc, None)" and len: "Suc (length vs) \ length stk" and n0: "n0 = length vs" by auto hence "P,obj\M'(ps),h \ (Val v'\M'(inline_calls (Val v) ps'), xs) \ ((v # drop (Suc (length vs)) stk) @ [v'], loc, length (compE2 obj) + Suc pc, None)" by-(rule bisim1_bisims1.bisim1CallParams) thus ?thesis using call' len n0 by(auto simp add: is_vals_conv) next case Call hence [simp]: "v' = Addr a" "M' = M" "ps' = map Val vs" by auto from bisim have "pc \ length (compEs2 ps)" by(auto dest: bisims1_pc_length_compEs2) moreover { assume pc: "pc < length (compEs2 ps)" - with bisim ins have False by(auto dest: bisims_Val_pc_not_Invoke simp add: nth_append) } + with bisim ins have False by(auto dest: bisims_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compEs2 ps)" by(cases "pc < length (compEs2 ps)") auto from bisim have [simp]: "stk = rev vs" "xs = loc" by(auto dest: bisims1_Val_length_compEs2D) hence "P,obj\M(ps),h \ (Val v, loc) \ ([v], loc, length (compE2 (obj\M(ps))), None)" by-(rule bisim1Val2, simp) moreover from bisim have "length ps = length ps'" by(rule bisims1_lengthD) ultimately show ?thesis using ins by(auto) qed next case bisim1BlockSome1 thus ?case by simp next case bisim1BlockSome2 thus ?case by simp next case bisim1BlockSome4 thus ?case by(auto intro: bisim1_bisims1.bisim1BlockSome4) next case bisim1BlockNone thus ?case by(auto intro: bisim1_bisims1.bisim1BlockNone) next case bisim1Sync1 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Sync1) next case bisim1Sync2 thus ?case by simp next case bisim1Sync3 thus ?case by simp next case bisim1Sync4 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 bisim1_bisims1.bisim1Sync4) 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 by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Seq1) next case bisim1Seq2 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2)(fastforce dest: bisim1_bisims1.bisim1Seq2) next case bisim1Cond1 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Cond1) next case (bisim1CondThen e1 stk loc pc e e2 e' xs) thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2) (fastforce dest: bisim1_bisims1.bisim1CondThen[where e=e and ?e2.0=e2]) next case (bisim1CondElse e2 stk loc pc e e1 e' xs) thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2) (fastforce dest: bisim1_bisims1.bisim1CondElse[where e=e and ?e1.0=e1]) next case bisim1While1 thus ?case by simp next case bisim1While3 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1While3) next case bisim1While4 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2)(fastforce dest: bisim1_bisims1.bisim1While4) next case bisim1While6 thus ?case by simp next case bisim1While7 thus ?case by simp next case bisim1Throw1 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Throw1) next case bisim1Try thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Try) next case bisim1TryCatch1 thus ?case by simp next case bisim1TryCatch2 thus ?case by(fastforce dest: bisim1_bisims1.bisim1TryCatch2) next case bisims1Nil thus ?case by simp next case (bisims1List1 e e' xs stk loc pc es) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0\ \ ?concl e n e' xs pc stk loc\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \calls1 (e' # es) = \(a, M, vs)\\ note ins = \compEs2 (e # es) ! pc = Invoke M n0\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisims1List1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e)" - with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } + with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 e)" by(cases "pc < length (compE2 e)") auto with ins call have False by(cases es)(auto) thus ?thesis .. qed next case (bisims1List2 es es' xs stk loc pc e v') note IH = \\calls1 es' = \(a, M, vs)\; compEs2 es ! pc = Invoke M n0\ \ ?concls es n es' xs pc stk loc\ note call = \calls1 (Val v' # es') = \(a, M, vs)\\ note bisim = \P,es,h \ (es', xs) [\] (stk, loc, pc, None)\ note ins = \compEs2 (e # es) ! (length (compE2 e) + pc) = Invoke M n0\ from call have call': "calls1 es' = \(a, M, vs)\" by simp with bisim have pc: "pc < length (compEs2 es)" by(rule bisims1_calls_pcD) - with ins have ins': "compEs2 es ! pc = Invoke M n0" by(simp add: nth_append) + with ins have ins': "compEs2 es ! pc = Invoke M n0" by(simp) from IH call ins pc show ?case by(auto split: if_split_asm dest: bisim1_bisims1.bisims1List2) qed lemma bisim1_fv: "P,e,h \ (e', xs) \ s \ fv e' \ fv e" and bisims1_fvs: "P,es,h \ (es', xs) [\] s \ fvs es' \ fvs es" apply(induct "(e', xs)" s and "(es', xs)" s arbitrary: e' xs and es' xs rule: bisim1_bisims1.inducts) apply(auto) done lemma bisim1_syncvars: "\ P,e,h \ (e', xs) \ s; syncvars e \ \ syncvars e'" and bisims1_syncvarss: "\ P,es,h \ (es', xs) [\] s; syncvarss es \ \ syncvarss es'" apply(induct "(e', xs)" s and "(es', xs)" s arbitrary: e' xs and es' xs rule: bisim1_bisims1.inducts) apply(auto dest: bisim1_fv) done declare pcs_stack_xlift [simp] lemma bisim1_Val_\red1r: "\ P, E, h \ (e, xs) \ ([v], loc, length (compE2 E), None); n + max_vars e \ length xs; \ E n \ \ \red1r P t h (e, xs) (Val v, loc)" and bisims1_Val_\Reds1r: "\ P, Es, h \ (es, xs) [\] (rev vs, loc, length (compEs2 Es), None); n + max_varss es \ length xs; \s Es n \ \ \reds1r P t h (es, xs) (map Val vs, loc)" proof(induct E n e xs stk\"[v]" loc pc\"length (compE2 E)" xcp\"None::'addr option" and Es n es xs stk\"rev vs" loc pc\"length (compEs2 Es)" xcp\"None::'addr option" arbitrary: v and vs rule: bisim1_bisims1_inducts_split) case bisim1BlockSome2 thus ?case by(simp (no_asm_use)) next case (bisim1BlockSome4 e n e' xs loc pc V T val) from \\ {V:T=\val\; e} n\ have [simp]: "n = V" and "\ e (Suc n)" by auto note len = \n + max_vars {V:T=None; e'} \ length xs\ hence V: "V < length xs" by simp from \P,e,h \ (e', xs) \ ([v], loc, pc, None)\ have lenxs: "length xs = length loc" by(auto dest: bisim1_length_xs) note IH = \\pc = length (compE2 e); Suc n + max_vars e' \ length xs; \ e (Suc n)\ \ \red1r P t h (e', xs) (Val v, loc)\ with len \Suc (Suc pc) = length (compE2 {V:T=\val\; e})\ \\ e (Suc n)\ have "\red1r P t h (e', xs) (Val v, loc)" by(simp) hence "\red1r P t h ({V:T=None; e'}, xs) ({V:T=None; Val v}, loc)" by(rule Block_None_\red1r_xt) thus ?case using V lenxs by(auto elim!: rtranclp.rtrancl_into_rtrancl intro: Red1Block \move1BlockRed) next case (bisim1BlockNone e n e' xs loc V T) from \\ {V:T=None; e} n\ have [simp]: "n = V" and "\ e (Suc n)" by auto note len = \n + max_vars {V:T=None; e'} \ length xs\ hence V: "V < length xs" by simp from \P,e,h \ (e', xs) \ ([v], loc, length (compE2 {V:T=None; e}), None)\ have lenxs: "length xs = length loc" by(auto dest: bisim1_length_xs) note IH = \\length (compE2 {V:T=None; e}) = length (compE2 e); Suc n + max_vars e' \ length xs; \ e (Suc n) \ \ \red1r P t h (e', xs) (Val v, loc)\ with len \\ e (Suc n)\ have "\red1r P t h (e', xs) (Val v, loc)" by(simp) hence "\red1r P t h ({V:T=None; e'}, xs) ({V:T=None; Val v}, loc)" by(rule Block_None_\red1r_xt) thus ?case using V lenxs by(auto elim!: rtranclp.rtrancl_into_rtrancl intro: Red1Block \move1BlockRed) next case (bisim1TryCatch2 e2 n e' xs loc pc e C V) from \\ (try e catch(C V) e2) n\ have [simp]: "n = V" and "\ e2 (Suc n)" by auto note len = \n + max_vars {V:Class C=None; e'} \ length xs\ hence V: "V < length xs" by simp from \P,e2,h \ (e', xs) \ ([v], loc, pc, None)\ have lenxs: "length xs = length loc" by(auto dest: bisim1_length_xs) note IH = \\pc = length (compE2 e2); Suc n + max_vars e' \ length xs; \ e2 (Suc n)\ \ \red1r P t h (e', xs) (Val v, loc)\ with len \Suc (Suc (length (compE2 e) + pc)) = length (compE2 (try e catch(C V) e2))\ \\ e2 (Suc n)\ have "\red1r P t h (e', xs) (Val v, loc)" by(simp) hence "\red1r P t h ({V:Class C=None; e'}, xs) ({V:Class C=None; Val v}, loc)" by(rule Block_None_\red1r_xt) thus ?case using V lenxs by(auto elim!: rtranclp.rtrancl_into_rtrancl intro: Red1Block \move1BlockRed) next case (bisims1List1 e n e' xs loc es) note bisim = \P,e,h \ (e', xs) \ (rev vs, loc, length (compEs2 (e # es)), None)\ then have es: "es = []" and pc: "length (compEs2 (e # es)) = length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) with bisim obtain val where stk: "rev vs = [val]" and e': "is_val e' \ e' = Val val" by(auto dest: bisim1_pc_length_compE2D) with es pc bisims1List1 have "\red1r P t h (e', xs) (Val val, loc)" by simp with stk es show ?case by(auto intro: \red1r_inj_\reds1r) next case (bisims1List2 es n es' xs stk loc pc e v) from \stk @ [v] = rev vs\ obtain vs' where vs: "vs = v # vs'" by(cases vs) auto with bisims1List2 show ?case by(auto intro: \reds1r_cons_\reds1r) qed(fastforce dest: bisim1_pc_length_compE2 bisims1_pc_length_compEs2)+ lemma exec_meth_stk_split: "\ P,E,h \ (e, xs) \ (stk, loc, pc, xcp); exec_meth_d (compP2 P) (compE2 E) (stack_xlift (length STK) (compxE2 E 0 0)) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp') \ \ \stk''. stk' = stk'' @ STK \ exec_meth_d (compP2 P) (compE2 E) (compxE2 E 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc', xcp')" (is "\ _; ?exec E stk STK loc pc xcp stk' loc' pc' xcp' \ \ ?concl E stk STK loc pc xcp stk' loc' pc' xcp'") and exec_meth_stk_splits: "\ P,Es,h \ (es,xs) [\] (stk,loc,pc,xcp); exec_meth_d (compP2 P) (compEs2 Es) (stack_xlift (length STK) (compxEs2 Es 0 0)) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp') \ \ \stk''. stk' = stk'' @ STK \ exec_meth_d (compP2 P) (compEs2 Es) (compxEs2 Es 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc', xcp')" (is "\ _; ?execs Es stk STK loc pc xcp stk' loc' pc' xcp' \ \ ?concls Es stk STK loc pc xcp stk' loc' pc' xcp'") proof(induct E "n :: nat" e xs stk loc pc xcp and Es "n :: nat" es xs stk loc pc xcp arbitrary: stk' loc' pc' xcp' STK and stk' loc' pc' xcp' STK rule: bisim1_bisims1_inducts_split) case bisim1InSync thus ?case by(auto elim!: exec_meth.cases intro!: exec_meth.intros) next case bisim1Val2 thus ?case by(auto dest: exec_meth_length_compE2_stack_xliftD) next case bisim1New thus ?case by (fastforce elim: exec_meth.cases intro: exec_meth.intros split: if_split_asm cong del: image_cong_simp) next case bisim1NewThrow thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1NewArray e n e' xs stk loc pc xcp T) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (newA T\e\) stk STK loc pc xcp stk' loc' pc' xcp'\ 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 e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply simp apply (erule exec_meth.cases) apply (auto 4 4 intro: exec_meth.intros split: if_split_asm cong del: image_cong_simp) done qed next case (bisim1NewArrayThrow e n a xs stk loc pc T) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (newA T\e\) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case bisim1NewArrayFail thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1Cast e n e' xs stk loc pc xcp T) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (Cast T e) stk STK loc pc xcp stk' loc' pc' xcp'\ 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 e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply(simp) by(erule exec_meth.cases)(auto intro!: exec_meth.intros split: if_split_asm) qed next case (bisim1CastThrow e n a xs stk loc pc T) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (Cast T e) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case bisim1CastFail thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1InstanceOf e n e' xs stk loc pc xcp T) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (e instanceof T) stk STK loc pc xcp stk' loc' pc' xcp'\ 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 e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply(simp) by(erule exec_meth.cases)(auto intro!: exec_meth.intros split: if_split_asm) qed next case (bisim1InstanceOfThrow e n a xs stk loc pc T) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e instanceof T) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case bisim1Val thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case bisim1Var thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1BinOp1 e1 n e1' xs stk loc pc xcp e2 bop) note IH1 = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?exec e2 [] STK xs 0 None stk' loc' pc' xcp' \ ?concl e2 [] STK xs 0 None 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 exec = \?exec (e1 \bop\ e2) stk STK loc pc xcp stk' loc' pc' xcp'\ 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 e1 stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 e1)" by simp with exec have "pc' \ length (compE2 e1)" by(simp add: compxE2_size_convs stack_xlift_compxE2)(auto split: bop.splits elim!: exec_meth_drop_xt_pc) then obtain PC where PC: "pc' = PC + length (compE2 e1)" by -(rule_tac PC34="pc' - length (compE2 e1)" in that, simp) from pc bisim1 obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec pc have "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0))) t h (stk @ STK, loc, length (compE2 e1) + 0, xcp) ta h' (stk', loc', pc', xcp')" by-(rule exec_meth_take, auto) hence "?exec e2 [] (v # STK) loc 0 None stk' loc' (pc' - length (compE2 e1)) xcp'" using \stk = [v]\ \xcp = None\ by -(rule exec_meth_drop_xt, auto simp add: stack_xlift_compxE2 shift_compxE2) from IH2[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ [BinOpInstr bop]) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v]) (compxE2 e2 0 0))) t h ([] @ [v], loc, length (compE2 e1) + 0, None) ta h' (stk'' @ [v], loc', length (compE2 e1) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using \stk = [v]\ \xcp = None\ stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1BinOp2 e2 n e2' xs stk loc pc xcp e1 bop v1) note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc 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 exec = \?exec (e1 \bop\ e2) (stk @ [v1]) STK loc (length (compE2 e1) + pc) xcp stk' loc' pc' xcp'\ 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_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ [BinOpInstr bop]) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e2 stk (v1 # STK) loc pc xcp stk' loc' (pc' - length (compE2 e1)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e1), xcp')" by blast from exec'' have "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'' @ [v1], loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e1) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ [BinOpInstr bop]) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e1) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 e1)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(simp add: stack_xlift_compxE2 shift_compxE2) next case False with pc have pc: "pc = length (compE2 e2)" by simp with bisim2 obtain v2 where [simp]: "stk = [v2]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec pc show ?thesis by(fastforce elim: exec_meth.cases split: sum.split_asm intro!: exec_meth.intros) qed next case (bisim1BinOpThrow1 e1 n a xs stk loc pc e2 bop) note bisim1 = \P,e1,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e1 \bop\ e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 e1)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e1 @ (compE2 e2 @ [BinOpInstr bop])) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH1[OF this] show ?case by(auto) next case (bisim1BinOpThrow2 e2 n a xs stk loc pc e1 bop v1) note bisim2 = \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e1 \bop\ e2) (stk @ [v1]) STK loc (length (compE2 e1) + pc) \a\ stk' loc' pc' xcp'\ from bisim2 have pc: "pc < length (compE2 e2)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ [BinOpInstr bop]) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, \a\) ta h' (stk', loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e2 stk (v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 e1)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 e1), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v1]) (compxE2 e2 0 0)) t h (stk @ [v1], loc, pc, \a\) ta h' (stk'' @ [v1], loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e1) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule append_exec_meth_xt)(auto) hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ [BinOpInstr bop]) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e1) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule exec_meth_append) moreover from exec' have pc': "pc' \ length (compE2 e1)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case bisim1BinOpThrow thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1LAss1 e n e' xs stk loc pc xcp V) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (V := e) stk STK loc pc xcp stk' loc' pc' xcp'\ 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 e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply(simp) by(erule exec_meth.cases)(auto intro!: exec_meth.intros) qed next case (bisim1LAss2 e n xs V) thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1LAssThrow e n a xs stk loc pc V) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (V := e) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case (bisim1AAcc1 a n a' xs stk loc pc xcp i) note IH1 = \\stk' loc' pc' xcp' STK. ?exec a stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl a stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?exec i [] STK xs 0 None stk' loc' pc' xcp' \ ?concl i [] STK xs 0 None 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 exec = \?exec (a\i\) stk STK loc pc xcp stk' loc' pc' xcp'\ 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 a stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 a)" by simp with exec have "pc' \ length (compE2 a)" by(simp add: compxE2_size_convs stack_xlift_compxE2)(auto elim!: exec_meth_drop_xt_pc) then obtain PC where PC: "pc' = PC + length (compE2 a)" by -(rule_tac PC34="pc' - length (compE2 a)" in that, simp) from pc bisim1 obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec pc have "exec_meth_d (compP2 P) (compE2 a @ compE2 i) (stack_xlift (length STK) (compxE2 a 0 0 @ compxE2 i (length (compE2 a)) (Suc 0))) t h (stk @ STK, loc, length (compE2 a) + 0, xcp) ta h' (stk', loc', pc', xcp')" by-(rule exec_meth_take, auto) hence "?exec i [] (v # STK) loc 0 None stk' loc' (pc' - length (compE2 a)) xcp'" using \stk = [v]\ \xcp = None\ by -(rule exec_meth_drop_xt, auto simp add: stack_xlift_compxE2 shift_compxE2) from IH2[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "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'' @ [v], loc', length (compE2 a) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using \stk = [v]\ \xcp = None\ stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1AAcc2 i n i' xs stk loc pc xcp a v1) note IH2 = \\stk' loc' pc' xcp' STK. ?exec i stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl i stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec (a\i\) (stk @ [v1]) STK loc (length (compE2 a) + pc) xcp stk' loc' pc' xcp'\ 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_meth_d (compP2 P) ((compE2 a @ compE2 i) @ [ALoad]) (stack_xlift (length STK) (compxE2 a 0 0) @ shift (length (compE2 a)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 a) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 a @ compE2 i) (stack_xlift (length STK) (compxE2 a 0 0) @ shift (length (compE2 a)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 a) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length STK) (compxE2 i 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec i stk (v1 # STK) loc pc xcp stk' loc' (pc' - length (compE2 a)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 a), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v1]) (compxE2 i 0 0)) t h (stk @ [v1], loc, pc, xcp) ta h' (stk'' @ [v1], loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 a @ compE2 i) (compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v1]) (compxE2 i 0 0))) t h (stk @ [v1], loc, length (compE2 a) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ [ALoad]) (compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v1]) (compxE2 i 0 0))) t h (stk @ [v1], loc, length (compE2 a) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 a)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(simp add: stack_xlift_compxE2 shift_compxE2) next case False with pc have pc: "pc = length (compE2 i)" by simp with bisim2 obtain v2 where [simp]: "stk = [v2]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec pc show ?thesis by(clarsimp)(erule exec_meth.cases, auto intro!: exec_meth.intros split: if_split_asm) qed next case (bisim1AAccThrow1 A n a xs stk loc pc i) note bisim1 = \P,A,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec A stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl A stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (A\i\) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 A)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 A @ (compE2 i @ [ALoad])) (stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec A stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH1[OF this] show ?case by(auto) next case (bisim1AAccThrow2 i n a xs stk loc pc A v1) note bisim2 = \P,i,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH2 = \\stk' loc' pc' xcp' STK. ?exec i stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl i stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (A\i\) (stk @ [v1]) STK loc (length (compE2 A) + pc) \a\ stk' loc' pc' xcp'\ from bisim2 have pc: "pc < length (compE2 i)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 A @ compE2 i) @ [ALoad]) (stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 A) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 A @ compE2 i) (stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 A) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length STK) (compxE2 i 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, \a\) ta h' (stk', loc', pc' - length (compE2 A), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec i stk (v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 A)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 A), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v1]) (compxE2 i 0 0)) t h (stk @ [v1], loc, pc, \a\) ta h' (stk'' @ [v1], loc', pc' - length (compE2 A), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 A @ compE2 i) (compxE2 A 0 0 @ shift (length (compE2 A)) (stack_xlift (length [v1]) (compxE2 i 0 0))) t h (stk @ [v1], loc, length (compE2 A) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 A) + (pc' - length (compE2 A)), xcp')" by(rule append_exec_meth_xt)(auto) hence "exec_meth_d (compP2 P) ((compE2 A @ compE2 i) @ [ALoad]) (compxE2 A 0 0 @ shift (length (compE2 A)) (stack_xlift (length [v1]) (compxE2 i 0 0))) t h (stk @ [v1], loc, length (compE2 A) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 A) + (pc' - length (compE2 A)), xcp')" by(rule exec_meth_append) moreover from exec' have pc': "pc' \ length (compE2 A)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case bisim1AAccFail thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1AAss1 a n a' xs stk loc pc xcp i e) note IH1 = \\stk' loc' pc' xcp' STK. ?exec a stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl a stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?exec i [] STK xs 0 None stk' loc' pc' xcp' \ ?concl i [] STK xs 0 None 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 exec = \?exec (a\i\ := e) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim1 have pc: "pc \ length (compE2 a)" by(rule bisim1_pc_length_compE2) from exec have exec': "exec_meth_d (compP2 P) (compE2 a @ compE2 i @ compE2 e @ [AStore, Push Unit]) (stack_xlift (length STK) (compxE2 a 0 0) @ shift (length (compE2 a)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0) @ compxE2 e (length (compE2 i)) (Suc (Suc 0))))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) show ?case proof(cases "pc < length (compE2 a)") case True with exec' have "?exec a stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 a)" by simp with exec' have "pc' \ length (compE2 a)" by -(erule exec_meth_drop_xt_pc, auto) then obtain PC where PC: "pc' = PC + length (compE2 a)" by -(rule_tac PC34="pc' - length (compE2 a)" in that, simp) from pc bisim1 obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec PC pc have "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e @ [AStore, Push Unit]) (stack_xlift (length STK) (compxE2 a 0 0 @ shift (length (compE2 a)) (compxE2 i 0 (Suc 0))) @ shift (length (compE2 a @ compE2 i)) (compxE2 e 0 (length STK + Suc (Suc 0)))) t h (v # STK, loc, length (compE2 a) + 0, None) ta h' (stk', loc', length (compE2 a) + PC, xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence "exec_meth_d (compP2 P) (compE2 a @ compE2 i) (stack_xlift (length STK) (compxE2 a 0 0 @ shift (length (compE2 a)) (compxE2 i 0 (Suc 0)))) t h (v # STK, loc, length (compE2 a) + 0, None) ta h' (stk', loc', length (compE2 a) + PC, xcp')" by(rule exec_meth_take_xt) simp hence "?exec i [] (v # STK) loc 0 None stk' loc' ((length (compE2 a) + PC) - length (compE2 a)) xcp'" by -(rule exec_meth_drop_xt, auto simp add: stack_xlift_compxE2 shift_compxE2) from IH2[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "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 a @ compE2 i)) (compxE2 e 0 (Suc (Suc 0)))) t h ([] @ [v], loc, length (compE2 a) + 0, None) ta h' (stk'' @ [v], loc', length (compE2 a) + PC, xcp')" apply - apply(rule exec_meth_append_xt) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using \stk = [v]\ \xcp = None\ stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1AAss2 i n i' xs stk loc pc xcp a e v) note IH2 = \\stk' loc' pc' xcp' STK. ?exec i stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl i stk STK loc pc xcp stk' loc' pc' xcp'\ note IH3 = \\xs stk' loc' pc' xcp' STK. ?exec e [] STK xs 0 None stk' loc' pc' xcp' \ ?concl e [] STK xs 0 None 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 exec = \?exec (a\i\ := e) (stk @ [v]) STK loc (length (compE2 a) + pc) xcp stk' loc' pc' xcp'\ from bisim2 have pc: "pc \ length (compE2 i)" by(rule bisim1_pc_length_compE2) from exec have exec': "\v'. exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e @ [AStore, Push Unit]) ((compxE2 a 0 (length STK) @ shift (length (compE2 a)) (stack_xlift (length (v # STK)) (compxE2 i 0 0))) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length (v' # v # STK)) (compxE2 e 0 0))) t h (stk @ v # STK, loc, length (compE2 a) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) show ?case proof(cases "pc < length (compE2 i)") case True with exec'[of arbitrary] have exec'': "exec_meth_d (compP2 P) (compE2 a @ compE2 i) (compxE2 a 0 (length STK) @ shift (length (compE2 a)) (stack_xlift (length (v # STK)) (compxE2 i 0 0))) t h (stk @ v # STK, loc, length (compE2 a) + pc, xcp) ta h' (stk', loc', pc', xcp')" by-(erule exec_meth_take_xt, simp) hence "?exec i stk (v # STK) loc pc xcp stk' loc' (pc' - length (compE2 a)) xcp'" by(rule exec_meth_drop_xt) auto from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec''': "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 a), xcp')" by blast from exec''' have "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 a @ compE2 i)) (compxE2 e 0 (Suc (Suc 0)))) t h (stk @ [v], loc, length (compE2 a) + pc, xcp) ta h' (stk'' @ [v], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" apply - apply(rule exec_meth_append_xt) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by auto moreover from exec'' have "pc' \ length (compE2 a)" by(rule exec_meth_drop_xt_pc) auto ultimately show ?thesis using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2) next case False with pc have pc: "pc = length (compE2 i)" by simp with exec'[of arbitrary] have "pc' \ length (compE2 a @ compE2 i)" by-(erule exec_meth_drop_xt_pc, auto simp add: shift_compxE2 stack_xlift_compxE2) then obtain PC where PC: "pc' = PC + length (compE2 a) + length (compE2 i)" by -(rule_tac PC34="pc' - length (compE2 a @ compE2 i)" in that, simp) from pc bisim2 obtain v' where stk: "stk = [v']" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) from exec'[of v'] have "exec_meth_d (compP2 P) (compE2 e @ [AStore, Push Unit]) (stack_xlift (length (v' # v # STK)) (compxE2 e 0 0)) t h (v' # v # STK, loc, 0, xcp) ta h' (stk', loc', pc' - length (compE2 a @ compE2 i), xcp')" unfolding stk pc append_Cons append_Nil by -(rule exec_meth_drop_xt, simp only: add_0_right length_append, auto simp add: shift_compxE2 stack_xlift_compxE2) with PC xcp have "?exec e [] (v' # v # STK) loc 0 None stk' loc' PC xcp'" by -(rule exec_meth_take,auto) from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v' # v # STK" and "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "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 ([] @ [v', v], loc, length (compE2 a @ compE2 i) + 0, None) ta h' (stk'' @ [v', v], loc', length (compE2 a @ compE2 i) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by auto thus ?thesis using stk xcp stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1AAss3 e n e' xs stk loc pc xcp a i v1 v2) note IH3 = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim3 = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec (a\i\ := e) (stk @ [v2, v1]) STK loc (length (compE2 a) + length (compE2 i) + pc) xcp stk' loc' pc' xcp'\ 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_meth_d (compP2 P) (((compE2 a @ compE2 i) @ compE2 e) @ [AStore, Push Unit]) ((compxE2 a 0 (length STK) @ compxE2 i (length (compE2 a)) (Suc (length STK))) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 a @ compE2 i) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) hence exec': "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e) ((compxE2 a 0 (length STK) @ compxE2 i (length (compE2 a)) (Suc (length STK))) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 a @ compE2 i) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "?exec e stk (v2 # v1 # STK) loc pc xcp stk' loc' (pc' - length (compE2 a @ compE2 i)) xcp'" by(rule exec_meth_drop_xt) auto from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v2 # v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 a @ compE2 i), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e) (stack_xlift (length [v2, v1]) (compxE2 e 0 0)) t h (stk @ [v2, v1], loc, pc, xcp) ta h' (stk'' @ [v2, v1], loc', pc' - length (compE2 a @ compE2 i), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e) ((compxE2 a 0 0 @ compxE2 i (length (compE2 a)) (Suc 0)) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length [v2, v1]) (compxE2 e 0 0))) t h (stk @ [v2, v1], loc, length (compE2 a @ compE2 i) + pc, xcp) ta h' (stk'' @ [v2, v1], loc', length (compE2 a @ compE2 i) + (pc' - length (compE2 a @ compE2 i)), xcp')" by(rule append_exec_meth_xt) auto hence "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, v1]) (compxE2 e 0 0))) t h (stk @ [v2, v1], loc, length (compE2 a @ compE2 i) + pc, xcp) ta h' (stk'' @ [v2, v1], loc', length (compE2 a @ compE2 i) + (pc' - length (compE2 a @ compE2 i)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 a @ compE2 i)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(simp add: stack_xlift_compxE2 shift_compxE2) next case False with pc have pc: "pc = length (compE2 e)" by simp with bisim3 obtain v3 where [simp]: "stk = [v3]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec pc show ?thesis apply(simp) by(erule exec_meth.cases)(auto intro!: exec_meth.intros split: if_split_asm) qed next case (bisim1AAssThrow1 A n a xs stk loc pc i e) note bisim1 = \P,A,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec A stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl A stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (A\i\ := e) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 A)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 A @ (compE2 i @ compE2 e @ [AStore, Push Unit])) (stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0) @ compxE2 e (length (compE2 i)) (Suc (Suc 0))))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec A stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH1[OF this] show ?case by(auto) next case (bisim1AAssThrow2 i n a xs stk loc pc A e v1) note bisim2 = \P,i,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH2 = \\stk' loc' pc' xcp' STK. ?exec i stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl i stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (A\i\ := e) (stk @ [v1]) STK loc (length (compE2 A) + pc) \a\ stk' loc' pc' xcp'\ from bisim2 have pc: "pc < length (compE2 i)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 A @ compE2 i) @ compE2 e @ [AStore, Push Unit]) ((stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) @ (shift (length (compE2 A @ compE2 i)) (compxE2 e 0 (Suc (Suc (length STK)))))) t h (stk @ v1 # STK, loc, length (compE2 A) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence exec': "exec_meth_d (compP2 P) (compE2 A @ compE2 i) (stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 A) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take_xt)(simp add: pc) hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length STK) (compxE2 i 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, \a\) ta h' (stk', loc', pc' - length (compE2 A), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec i stk (v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 A)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 A), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v1]) (compxE2 i 0 0)) t h (stk @ [v1], loc, pc, \a\) ta h' (stk'' @ [v1], loc', pc' - length (compE2 A), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 A @ compE2 i) (compxE2 A 0 0 @ shift (length (compE2 A)) (stack_xlift (length [v1]) (compxE2 i 0 0))) t h (stk @ [v1], loc, length (compE2 A) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 A) + (pc' - length (compE2 A)), xcp')" by(rule append_exec_meth_xt)(auto) hence "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 [v1]) (compxE2 i 0 0))) @ (shift (length (compE2 A @ compE2 i)) (compxE2 e 0 (Suc (Suc 0))))) t h (stk @ [v1], loc, length (compE2 A) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 A) + (pc' - length (compE2 A)), xcp')" by(rule exec_meth_append_xt) moreover from exec' have pc': "pc' \ length (compE2 A)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case (bisim1AAssThrow3 e n a xs stk loc pc A i v2 v1) note bisim3 = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH3 = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (A\i\ := e) (stk @ [v2, v1]) STK loc (length (compE2 A) + length (compE2 i) + pc) \a\ stk' loc' pc' xcp'\ from bisim3 have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (((compE2 A @ compE2 i) @ compE2 e) @ [AStore, Push Unit]) ((stack_xlift (length STK) (compxE2 A 0 0 @ compxE2 i (length (compE2 A)) (Suc 0))) @ shift (length (compE2 A @ compE2 i)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 A @ compE2 i) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence exec': "exec_meth_d (compP2 P) ((compE2 A @ compE2 i) @ compE2 e) ((stack_xlift (length STK) (compxE2 A 0 0 @ compxE2 i (length (compE2 A)) (Suc 0))) @ shift (length (compE2 A @ compE2 i)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 A @ compE2 i) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "?exec e stk (v2 # v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 A @ compE2 i)) xcp'" by(rule exec_meth_drop_xt) auto from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v2 # v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 A @ compE2 i), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e) (stack_xlift (length [v2, v1]) (compxE2 e 0 0)) t h (stk @ [v2, v1], loc, pc, \a\) ta h' (stk'' @ [v2, v1], loc', pc' - length (compE2 A @ compE2 i), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) ((compE2 A @ compE2 i) @ compE2 e) ((compxE2 A 0 0 @ compxE2 i (length (compE2 A)) (Suc 0)) @ shift (length (compE2 A @ compE2 i)) (stack_xlift (length [v2, v1]) (compxE2 e 0 0))) t h (stk @ [v2, v1], loc, length (compE2 A @ compE2 i) + pc, \a\) ta h' (stk'' @ [v2, v1], loc', length (compE2 A @ compE2 i) + (pc' - length (compE2 A @ compE2 i)), xcp')" by(rule append_exec_meth_xt)(auto) hence "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, v1]) (compxE2 e 0 0))) t h (stk @ [v2, v1], loc, length (compE2 A @ compE2 i) + pc, \a\) ta h' (stk'' @ [v2, v1], loc', length (compE2 A @ compE2 i) + (pc' - length (compE2 A @ compE2 i)), xcp')" by(rule exec_meth_append) moreover from exec' have pc': "pc' \ length (compE2 A @ compE2 i)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case bisim1AAssFail thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case bisim1AAss4 thus ?case by -(erule exec_meth.cases, auto intro!: exec_meth.exec_instr) next case (bisim1ALength a n a' xs stk loc pc xcp) note bisim = \P,a,h \ (a', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec a stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl a stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (a\length) stk STK loc pc xcp stk' loc' pc' xcp'\ 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 a stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 a)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply(simp) by(erule exec_meth.cases)(auto intro!: exec_meth.intros split: if_split_asm) qed next case (bisim1ALengthThrow e n a xs stk loc pc) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e\length) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case bisim1ALengthNull thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1FAcc e n e' xs stk loc pc xcp F D) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (e\F{D}) stk STK loc pc xcp stk' loc' pc' xcp'\ 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 e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply(simp) by(erule exec_meth.cases)(fastforce intro!: exec_meth.intros simp add: is_Ref_def split: if_split_asm)+ qed next case (bisim1FAccThrow e n a xs stk loc pc F D) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e\F{D}) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case bisim1FAccNull thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1FAss1 e n e' xs stk loc pc xcp e2 F D) note IH1 = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?exec e2 [] STK xs 0 None stk' loc' pc' xcp' \ ?concl e2 [] STK xs 0 None 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 exec = \?exec (e\F{D} := e2) stk STK loc pc xcp stk' loc' pc' xcp'\ 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 e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 e)" by simp with exec have "pc' \ length (compE2 e)" by(simp add: compxE2_size_convs stack_xlift_compxE2)(auto elim!: exec_meth_drop_xt_pc) then obtain PC where PC: "pc' = PC + length (compE2 e)" by -(rule_tac PC34="pc' - length (compE2 e)" in that, simp) from pc bisim1 obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec pc have "exec_meth_d (compP2 P) (compE2 e @ compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0 @ compxE2 e2 (length (compE2 e)) (Suc 0))) t h (stk @ STK, loc, length (compE2 e) + 0, xcp) ta h' (stk', loc', pc', xcp')" by-(rule exec_meth_take, auto) hence "?exec e2 [] (v # STK) loc 0 None stk' loc' (pc' - length (compE2 e)) xcp'" using \stk = [v]\ \xcp = None\ by -(rule exec_meth_drop_xt, auto simp add: stack_xlift_compxE2 shift_compxE2) from IH2[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "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'' @ [v], loc', length (compE2 e) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using \stk = [v]\ \xcp = None\ stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1FAss2 e2 n e' xs stk loc pc xcp e F D v1) note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim2 = \P,e2,h \ (e', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec (e\F{D} := e2) (stk @ [v1]) STK loc (length (compE2 e) + pc) xcp stk' loc' pc' xcp'\ 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_meth_d (compP2 P) ((compE2 e @ compE2 e2) @ [Putfield F D, Push Unit]) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 e @ compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e2 stk (v1 # STK) loc pc xcp stk' loc' (pc' - length (compE2 e)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" by blast from exec'' have "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'' @ [v1], loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e @ compE2 e2) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule append_exec_meth_xt) auto hence "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 [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(simp add: stack_xlift_compxE2 shift_compxE2) next case False with pc have pc: "pc = length (compE2 e2)" by simp with bisim2 obtain v2 where [simp]: "stk = [v2]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec pc show ?thesis apply(simp) by(erule exec_meth.cases)(fastforce intro!: exec_meth.intros split: if_split_asm)+ qed next case (bisim1FAssThrow1 e n a xs stk loc pc e2 F D) note bisim1 = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e\F{D} := e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e @ (compE2 e2 @ [Putfield F D, Push Unit])) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH1[OF this] show ?case by(auto) next case (bisim1FAssThrow2 e2 n a xs stk loc pc e F D v1) note bisim2 = \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e\F{D} := e2) (stk @ [v1]) STK loc (length (compE2 e) + pc) \a\ stk' loc' pc' xcp'\ from bisim2 have pc: "pc < length (compE2 e2)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 e @ compE2 e2) @ [Putfield F D, Push Unit]) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 e @ compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, \a\) ta h' (stk', loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e2 stk (v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 e)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v1]) (compxE2 e2 0 0)) t h (stk @ [v1], loc, pc, \a\) ta h' (stk'' @ [v1], loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e @ compE2 e2) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule append_exec_meth_xt)(auto) hence "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 [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule exec_meth_append) moreover from exec' have pc': "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case bisim1FAssNull thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case bisim1FAss3 thus ?case by -(erule exec_meth.cases, auto intro!: exec_meth.exec_instr) next case (bisim1CAS1 e1 n e1' xs stk loc pc xcp e2 e3 D F) note IH1 = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?exec e2 [] STK xs 0 None stk' loc' pc' xcp' \ ?concl e2 [] STK xs 0 None 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 exec = \?exec _ stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim1 have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) from exec have exec': "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2 @ compE2 e3 @ [CAS F D]) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0) @ compxE2 e3 (length (compE2 e2)) (Suc (Suc 0))))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) show ?case proof(cases "pc < length (compE2 e1)") case True with exec' have "?exec e1 stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 e1)" by simp with exec' have "pc' \ length (compE2 e1)" by -(erule exec_meth_drop_xt_pc, auto) then obtain PC where PC: "pc' = PC + length (compE2 e1)" by -(rule_tac PC34="pc' - length (compE2 e1)" in that, simp) from pc bisim1 obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec PC pc have "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3 @ [CAS F D]) (stack_xlift (length STK) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (compxE2 e2 0 (Suc 0))) @ shift (length (compE2 e1 @ compE2 e2)) (compxE2 e3 0 (length STK + Suc (Suc 0)))) t h (v # STK, loc, length (compE2 e1) + 0, None) ta h' (stk', loc', length (compE2 e1) + PC, xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (compxE2 e2 0 (Suc 0)))) t h (v # STK, loc, length (compE2 e1) + 0, None) ta h' (stk', loc', length (compE2 e1) + PC, xcp')" by(rule exec_meth_take_xt) simp hence "?exec e2 [] (v # STK) loc 0 None stk' loc' ((length (compE2 e1) + PC) - length (compE2 e1)) xcp'" by -(rule exec_meth_drop_xt, auto simp add: stack_xlift_compxE2 shift_compxE2) from IH2[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ (compE2 e3 @ [CAS F D])) ((compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v]) (compxE2 e2 0 0))) @ shift (length (compE2 e1 @ compE2 e2)) (compxE2 e3 0 (Suc (Suc 0)))) t h ([] @ [v], loc, length (compE2 e1) + 0, None) ta h' (stk'' @ [v], loc', length (compE2 e1) + PC, xcp')" apply - apply(rule exec_meth_append_xt) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using \stk = [v]\ \xcp = None\ stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1CAS2 e2 n e2' xs stk loc pc xcp e1 e3 D F v) note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note IH3 = \\xs stk' loc' pc' xcp' STK. ?exec e3 [] STK xs 0 None stk' loc' pc' xcp' \ ?concl e3 [] STK xs 0 None stk' loc' pc' xcp'\ note bisim2 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, xcp)\ note bisim3 = \P,e3,h \ (e3, loc) \ ([], loc, 0, None)\ note exec = \?exec _ (stk @ [v]) STK loc (length (compE2 e1) + pc) xcp stk' loc' pc' xcp'\ from bisim2 have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) from exec have exec': "\v'. exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3 @ [CAS F D]) ((compxE2 e1 0 (length STK) @ shift (length (compE2 e1)) (stack_xlift (length (v # STK)) (compxE2 e2 0 0))) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length (v' # v # STK)) (compxE2 e3 0 0))) t h (stk @ v # STK, loc, length (compE2 e1) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) show ?case proof(cases "pc < length (compE2 e2)") case True with exec'[of undefined] have exec'': "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (compxE2 e1 0 (length STK) @ shift (length (compE2 e1)) (stack_xlift (length (v # STK)) (compxE2 e2 0 0))) t h (stk @ v # STK, loc, length (compE2 e1) + pc, xcp) ta h' (stk', loc', pc', xcp')" by-(erule exec_meth_take_xt, simp) hence "?exec e2 stk (v # STK) loc pc xcp stk' loc' (pc' - length (compE2 e1)) xcp'" by(rule exec_meth_drop_xt) auto from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec''': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e1), xcp')" by blast from exec''' have "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3 @ [CAS F D]) ((compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v]) (compxE2 e2 0 0))) @ shift (length (compE2 e1 @ compE2 e2)) (compxE2 e3 0 (Suc (Suc 0)))) t h (stk @ [v], loc, length (compE2 e1) + pc, xcp) ta h' (stk'' @ [v], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" apply - apply(rule exec_meth_append_xt) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by auto moreover from exec'' have "pc' \ length (compE2 e1)" by(rule exec_meth_drop_xt_pc) auto ultimately show ?thesis using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2) next case False with pc have pc: "pc = length (compE2 e2)" by simp with exec'[of undefined] have "pc' \ length (compE2 e1 @ compE2 e2)" by-(erule exec_meth_drop_xt_pc, auto simp add: shift_compxE2 stack_xlift_compxE2) then obtain PC where PC: "pc' = PC + length (compE2 e1) + length (compE2 e2)" by -(rule_tac PC34="pc' - length (compE2 e1 @ compE2 e2)" in that, simp) from pc bisim2 obtain v' where stk: "stk = [v']" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) from exec'[of v'] have "exec_meth_d (compP2 P) (compE2 e3 @ [CAS F D]) (stack_xlift (length (v' # v # STK)) (compxE2 e3 0 0)) t h (v' # v # STK, loc, 0, xcp) ta h' (stk', loc', pc' - length (compE2 e1 @ compE2 e2), xcp')" unfolding stk pc append_Cons append_Nil by -(rule exec_meth_drop_xt, simp only: add_0_right length_append, auto simp add: shift_compxE2 stack_xlift_compxE2) with PC xcp have "?exec e3 [] (v' # v # STK) loc 0 None stk' loc' PC xcp'" by -(rule exec_meth_take,auto) from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v' # v # STK" and "exec_meth_d (compP2 P) (compE2 e3) (compxE2 e3 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "exec_meth_d (compP2 P) (((compE2 e1 @ compE2 e2) @ compE2 e3) @ [CAS F D]) ((compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0)) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length [v', v]) (compxE2 e3 0 0))) t h ([] @ [v', v], loc, length (compE2 e1 @ compE2 e2) + 0, None) ta h' (stk'' @ [v', v], loc', length (compE2 e1 @ compE2 e2) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by auto thus ?thesis using stk xcp stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1CAS3 e3 n e3' xs stk loc pc xcp e1 e2 D F v1 v2) note IH3 = \\stk' loc' pc' xcp' STK. ?exec e3 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e3 stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim3 = \P,e3,h \ (e3', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec _ (stk @ [v2, v1]) STK loc (length (compE2 e1) + length (compE2 e2) + pc) xcp stk' loc' pc' xcp'\ from bisim3 have pc: "pc \ length (compE2 e3)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e3)") case True from exec have "exec_meth_d (compP2 P) (((compE2 e1 @ compE2 e2) @ compE2 e3) @ [CAS F D]) ((compxE2 e1 0 (length STK) @ compxE2 e2 (length (compE2 e1)) (Suc (length STK))) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e3 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 e1 @ compE2 e2) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) hence exec': "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3) ((compxE2 e1 0 (length STK) @ compxE2 e2 (length (compE2 e1)) (Suc (length STK))) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e3 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 e1 @ compE2 e2) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "?exec e3 stk (v2 # v1 # STK) loc pc xcp stk' loc' (pc' - length (compE2 e1 @ compE2 e2)) xcp'" by(rule exec_meth_drop_xt) auto from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v2 # v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e3) (compxE2 e3 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e1 @ compE2 e2), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e3) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0)) t h (stk @ [v2, v1], loc, pc, xcp) ta h' (stk'' @ [v2, v1], loc', pc' - length (compE2 e1 @ compE2 e2), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3) ((compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0)) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0))) t h (stk @ [v2, v1], loc, length (compE2 e1 @ compE2 e2) + pc, xcp) ta h' (stk'' @ [v2, v1], loc', length (compE2 e1 @ compE2 e2) + (pc' - length (compE2 e1 @ compE2 e2)), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) (((compE2 e1 @ compE2 e2) @ compE2 e3) @ [CAS F D]) ((compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0)) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0))) t h (stk @ [v2, v1], loc, length (compE2 e1 @ compE2 e2) + pc, xcp) ta h' (stk'' @ [v2, v1], loc', length (compE2 e1 @ compE2 e2) + (pc' - length (compE2 e1 @ compE2 e2)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 e1 @ compE2 e2)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(simp add: stack_xlift_compxE2 shift_compxE2) next case False with pc have pc: "pc = length (compE2 e3)" by simp with bisim3 obtain v3 where [simp]: "stk = [v3]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec pc show ?thesis apply(simp) by(erule exec_meth.cases)(fastforce intro!: exec_meth.intros split: if_split_asm)+ qed next case (bisim1CASThrow1 e1 n a xs stk loc pc e2 e3 D F) note bisim1 = \P,e1,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec _ stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 e1)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e1 @ (compE2 e2 @ compE2 e3 @ [CAS F D])) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0) @ compxE2 e3 (length (compE2 e2)) (Suc (Suc 0))))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH1[OF this] show ?case by(auto) next case (bisim1CASThrow2 e2 n a xs stk loc pc e1 e3 D F v1) note bisim2 = \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec _ (stk @ [v1]) STK loc (length (compE2 e1) + pc) \a\ stk' loc' pc' xcp'\ from bisim2 have pc: "pc < length (compE2 e2)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3 @ [CAS F D]) ((stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) @ (shift (length (compE2 e1 @ compE2 e2)) (compxE2 e3 0 (Suc (Suc (length STK)))))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence exec': "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take_xt)(simp add: pc) hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, \a\) ta h' (stk', loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e2 stk (v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 e1)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 e1), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v1]) (compxE2 e2 0 0)) t h (stk @ [v1], loc, pc, \a\) ta h' (stk'' @ [v1], loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e1) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule append_exec_meth_xt)(auto) hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3 @ [CAS F D]) ((compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) @ (shift (length (compE2 e1 @ compE2 e2)) (compxE2 e3 0 (Suc (Suc 0))))) t h (stk @ [v1], loc, length (compE2 e1) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule exec_meth_append_xt) moreover from exec' have pc': "pc' \ length (compE2 e1)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case (bisim1CASThrow3 e3 n a xs stk loc pc e1 e2 D F v2 v1) note bisim3 = \P,e3,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH3 = \\stk' loc' pc' xcp' STK. ?exec e3 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e3 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec _ (stk @ [v2, v1]) STK loc (length (compE2 e1) + length (compE2 e2) + pc) \a\ stk' loc' pc' xcp'\ from bisim3 have pc: "pc < length (compE2 e3)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (((compE2 e1 @ compE2 e2) @ compE2 e3) @ [CAS F D]) ((stack_xlift (length STK) (compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0))) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e3 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 e1 @ compE2 e2) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence exec': "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3) ((stack_xlift (length STK) (compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0))) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e3 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 e1 @ compE2 e2) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "?exec e3 stk (v2 # v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 e1 @ compE2 e2)) xcp'" by(rule exec_meth_drop_xt) auto from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v2 # v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e3) (compxE2 e3 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 e1 @ compE2 e2), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e3) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0)) t h (stk @ [v2, v1], loc, pc, \a\) ta h' (stk'' @ [v2, v1], loc', pc' - length (compE2 e1 @ compE2 e2), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3) ((compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0)) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0))) t h (stk @ [v2, v1], loc, length (compE2 e1 @ compE2 e2) + pc, \a\) ta h' (stk'' @ [v2, v1], loc', length (compE2 e1 @ compE2 e2) + (pc' - length (compE2 e1 @ compE2 e2)), xcp')" by(rule append_exec_meth_xt)(auto) hence "exec_meth_d (compP2 P) (((compE2 e1 @ compE2 e2) @ compE2 e3) @ [CAS F D]) ((compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0)) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0))) t h (stk @ [v2, v1], loc, length (compE2 e1 @ compE2 e2) + pc, \a\) ta h' (stk'' @ [v2, v1], loc', length (compE2 e1 @ compE2 e2) + (pc' - length (compE2 e1 @ compE2 e2)), xcp')" by(rule exec_meth_append) moreover from exec' have pc': "pc' \ length (compE2 e1 @ compE2 e2)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case bisim1CASFail thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1Call1 obj n obj' xs stk loc pc xcp ps M') note bisimObj = \P,obj,h \ (obj', xs) \ (stk, loc, pc, xcp)\ note IHobj = \\stk' loc' pc' xcp' STK. ?exec obj stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl obj stk STK loc pc xcp stk' loc' pc' xcp'\ note IHparams = \\xs stk' loc' pc' xcp' STK. ?execs ps [] STK xs 0 None stk' loc' pc' xcp' \ ?concls ps [] STK xs 0 None stk' loc' pc' xcp'\ note exec = \?exec (obj\M'(ps)) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisimObj have pc: "pc \ length (compE2 obj)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 obj)") case True from exec have "?exec obj stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxEs2_size_convs)(erule exec_meth_take_xt[OF _ True]) from IHobj[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 obj)" by simp with exec have "pc' \ length (compE2 obj)" by(simp add: compxEs2_size_convs stack_xlift_compxE2)(auto elim!: exec_meth_drop_xt_pc) then obtain PC where PC: "pc' = PC + length (compE2 obj)" by -(rule_tac PC34="pc' - length (compE2 obj)" in that, simp) from pc bisimObj obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) show ?thesis proof(cases ps) case Cons with exec pc have "exec_meth_d (compP2 P) (compE2 obj @ compEs2 ps) (stack_xlift (length STK) (compxE2 obj 0 0 @ compxEs2 ps (length (compE2 obj)) (Suc 0))) t h (stk @ STK, loc, length (compE2 obj) + 0, xcp) ta h' (stk', loc', pc', xcp')" by -(rule exec_meth_take, auto) hence "?execs ps [] (v # STK) loc 0 None stk' loc' (pc' - length (compE2 obj)) xcp'" apply - apply(rule exec_meth_drop_xt) apply(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) apply(auto simp add: stack_xlift_compxE2) done from IHparams[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec': "exec_meth_d (compP2 P) (compEs2 ps) (compxEs2 ps 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto from exec' have "exec_meth_d (compP2 P) ((compE2 obj @ compEs2 ps) @ [Invoke M' (length ps)]) (compxE2 obj 0 0 @ shift (length (compE2 obj)) (stack_xlift (length [v]) (compxEs2 ps 0 0))) t h ([] @ [v], loc, length (compE2 obj) + 0, None) ta h' (stk'' @ [v], loc', length (compE2 obj) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using stk' PC by(clarsimp simp add: shift_compxEs2 stack_xlift_compxEs2 ac_simps) next case Nil with exec pc show ?thesis apply(auto elim!: exec_meth.cases intro!: exec_meth.intros simp add: split_beta split: if_split_asm) apply(auto split: extCallRet.split_asm intro!: exec_meth.intros) apply(force intro!: exI) apply(force intro!: exI) apply(force intro!: exI) done qed qed next case (bisim1CallParams ps n ps' xs stk loc pc xcp obj M' v) note bisimParam = \P,ps,h \ (ps',xs) [\] (stk,loc,pc,xcp)\ note IHparam = \\stk' loc' pc' xcp' STK. ?execs ps stk STK loc pc xcp stk' loc' pc' xcp' \ ?concls ps stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (obj\M'(ps)) (stk @ [v]) STK loc (length (compE2 obj) + pc) xcp stk' loc' pc' xcp'\ show ?case proof(cases ps) case Nil with bisimParam have "pc = 0" "xcp = None" by(auto elim: bisims1.cases) with exec Nil show ?thesis apply(auto elim!: exec_meth.cases intro!: exec_meth.intros simp add: split_beta extRet2JVM_def split: if_split_asm) apply(auto split: extCallRet.split_asm simp add: neq_Nil_conv) apply(force intro!: exec_meth.intros)+ done next case Cons from bisimParam have pc: "pc \ length (compEs2 ps)" by(rule bisims1_pc_length_compEs2) show ?thesis proof(cases "pc < length (compEs2 ps)") case True from exec have "exec_meth_d (compP2 P) ((compE2 obj @ compEs2 ps) @ [Invoke M' (length ps)]) (stack_xlift (length STK) (compxE2 obj 0 0) @ shift (length (compE2 obj)) (stack_xlift (length (v # STK)) (compxEs2 ps 0 0))) t h (stk @ v # STK, loc, length (compE2 obj) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) hence exec': "exec_meth_d (compP2 P) (compE2 obj @ compEs2 ps) (stack_xlift (length STK) (compxE2 obj 0 0) @ shift (length (compE2 obj)) (stack_xlift (length (v # STK)) (compxEs2 ps 0 0))) t h (stk @ v # STK, loc, length (compE2 obj) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "?execs ps stk (v # STK) loc pc xcp stk' loc' (pc' - length (compE2 obj)) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IHparam[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec'': "exec_meth_d (compP2 P) (compEs2 ps) (compxEs2 ps 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 obj), xcp')" by blast from exec'' have "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'' @ [v], loc', pc' - length (compE2 obj), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 obj @ compEs2 ps) (compxE2 obj 0 0 @ shift (length (compE2 obj)) (stack_xlift (length [v]) (compxEs2 ps 0 0))) t h (stk @ [v], loc, length (compE2 obj) + pc, xcp) ta h' (stk'' @ [v], loc', length (compE2 obj) + (pc' - length (compE2 obj)), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) ((compE2 obj @ compEs2 ps) @ [Invoke M' (length ps)]) (compxE2 obj 0 0 @ shift (length (compE2 obj)) (stack_xlift (length [v]) (compxEs2 ps 0 0))) t h (stk @ [v], loc, length (compE2 obj) + pc, xcp) ta h' (stk'' @ [v], loc', length (compE2 obj) + (pc' - length (compE2 obj)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 obj)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(auto simp add: shift_compxEs2 stack_xlift_compxEs2) next case False with pc have pc: "pc = length (compEs2 ps)" by simp with bisimParam obtain vs where "stk = vs" "length vs = length ps" "xcp = None" by(auto dest: bisims1_pc_length_compEs2D) with exec pc Cons show ?thesis apply(auto elim!: exec_meth.cases intro!: exec_meth.intros simp add: split_beta extRet2JVM_def split: if_split_asm) apply(auto simp add: neq_Nil_conv split: extCallRet.split_asm) apply(force intro!: exec_meth.intros)+ done qed qed next case (bisim1CallThrowObj obj n a xs stk loc pc ps M') note bisim = \P,obj,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec obj stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl obj stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (obj\M'(ps)) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have "pc < length (compE2 obj)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) with exec have "?exec obj stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs)(erule exec_meth_take_xt) from IH[OF this] show ?case by auto next case (bisim1CallThrowParams ps n vs a ps' xs stk loc pc obj M' v) note bisim = \P,ps,h \ (map Val vs @ Throw a # ps',xs) [\] (stk,loc,pc,\a\)\ note IH = \\stk' loc' pc' xcp' STK. ?execs ps stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concls ps stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (obj\M'(ps)) (stk @ [v]) STK loc (length (compE2 obj) + pc) \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compEs2 ps)" "loc = xs" by(auto dest: bisims1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 obj @ compEs2 ps) @ [Invoke M' (length ps)]) (stack_xlift (length STK) (compxE2 obj 0 0) @ shift (length (compE2 obj)) (stack_xlift (length (v # STK)) (compxEs2 ps 0 0))) t h (stk @ v # STK, loc, length (compE2 obj) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) hence exec': "exec_meth_d (compP2 P) (compE2 obj @ compEs2 ps) (stack_xlift (length STK) (compxE2 obj 0 0) @ shift (length (compE2 obj)) (stack_xlift (length (v # STK)) (compxEs2 ps 0 0))) t h (stk @ v # STK, loc, length (compE2 obj) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "?execs ps stk (v # STK) loc pc \a\ stk' loc' (pc' - length (compE2 obj)) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec'': "exec_meth_d (compP2 P) (compEs2 ps) (compxEs2 ps 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 obj), xcp')" by auto from exec'' have "exec_meth_d (compP2 P) ((compE2 obj @ compEs2 ps) @ [Invoke M' (length ps)]) (compxE2 obj 0 0 @ shift (length (compE2 obj)) (stack_xlift (length [v]) (compxEs2 ps 0 0))) t h (stk @ [v], loc, length (compE2 obj) + pc, \a\) ta h' (stk'' @ [v], loc', length (compE2 obj) + (pc' - length (compE2 obj)), xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) apply auto done moreover from exec' have "pc' \ length (compE2 obj)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) next case bisim1CallThrow thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case bisim1BlockSome1 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case bisim1BlockSome2 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1BlockSome4 e n e' xs stk loc pc xcp V T v) note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec {V:T=\v\; e} stk STK loc (Suc (Suc pc)) xcp stk' loc' pc' xcp'\ hence exec': "exec_meth_d (compP2 P) ([Push v, Store V] @ compE2 e) (shift (length [Push v, Store V]) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, length [Push v, Store V] + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e stk STK loc pc xcp stk' loc' (pc' - length [Push v, Store V]) xcp'" by(rule exec_meth_drop) auto from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length [Push v, Store V], xcp')" by auto from exec'' have "exec_meth_d (compP2 P) ([Push v, Store V] @ compE2 e) (shift (length [Push v, Store V]) (compxE2 e 0 0)) t h (stk, loc, length [Push v, Store V] + pc, xcp) ta h' (stk'', loc', length [Push v, Store V] + (pc' - length [Push v, Store V]), xcp')" by(rule append_exec_meth) auto moreover from exec' have "pc' \ length [Push v, Store V]" by(rule exec_meth_drop_pc)(auto simp add: stack_xlift_compxE2) hence "Suc (Suc (pc' - Suc (Suc 0))) = pc'" by(simp) ultimately show ?case using stk' by(auto simp add: compxE2_size_convs) next case (bisim1BlockThrowSome e n a xs stk loc pc V T v) note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec {V:T=\v\; e} stk STK loc (Suc (Suc pc)) \a\ stk' loc' pc' xcp'\ hence exec': "exec_meth_d (compP2 P) ([Push v, Store V] @ compE2 e) (shift (length [Push v, Store V]) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, length [Push v, Store V] + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e stk STK loc pc \a\ stk' loc' (pc' - length [Push v, Store V]) xcp'" by(rule exec_meth_drop) auto from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length [Push v, Store V], xcp')" by auto from exec'' have "exec_meth_d (compP2 P) ([Push v, Store V] @ compE2 e) (shift (length [Push v, Store V]) (compxE2 e 0 0)) t h (stk, loc, length [Push v, Store V] + pc, \a\) ta h' (stk'', loc', length [Push v, Store V] + (pc' - length [Push v, Store V]), xcp')" by(rule append_exec_meth) auto moreover from exec' have "pc' \ length [Push v, Store V]" by(rule exec_meth_drop_pc)(auto simp add: stack_xlift_compxE2) hence "Suc (Suc (pc' - Suc (Suc 0))) = pc'" by(simp) ultimately show ?case using stk' by(auto simp add: compxE2_size_convs) next case bisim1BlockNone thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case bisim1BlockThrowNone thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1Sync1 e1 n e1' xs stk loc pc xcp e2 V) note bisim = \P,e1,h \ (e1', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e1)") case True from exec have "exec_meth_d (compP2 P) (compE2 e1 @ (Dup # Store V # MEnter # compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc])) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 3 0) @ stack_xlift (length STK) [(3, 3 + length (compE2 e2), None, 6 + length (compE2 e2), 0)])) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence "?exec e1 stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule True) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec': "exec_meth_d (compP2 P) (compE2 e1) (compxE2 e1 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc', xcp')" by blast from exec' have "exec_meth_d (compP2 P) (compE2 e1 @ (Dup # Store V # MEnter # compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc])) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (compxE2 e2 3 0 @ [(3, 3 + length (compE2 e2), None, 6 + length (compE2 e2), 0)])) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc', xcp')" by(rule exec_meth_append_xt) thus ?thesis using stk' by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) next case False with pc have [simp]: "pc = length (compE2 e1)" by simp with bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) thus ?thesis using exec by(auto elim!: exec_meth.cases intro!: exec_meth.intros) qed next case bisim1Sync2 thus ?case by(fastforce elim!: exec_meth.cases intro!: exec_meth.intros) next case bisim1Sync3 thus ?case by(fastforce elim!: exec_meth.cases intro!: exec_meth.intros split: if_split_asm) next case (bisim1Sync4 e2 n e2' xs stk loc pc xcp e1 V a) note IH = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e2,h \ (e2', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) stk STK loc (Suc (Suc (Suc (length (compE2 e1) + pc)))) xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e2)") case True let ?pre = "compE2 e1 @ [Dup, Store V, MEnter]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e2 0 0) @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), length STK)])) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: stack_xlift_compxE2 shift_compxE2 eval_nat_numeral ac_simps) hence exec'': "exec_meth_d (compP2 P) (compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (stack_xlift (length STK) (compxE2 e2 0 0) @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), length STK)]) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc' - length ?pre, xcp')" by(rule exec_meth_drop_xt[where n=1])(auto simp add: stack_xlift_compxE2) from exec' have pc': "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc[where n'=1])(auto simp add: stack_xlift_compxE2) hence pc'': "(Suc (Suc (Suc (pc' - Suc (Suc (Suc 0)))))) = pc'" by simp show ?thesis proof(cases xcp) case None from exec'' None True have "?exec e2 stk STK loc pc xcp stk' loc' (pc' - length ?pre) xcp'" apply - apply (erule exec_meth.cases) apply (cases "compE2 e2 ! pc") apply (fastforce simp add: is_Ref_def intro: exec_meth.intros split: if_split_asm cong del: image_cong_simp)+ done from IH[OF this] obtain stk'' where stk: "stk' = stk'' @ STK" and exec''': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by blast from exec''' have "exec_meth_d (compP2 P) (compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (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_append_xt) hence "exec_meth_d (compP2 P) (?pre @ compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (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', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt[where n=1]) auto thus ?thesis using stk pc' pc'' by(simp add: eval_nat_numeral shift_compxE2 ac_simps) next case (Some a) with exec'' have [simp]: "h' = h" "xcp' = None" "loc' = loc" "ta = \" by(auto elim!: exec_meth.cases simp add: match_ex_table_append split: if_split_asm dest!: match_ex_table_stack_xliftD) show ?thesis proof(cases "match_ex_table (compP2 P) (cname_of h a) pc (compxE2 e2 0 0)") case None with Some exec'' True have [simp]: "stk' = Addr a # STK" and pc': "pc' = length (compE2 e1) + length (compE2 e2) + 6" by(auto elim!: exec_meth.cases simp add: match_ex_table_append split: if_split_asm dest!: match_ex_table_stack_xliftD) with exec'' Some None have "exec_meth_d (compP2 P) (compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (compxE2 e2 0 0 @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), 0)]) t h (stk, loc, pc, \a\) \ h (Addr a # drop (length stk - 0) stk, loc, pc' - length ?pre, None)" by -(rule exec_catch, auto elim!: exec_meth.cases simp add: match_ex_table_append matches_ex_entry_def split: if_split_asm dest!: match_ex_table_stack_xliftD) hence "exec_meth_d (compP2 P) (?pre @ compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (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, \a\) \ h (Addr a # drop (length stk - 0) stk, loc, length ?pre + (pc' - length ?pre), None)" by(rule append_exec_meth_xt[where n=1]) auto with pc' Some show ?thesis by(simp add: eval_nat_numeral shift_compxE2 ac_simps) next case (Some pcd) with \xcp = \a\\ exec'' True have "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, \a\) \ h (Addr a # drop (length stk - snd pcd) stk, loc, pc' - length ?pre, None)" apply - apply(rule exec_catch) apply(auto elim!: exec_meth.cases simp add: match_ex_table_append split: if_split_asm dest!: match_ex_table_stack_xliftD) done hence "exec_meth_d (compP2 P) (compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (compxE2 e2 0 0 @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), 0)]) t h (stk, loc, pc, \a\) \ h (Addr a # drop (length stk - snd pcd) stk, loc, pc' - length ?pre, None)" by(rule exec_meth_append_xt) hence "exec_meth_d (compP2 P) (?pre @ compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (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, \a\) \ h (Addr a # drop (length stk - snd pcd) stk, loc, length ?pre + (pc' - length ?pre), None)" by(rule append_exec_meth_xt[where n=1])(auto) moreover from Some \xcp = \a\\ exec'' True pc' have "pc' = length (compE2 e1) + 3 + fst pcd" "stk' = Addr a # drop (length stk - snd pcd) stk @ STK" by(auto elim!: exec_meth.cases dest!: match_ex_table_stack_xliftD simp: match_ex_table_append split: if_split_asm) ultimately show ?thesis using \xcp = \a\\ by(auto simp add: eval_nat_numeral shift_compxE2 ac_simps) qed qed next case False with pc have [simp]: "pc = length (compE2 e2)" by simp with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros split: if_split_asm simp add: match_ex_table_append_not_pcs eval_nat_numeral)(simp_all add: matches_ex_entry_def) qed next case bisim1Sync5 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros split: if_split_asm) next case bisim1Sync6 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros split: if_split_asm) next case bisim1Sync7 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros split: if_split_asm) next case bisim1Sync8 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros split: if_split_asm) next case (bisim1Sync9 e1 n e2 V a xs) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [Addr a] STK xs (8 + length (compE2 e1) + length (compE2 e2)) None stk' loc' pc' xcp'\ let ?pre = "compE2 e1 @ Dup # Store V # MEnter # compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ [ThrowExc]) (stack_xlift (length STK) (compxE2 (sync\<^bsub>V\<^esub> (e1) e2) 0 0) @ shift (length ?pre) []) t h (Addr a # STK, xs, length ?pre + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: eval_nat_numeral) hence "exec_meth_d (compP2 P) [ThrowExc] [] t h (Addr a # STK, xs, 0, None) ta h' (stk', loc', pc' - length ?pre, xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) moreover from exec' have "pc' = 8 + length (compE2 e1) + length (compE2 e2)" "stk' = Addr a # STK" by(auto elim!: exec_meth.cases) ultimately show ?case by(fastforce elim!: exec_meth.cases intro: exec_meth.intros) next case (bisim1Sync10 e1 n e2 V a xs) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [Addr a] STK xs (8 + length (compE2 e1) + length (compE2 e2)) \a\ stk' loc' pc' xcp'\ hence "match_ex_table (compP2 P) (cname_of h a) (8 + length (compE2 e1) + length (compE2 e2)) (stack_xlift (length STK) (compxE2 (sync\<^bsub>V\<^esub> (e1) e2) 0 0)) \ None" by(rule exec_meth.cases) auto hence False by(auto split: if_split_asm simp add: match_ex_table_append_not_pcs)(simp add: matches_ex_entry_def) thus ?case .. next case (bisim1Sync11 e1 n e2 V xs) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [Null] STK xs (Suc (Suc (length (compE2 e1)))) \addr_of_sys_xcpt NullPointer\ stk' loc' pc' xcp'\ hence "match_ex_table (compP2 P) (cname_of h (addr_of_sys_xcpt NullPointer)) (2 + length (compE2 e1)) (stack_xlift (length STK) (compxE2 (sync\<^bsub>V\<^esub> (e1) e2) 0 0)) \ None" by(rule exec_meth.cases)(auto split: if_split_asm) hence False by(auto split: if_split_asm simp add: match_ex_table_append_not_pcs)(simp add: matches_ex_entry_def) thus ?case .. next case (bisim1SyncThrow e1 n a xs stk loc pc e2 V) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ note IH = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc \a\ stk' loc' pc' xcp'\ note bisim = \P,e1,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e1)" and [simp]: "loc = xs" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e1 @ Dup # Store V # MEnter # compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 3 0) @ [(3, 3 + length (compE2 e2), None, 6 + length (compE2 e2), length STK)])) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence "?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH[OF this] show ?case by auto next case (bisim1Seq1 e1 n e1' xs stk loc pc xcp e2) note bisim1 = \P,e1,h \ (e1', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (e1;;e2) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim1 have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e1)") case True from exec have "exec_meth_d (compP2 P) (compE2 e1 @ Pop # compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 (Suc 0) 0))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) hence "?exec e1 stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule True) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e1)" by simp with bisim1 obtain v where "xcp = None" "stk = [v]" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(fastforce elim: exec_meth.cases intro: exec_meth.intros) qed next case (bisim1SeqThrow1 e1 n a xs stk loc pc e2) note bisim1 = \P,e1,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e1;;e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 e1)" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e1 @ Pop # compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 (Suc 0) 0))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) hence "?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH[OF this] show ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1Seq2 e2 n e2' xs stk loc pc xcp e1) note bisim2 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (e1;;e2) stk STK loc (Suc (length (compE2 e1) + pc)) xcp stk' loc' pc' xcp'\ from bisim2 have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e2)") case False with pc have [simp]: "pc = length (compE2 e2)" by simp from bisim2 have "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec have False by(auto elim: exec_meth.cases) thus ?thesis .. next case True from exec have exec': "exec_meth_d (compP2 P) ((compE2 e1 @ [Pop]) @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1 @ [Pop])) (stack_xlift (length STK) (compxE2 e2 0 0))) t h (stk @ STK, loc, length ((compE2 e1) @ [Pop]) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e2 stk STK loc pc xcp stk' loc' (pc' - length ((compE2 e1) @ [Pop])) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e1 @ [Pop]), xcp')" by auto from exec'' have "exec_meth_d (compP2 P) ((compE2 e1 @ [Pop]) @ compE2 e2) (compxE2 e1 0 0 @ shift (length (compE2 e1 @ [Pop])) (compxE2 e2 0 0)) t h (stk, loc, length ((compE2 e1) @ [Pop]) + pc, xcp) ta h' (stk'', loc', length ((compE2 e1) @ [Pop]) + (pc' - length ((compE2 e1) @ [Pop])), xcp')" by(rule append_exec_meth_xt) auto moreover from exec' have "pc' \ length ((compE2 e1) @ [Pop])" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2) qed next case (bisim1Cond1 e n e' xs stk loc pc xcp e1 e2) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (if (e) e1 else e2) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim 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_meth_d (compP2 P) (compE2 e @ IfFalse (2 + int (length (compE2 e1))) # compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e1 (Suc 0) 0) @ stack_xlift (length STK) (compxE2 e2 (Suc (Suc (length (compE2 e1)))) 0))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: stack_xlift_compxE2 shift_compxE2 ac_simps) hence "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule True) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros) qed next case (bisim1CondThen e1 n e1' xs stk loc pc xcp e e2) note bisim = \P,e1,h \ (e1', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (if (e) e1 else e2) stk STK loc (Suc (length (compE2 e) + pc)) xcp stk' loc' pc' xcp'\ 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 ?pre = "compE2 e @ [IfFalse (2 + int (length (compE2 e1)))]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 (Suc 0) 0)))) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: stack_xlift_compxE2 shift_compxE2 ac_simps) hence "exec_meth_d (compP2 P) (compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 (Suc 0) 0))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc' - length ?pre, xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e1 stk STK loc pc xcp stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_take_xt)(rule True) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e1) (compxE2 e1 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (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_append_xt) hence "exec_meth_d (compP2 P) (?pre @ compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (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', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt)(auto) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) next case False with pc have [simp]: "pc = length (compE2 e1)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros) qed next case (bisim1CondElse e2 n e2' xs stk loc pc xcp e e1) note IH = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (if (e) e1 else e2) stk STK loc (Suc (Suc (length (compE2 e) + length (compE2 e1) + pc))) xcp stk' loc' pc' xcp'\ note bisim = \P,e2,h \ (e2', xs) \ (stk, loc, pc, xcp)\ from bisim have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) 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) (stack_xlift (length STK) (compxE2 e 0 0 @ compxE2 e1 (Suc (length (compE2 e))) 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e2 0 0))) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: stack_xlift_compxE2 shift_compxE2 ac_simps) hence "?exec e2 stk STK loc pc xcp stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2 shift_compxEs2 ac_simps) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by blast from exec'' have "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', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt)(auto) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) moreover hence "(Suc (Suc (pc' - Suc (Suc 0)))) = pc'" by simp ultimately show ?case using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2 ac_simps eval_nat_numeral) next case (bisim1CondThrow e n a xs stk loc pc e1 e2) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (if (e) e1 else e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e @ IfFalse (2 + int (length (compE2 e1))) # compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e1 (Suc 0) 0) @ stack_xlift (length STK) (compxE2 e2 (Suc (Suc (length (compE2 e1)))) 0))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: stack_xlift_compxE2 shift_compxE2 ac_simps) hence "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH[OF this] show ?case by auto next case (bisim1While1 c n e xs) note IH = \\stk' loc' pc' xcp' STK. ?exec c [] STK xs 0 None stk' loc' pc' xcp' \ ?concl c [] STK xs 0 None stk' loc' pc' xcp'\ note exec = \?exec (while (c) e) [] STK xs 0 None stk' loc' pc' xcp'\ hence "exec_meth_d (compP2 P) (compE2 c @ IfFalse (int (length (compE2 e)) + 3) # compE2 e @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length (compE2 c)) (stack_xlift (length STK) (compxE2 e (Suc 0) 0))) t h ([] @ STK, xs, 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec c [] STK xs 0 None stk' loc' pc' xcp'" by(rule exec_meth_take_xt) simp from IH[OF this] show ?case by auto next case (bisim1While3 c n c' xs stk loc pc xcp e) note bisim = \P,c,h \ (c', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec c stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl c stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (while (c) e) stk 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 from exec have "exec_meth_d (compP2 P) (compE2 c @ IfFalse (int (length (compE2 e)) + 3) # compE2 e @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length (compE2 c)) (stack_xlift (length STK) (compxE2 e (Suc 0) 0))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec c stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule True) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 c)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros) qed next case (bisim1While4 e n e' xs stk loc pc xcp c) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (while (c) e) stk STK loc (Suc (length (compE2 c) + pc)) xcp stk' loc' pc' xcp'\ 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 ?pre = "compE2 c @ [IfFalse (int (length (compE2 e)) + 3)]" from exec have "exec_meth_d (compP2 P) ((?pre @ compE2 e) @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (?pre @ compE2 e) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(auto intro: True) hence "?exec e stk STK loc pc xcp stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by auto from exec'' have "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', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) ((?pre @ compE2 e) @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (compxE2 c 0 0 @ shift (length ?pre) (compxE2 e 0 0)) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) moreover have "-2 + (- int (length (compE2 e)) - int (length (compE2 c))) = - int (length (compE2 c)) + (-2 - int (length (compE2 e)))" by simp ultimately show ?thesis using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2 algebra_simps uminus_minus_left_commute) next case False with pc have [simp]: "pc = length (compE2 e)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros) qed next case (bisim1While6 c n e xs) note exec = \?exec (while (c) e) [] STK xs (Suc (Suc (length (compE2 c) + length (compE2 e)))) None stk' loc' pc' xcp'\ thus ?case by(rule exec_meth.cases)(simp_all, auto intro!: exec_meth.intros) next case (bisim1While7 c n e xs) note exec = \?exec (while (c) e) [] STK xs (Suc (Suc (Suc (length (compE2 c) + length (compE2 e))))) None stk' loc' pc' xcp'\ thus ?case by(rule exec_meth.cases)(simp_all, auto intro!: exec_meth.intros) next case (bisim1WhileThrow1 c n a xs stk loc pc e) note bisim = \P,c,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec c stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl c stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (while (c) e) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 c)" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 c @ IfFalse (int (length (compE2 e)) + 3) # compE2 e @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length (compE2 c)) (stack_xlift (length STK) (compxE2 e (Suc 0) 0))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec c stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH[OF this] show ?case by auto next case (bisim1WhileThrow2 e n a xs stk loc pc c) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (while (c) e) stk STK loc (Suc (length (compE2 c) + pc)) \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) let ?pre = "compE2 c @ [IfFalse (int (length (compE2 e)) + 3)]" from exec have "exec_meth_d (compP2 P) ((?pre @ compE2 e) @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, length ?pre + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (?pre @ compE2 e) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, (length ?pre + pc), \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(auto intro: pc) hence "?exec e stk STK loc pc \a\ stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length ?pre, xcp')" by auto from exec'' have "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, \a\) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) ((?pre @ compE2 e) @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (compxE2 c 0 0 @ shift (length ?pre) (compxE2 e 0 0)) t h (stk, loc, length ?pre + pc, \a\) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) moreover have "-2 + (- int (length (compE2 e)) - int (length (compE2 c))) = - int (length (compE2 c)) + (-2 - int (length (compE2 e)))" by simp ultimately show ?case using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2 algebra_simps uminus_minus_left_commute) next case (bisim1Throw1 e n e' xs stk loc pc xcp) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (throw e) stk STK loc pc xcp stk' loc' pc' xcp'\ 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 e stk STK loc pc xcp stk' loc' pc' xcp'" by(auto elim: exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros split: if_split_asm) qed next case bisim1Throw2 thus ?case apply(auto elim!:exec_meth.cases intro: exec_meth.intros dest!: match_ex_table_stack_xliftD) apply(auto intro: exec_meth.intros dest!: match_ex_table_stack_xliftD intro!: exI) apply(auto simp add: le_Suc_eq) done next case bisim1ThrowNull thus ?case apply(auto elim!:exec_meth.cases intro: exec_meth.intros dest!: match_ex_table_stack_xliftD) apply(auto intro: exec_meth.intros dest!: match_ex_table_stack_xliftD intro!: exI) apply(auto simp add: le_Suc_eq) done next case (bisim1ThrowThrow e n a xs stk loc pc) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (throw e) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) with exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(auto elim: exec_meth_take simp add: compxE2_size_convs) from IH[OF this] show ?case by auto next case (bisim1Try e n e' xs stk loc pc xcp e2 C' V) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (try e catch(C' V) e2) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim 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 e @ Goto (int (length (compE2 e2)) + 2) # Store V # compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 (Suc (Suc 0)) 0)) @ [(0, length (compE2 e), \C'\, Suc (length (compE2 e)), length STK)]) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) show ?thesis proof(cases xcp) case None with exec' True have "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" apply - apply (erule exec_meth.cases) apply (cases "compE2 e ! pc") apply (fastforce simp add: is_Ref_def intro: exec_meth.intros split: if_split_asm cong del: image_cong_simp)+ done from IH[OF this] show ?thesis by auto next case (Some a) with exec' have [simp]: "h' = h" "loc' = loc" "xcp' = None" "ta = \" by(auto elim: exec_meth.cases) show ?thesis proof(cases "match_ex_table (compP2 P) (cname_of h a) pc (compxE2 e 0 0)") case (Some pcd) from exec \xcp = \a\\ Some pc have stk': "stk' = Addr a # (drop (length stk - snd pcd) stk) @ STK" by(auto elim!: exec_meth.cases simp add: match_ex_table_append split: if_split_asm dest!: match_ex_table_stack_xliftD) from exec' \xcp = \a\\ Some pc have "exec_meth_d (compP2 P) (compE2 e) (stack_xlift (length STK) (compxE2 e 0 0)) t h (stk @ STK, loc, pc, \a\) \ h (Addr a # (drop (length (stk @ STK) - (snd pcd + length STK)) (stk @ STK)), loc, pc', None)" apply - apply(rule exec_meth.intros) apply(auto elim!: exec_meth.cases simp add: match_ex_table_append split: if_split_asm dest!: match_ex_table_shift_pcD match_ex_table_stack_xliftD) done from IH[unfolded \ta = \\ \xcp = \a\\ \h' = h\, OF this] have stk: "Addr a # drop (length stk - snd pcd) (stk @ STK) = Addr a # drop (length stk - snd pcd) stk @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, \a\) \ h (Addr a # drop (length stk - snd pcd) stk, loc, pc', None)" by auto thus ?thesis using Some stk' \xcp = \a\\ by(auto) next case None with Some exec pc have stk': "stk' = Addr a # STK" and pc': "pc' = Suc (length (compE2 e))" and subcls: "compP2 P \ cname_of h a \\<^sup>* C'" by(auto elim!: exec_meth.cases split: if_split_asm simp add: match_ex_table_append_not_pcs)(simp add: matches_ex_entry_def) moreover from Some True None pc' subcls have "exec_meth_d (compP2 P) (compE2 (try e catch(C' V) e2)) (compxE2 (try e catch(C' V) e2) 0 0) t h (stk, loc, pc, \a\) \ h (Addr a # drop (length stk - 0) stk, loc, pc', None)" by -(rule exec_catch,auto simp add: match_ex_table_append_not_pcs matches_ex_entry_def) ultimately show ?thesis using Some by auto qed qed next case False with pc have [simp]: "pc = length (compE2 e)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros split: if_split_asm) qed next case bisim1TryCatch1 thus ?case by(auto elim!: exec_meth.cases intro!: exec_meth.intros split: if_split_asm) next case (bisim1TryCatch2 e2 n e' xs stk loc pc xcp e C' V) note bisim = \P,e2,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (try e catch(C' V) e2) stk STK loc (Suc (Suc (length (compE2 e) + pc))) xcp stk' loc' pc' xcp'\ let ?pre = "compE2 e @ [Goto (int (length (compE2 e2)) + 2), Store V]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e2 0 0))) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" proof(cases) case (exec_catch xcp'' d) let ?stk = "stk @ STK" and ?PC = "Suc (Suc (length (compE2 e) + pc))" note s = \stk' = Addr xcp'' # drop (length ?stk - d) ?stk\ \ta = \\ \h' = h\ \xcp' = None\ \xcp = \xcp''\\ \loc' = loc\ from \match_ex_table (compP2 P) (cname_of h xcp'') ?PC (stack_xlift (length STK) (compxE2 (try e catch(C' V) e2) 0 0)) = \(pc', d)\\ \d \ length ?stk\ show ?thesis unfolding s by -(rule exec_meth.exec_catch, simp_all add: shift_compxE2 stack_xlift_compxE2, simp add: match_ex_table_append add: matches_ex_entry_def) qed(auto intro: exec_meth.intros simp add: shift_compxE2 stack_xlift_compxE2) hence "?exec e2 stk STK loc pc xcp stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by auto from exec'' have "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', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt) auto hence "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', length ?pre + (pc' - length ?pre), xcp')" by(rule exec_meth.cases)(auto intro: exec_meth.intros simp add: match_ex_table_append_not_pcs) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) moreover hence "(Suc (Suc (pc' - Suc (Suc 0)))) = pc'" by simp ultimately show ?case using stk' by(auto simp add: shift_compxE2 eval_nat_numeral) 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\)\ note exec = \?exec (try e catch(C'' V) e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ note a = \typeof_addr h a = \Class_type C'\\ \\ P \ C' \\<^sup>* C''\ 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) moreover from bisim have "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) ultimately have False using exec a apply(auto elim!: exec_meth.cases simp add: outside_pcs_compxE2_not_matches_entry outside_pcs_not_matches_entry split: if_split_asm) apply(auto simp add: compP2_def match_ex_entry match_ex_table_append_not_pcs cname_of_def split: if_split_asm) done 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\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (try e catch(C' V) e2) stk STK loc (Suc (Suc (length (compE2 e) + pc))) \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e2)" by(auto dest: bisim1_ThrowD) let ?pre = "compE2 e @ [Goto (int (length (compE2 e2)) + 2), Store V]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e2 0 0))) t h (stk @ STK, loc, length ?pre + pc, \a\) ta h' (stk', loc', pc', xcp')" proof(cases) case (exec_catch d) let ?stk = "stk @ STK" and ?PC = "Suc (Suc (length (compE2 e) + pc))" note s = \stk' = Addr a # drop (length ?stk - d) ?stk\ \loc' = loc\ \ta = \\ \h' = h\ \xcp' = None\ from \match_ex_table (compP2 P) (cname_of h a) ?PC (stack_xlift (length STK) (compxE2 (try e catch(C' V) e2) 0 0)) = \(pc', d)\\ \d \ length ?stk\ show ?thesis unfolding s by -(rule exec_meth.exec_catch, simp_all add: shift_compxE2 stack_xlift_compxE2, simp add: match_ex_table_append add: matches_ex_entry_def) qed hence "?exec e2 stk STK loc pc \a\ stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length ?pre, xcp')" by auto from exec'' have "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, \a\) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt) auto hence "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, \a\) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule exec_meth.cases)(auto intro!: exec_meth.intros simp add: match_ex_table_append_not_pcs) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) moreover hence "(Suc (Suc (pc' - Suc (Suc 0)))) = pc'" by simp ultimately show ?case using stk' by(auto simp add: shift_compxE2 eval_nat_numeral) next case bisims1Nil thus ?case by(auto elim: exec_meth.cases) next case (bisims1List1 e n e' xs stk loc pc xcp es) note bisim1 = \P,e,h \ (e', xs ) \ (stk, loc, pc, xcp)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?execs es [] STK xs 0 None stk' loc' pc' xcp' \ ?concls es [] STK xs 0 None stk' loc' pc' xcp'\ note exec = \?execs (e # es) stk STK loc pc xcp stk' loc' pc' xcp'\ 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 e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxEs2_size_convs)(erule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 e)" by simp with bisim1 obtain v where s: "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec pc have exec': "exec_meth_d (compP2 P) (compE2 e @ compEs2 es) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length (v # STK)) (compxEs2 es 0 0))) t h ([] @ v # STK, loc, length (compE2 e) + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) hence "?execs es [] (v # STK) loc 0 None stk' loc' (pc' - length (compE2 e)) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec'': "exec_meth_d (compP2 P) (compEs2 es) (compxEs2 es 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" by auto from exec'' have "exec_meth_d (compP2 P) (compEs2 es) (stack_xlift (length [v]) (compxEs2 es 0 0)) t h ([] @ [v], loc, 0, None) ta h' (stk'' @ [v], loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_stk_offer) hence "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'' @ [v], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule append_exec_meth_xt) auto moreover from exec' have "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using s pc stk' by(auto simp add: shift_compxEs2 stack_xlift_compxEs2) qed next case (bisims1List2 es n es' xs stk loc pc xcp e v) note bisim = \P,es,h \ (es',xs) [\] (stk,loc,pc,xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?execs es stk STK loc pc xcp stk' loc' pc' xcp' \ ?concls es stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?execs (e # es) (stk @ [v]) STK loc (length (compE2 e) + pc) xcp stk' loc' pc' xcp'\ from exec have exec': "exec_meth_d (compP2 P) (compE2 e @ compEs2 es) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length (v # STK)) (compxEs2 es 0 0))) t h (stk @ v # STK, loc, length (compE2 e) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) hence "?execs es stk (v # STK) loc pc xcp stk' loc' (pc' - length (compE2 e)) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec'': "exec_meth_d (compP2 P) (compEs2 es) (compxEs2 es 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" by auto from exec'' have "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'' @ [v], loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_stk_offer) hence "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'' @ [v], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule append_exec_meth_xt) auto moreover from exec' have "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: shift_compxEs2 stack_xlift_compxEs2) next case (bisim1Sync12 e1 n e2 V a xs v v') note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [v, v'] STK xs (4 + length (compE2 e1) + length (compE2 e2)) \a\ stk' loc' pc' xcp'\ thus ?case by(auto elim!: exec_meth.cases split: if_split_asm simp add: match_ex_table_append_not_pcs)(simp add: matches_ex_entry_def) next case (bisim1Sync14 e1 n e2 V a xs v a') note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [v, Addr a'] STK xs (7 + length (compE2 e1) + length (compE2 e2)) \a\ stk' loc' pc' xcp'\ thus ?case by(auto elim!: exec_meth.cases split: if_split_asm simp add: match_ex_table_append_not_pcs)(simp add: matches_ex_entry_def) qed lemma shows bisim1_callD: "\ P,e,h \ (e', xs) \ (stk, loc, pc, xcp); call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M' n0 \ \ M = M'" and bisims1_callD: "\ P,es,h \ (es',xs) [\] (stk,loc,pc, xcp); calls1 es' = \(a, M, vs)\; compEs2 es ! pc = Invoke M' n0 \ \ M = M'" proof(induct e "n :: nat" e' xs stk loc pc xcp and es "n :: nat" es' xs stk loc pc xcp rule: bisim1_bisims1_inducts_split) case bisim1AAss1 thus ?case apply(simp (no_asm_use) split: if_split_asm add: is_val_iff) apply(fastforce dest: bisim_Val_pc_not_Invoke) apply(fastforce dest: bisim_Val_pc_not_Invoke) apply(fastforce dest: bisim_Val_pc_not_Invoke bisim1_pc_length_compE2)+ done next case bisim1Call1 thus ?case apply(clarsimp split: if_split_asm simp add: is_vals_conv) apply(drule bisim_Val_pc_not_Invoke, simp, fastforce) apply(drule bisim_Val_pc_not_Invoke, simp, fastforce) apply(drule bisim1_pc_length_compE2, clarsimp simp add: neq_Nil_conv) apply(drule bisim1_pc_length_compE2, simp) apply(drule bisim1_pc_length_compE2, simp) apply(drule bisim1_pc_length_compE2, simp) apply(drule bisim1_call_pcD, simp, simp) apply(drule bisim1_call_pcD, simp, simp) done next case bisim1CallParams thus ?case apply(clarsimp split: if_split_asm simp add: is_vals_conv) apply(drule bisims_Val_pc_not_Invoke, simp, fastforce) apply(drule bisims1_pc_length_compEs2, simp) apply(drule bisims1_calls_pcD, simp, simp) done qed(fastforce split: if_split_asm dest: bisim1_pc_length_compE2 bisims1_pc_length_compEs2 bisim1_call_pcD bisims1_calls_pcD bisim1_call_xcpNone bisims1_calls_xcpNone bisim_Val_pc_not_Invoke bisims_Val_pc_not_Invoke)+ lemma bisim1_xcpD: "P,e,h \ (e', xs) \ (stk, loc, pc, \a\) \ pc < length (compE2 e)" and bisims1_xcpD: "P,es,h \ (es', xs) [\] (stk, loc, pc, \a\) \ pc < length (compEs2 es)" by(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) simp_all lemma bisim1_match_Some_stk_length: "\ P,E,h \ (e, xs) \ (stk, loc, pc, \a\); match_ex_table (compP2 P) (cname_of h a) pc (compxE2 E 0 0) = \(pc', d)\ \ \ d \ length stk" and bisims1_match_Some_stk_length: "\ P,Es,h \ (es, xs) [\] (stk, loc, pc, \a\); match_ex_table (compP2 P) (cname_of h a) pc (compxEs2 Es 0 0) = \(pc', d)\ \ \ d \ length stk" proof(induct "(e, xs)" "(stk, loc, pc, \a :: 'addr\)" and "(es, xs)" "(stk, loc, pc, \a :: 'addr\)" arbitrary: pc' d e xs stk loc pc and pc' d es xs stk loc pc rule: bisim1_bisims1.inducts) case bisim1Call1 thus ?case by(fastforce dest: bisim1_xcpD simp add: match_ex_table_append match_ex_table_not_pcs_None) next case bisim1CallThrowObj thus ?case by(fastforce dest: bisim1_xcpD simp add: match_ex_table_append match_ex_table_not_pcs_None) next case bisim1Sync4 thus ?case apply(clarsimp simp add: match_ex_table_not_pcs_None match_ex_table_append matches_ex_entry_def split: if_split_asm) apply(fastforce simp add: match_ex_table_compxE2_shift_conv dest: bisim1_xcpD) done next case bisim1Try thus ?case by(fastforce simp add: match_ex_table_append matches_ex_entry_def match_ex_table_not_pcs_None dest: bisim1_xcpD split: if_split_asm) next case bisim1TryCatch2 thus ?case apply(clarsimp simp add: match_ex_table_not_pcs_None match_ex_table_append matches_ex_entry_def split: if_split_asm) apply(fastforce simp add: match_ex_table_compxE2_shift_conv dest: bisim1_xcpD) done next case bisim1TryFail thus ?case by(fastforce simp add: match_ex_table_append matches_ex_entry_def match_ex_table_not_pcs_None dest: bisim1_xcpD split: if_split_asm) next case bisim1TryCatchThrow thus ?case apply(clarsimp simp add: match_ex_table_not_pcs_None match_ex_table_append matches_ex_entry_def split: if_split_asm) apply(fastforce simp add: match_ex_table_compxE2_shift_conv dest: bisim1_xcpD) done next case bisims1List1 thus ?case by(fastforce simp add: match_ex_table_append split: if_split_asm dest: bisim1_xcpD match_ex_table_pcsD) qed(fastforce simp add: match_ex_table_not_pcs_None match_ex_table_append match_ex_table_compxE2_shift_conv match_ex_table_compxEs2_shift_conv match_ex_table_compxE2_stack_conv match_ex_table_compxEs2_stack_conv matches_ex_entry_def dest: bisim1_xcpD)+ end locale J1_JVM_heap_conf_base = J1_JVM_heap_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write + J1_heap_conf_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf P + JVM_heap_conf_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf "compP2 P" for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and hconf :: "'heap \ bool" and P :: "'addr J1_prog" begin inductive bisim1_list1 :: "'thread_id \ 'heap \ 'addr expr1 \ 'addr locals1 \ ('addr expr1 \ 'addr locals1) list \ 'addr option \ 'addr frame list \ bool" for t :: 'thread_id and h :: 'heap where bl1_Normal: "\ compTP P \ t:(xcp, h, (stk, loc, C, M, pc) # frs) \; P \ C sees M : Ts\T = \body\ in D; P,blocks1 0 (Class D#Ts) body, h \ (e, xs) \ (stk, loc, pc, xcp); max_vars e \ length xs; list_all2 (bisim1_fr P h) exs frs \ \ bisim1_list1 t h (e, xs) exs xcp ((stk, loc, C, M, pc) # frs)" | bl1_finalVal: "\ hconf h; preallocated h \ \ bisim1_list1 t h (Val v, xs) [] None []" | bl1_finalThrow: "\ hconf h; preallocated h \ \ bisim1_list1 t h (Throw a, xs) [] \a\ []" fun wbisim1 :: "'thread_id \ ((('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list) \ 'heap, ('addr option \ 'addr frame list) \ 'heap) bisim" where "wbisim1 t ((ex, exs), h) ((xcp, frs), h') \ h = h' \ bisim1_list1 t h ex exs xcp frs" lemma new_thread_conf_compTP: assumes hconf: "hconf h" "preallocated h" and ha: "typeof_addr h a = \Class_type C\" and sub: "typeof_addr h (thread_id2addr t) = \Class_type C'\" "P \ C' \\<^sup>* Thread" and sees: "P \ C sees M: []\T = \meth\ in D" shows "compTP P \ t:(None, h, [([], Addr a # replicate (max_vars meth) undefined_value, D, M, 0)]) \" proof - from ha sees_method_decl_above[OF sees] have "P,h \ Addr a :\ Class D" by(simp add: conf_def) moreover hence "compP2 P,h \ Addr a :\ Class D" by(simp add: compP2_def) hence "compP2 P,h \ Addr a # replicate (max_vars meth) undefined_value [:\\<^sub>\] map (\i. if i = 0 then OK ([Class D] ! i) else Err) [0..i. if i = 0 then OK ([Class D] ! i) else Err) [0..l (Suc (max_vars meth)) [Class D] {0}) (compE2 meth @ [Return]) ([], Addr a # replicate (max_vars meth) undefined_value, D, M, 0)" by(simp add: TC0.ty\<^sub>l_def conf_f_def2 compP2_def) with hconf ha sub sees_method_compP[OF sees, where f="\C M Ts T. compMb2"] sees_method_idemp[OF sees] show ?thesis by(auto simp add: TC0.ty\<^sub>i'_def correct_state_def compTP_def tconf_def)(fastforce simp add: compP2_def compMb2_def tconf_def intro: sees_method_idemp)+ qed lemma ta_bisim12_extTA2J1_extTA2JVM: assumes nt: "\n T C M a h. \ n < length \ta\\<^bsub>t\<^esub>; \ta\\<^bsub>t\<^esub> ! n = NewThread T (C, M, a) h \ \ typeof_addr h a = \Class_type C\ \ (\C'. typeof_addr h (thread_id2addr T) = \Class_type C'\ \ P \ C' \\<^sup>* Thread) \ (\T meth D. P \ C sees M:[]\T =\meth\ in D) \ hconf h \ preallocated h" shows "ta_bisim wbisim1 (extTA2J1 P ta) (extTA2JVM (compP2 P) ta)" proof - { fix n t C M a m assume "n < length \ta\\<^bsub>t\<^esub>" and "\ta\\<^bsub>t\<^esub> ! n = NewThread t (C, M, a) m" from nt[OF this] obtain T meth D C' where ma: "typeof_addr m a = \Class_type C\" and sees: "P \ C sees M: []\T = \meth\ in D" and sub: "typeof_addr m (thread_id2addr t) = \Class_type C'\" "P \ C' \\<^sup>* Thread" and mconf: "hconf m" "preallocated m" by fastforce from sees_method_compP[OF sees, where f="\C M Ts T. compMb2"] have sees': "compP2 P \ C sees M: []\T = \(max_stack meth, max_vars meth, compE2 meth @ [Return], compxE2 meth 0 0)\ in D" by(simp add: compMb2_def compP2_def) have "bisim1_list1 t m ({0:Class D=None; meth}, Addr a # replicate (max_vars meth) undefined_value) ([]) None [([], Addr a # replicate (max_vars meth) undefined_value, D, M, 0)]" proof from mconf ma sub sees show "compTP P \ t:(None, m, [([], Addr a # replicate (max_vars meth) undefined_value, D, M, 0)]) \" by(rule new_thread_conf_compTP) from sees show "P \ D sees M: []\T = \meth\ in D" by(rule sees_method_idemp) show "list_all2 (bisim1_fr P m) [] []" by simp show "P,blocks1 0 [Class D] meth,m \ ({0:Class D=None; meth}, Addr a # replicate (max_vars meth) undefined_value) \ ([], Addr a # replicate (max_vars meth) undefined_value, 0, None)" by simp(rule bisim1_refl) qed simp with sees sees' have "bisim1_list1 t m ({0:Class (fst (method P C M))=None; the (snd (snd (snd (method P C M))))}, Addr a # replicate (max_vars (the (snd (snd (snd (method P C M)))))) undefined_value) [] None [([], Addr a # replicate (fst (snd (the (snd (snd (snd (method (compP2 P) C M))))))) undefined_value, fst (method (compP2 P) C M), M, 0)]" by simp } thus ?thesis apply(auto simp add: ta_bisim_def intro!: list_all2_all_nthI) apply(case_tac "\ta\\<^bsub>t\<^esub> ! n", auto simp add: extNTA2JVM_def) done qed end definition no_call2 :: "'addr expr1 \ pc \ bool" where "no_call2 e pc \ (pc \ length (compE2 e)) \ (pc < length (compE2 e) \ (\M n. compE2 e ! pc \ Invoke M n))" definition no_calls2 :: "'addr expr1 list \ pc \ bool" where "no_calls2 es pc \ (pc \ length (compEs2 es)) \ (pc < length (compEs2 es) \ (\M n. compEs2 es ! pc \ Invoke M n))" locale J1_JVM_conf_read = J1_JVM_heap_conf_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf P + JVM_conf_read addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf "compP2 P" for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and hconf :: "'heap \ bool" and P :: "'addr J1_prog" locale J1_JVM_heap_conf = J1_JVM_heap_conf_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf P + JVM_heap_conf addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf "compP2 P" for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and hconf :: "'heap \ bool" and P :: "'addr J1_prog" begin lemma red_external_ta_bisim21: "\ wf_prog wf_md P; P,t \ \a\M(vs), h\ -ta\ext \va, h'\; hconf h'; preallocated h' \ \ ta_bisim wbisim1 (extTA2J1 P ta) (extTA2JVM (compP2 P) ta)" apply(rule ta_bisim12_extTA2J1_extTA2JVM) apply(frule (1) red_external_new_thread_sees) apply(fastforce simp add: in_set_conv_nth) apply(frule red_ext_new_thread_heap) apply(fastforce simp add: in_set_conv_nth) apply(frule red_external_new_thread_exists_thread_object[unfolded compP2_def, simplified]) apply(fastforce simp add: in_set_conv_nth) apply simp done lemma ta_bisim_red_extTA2J1_extTA2JVM: assumes wf: "wf_prog wf_md P" and red: "uf,P,t' \1 \e, s\ -ta\ \e', s'\" and hconf: "hconf (hp s')" "preallocated (hp s')" shows "ta_bisim wbisim1 (extTA2J1 P ta) (extTA2JVM (compP2 P) ta)" proof - { fix n t C M a H assume len: "n < length \ta\\<^bsub>t\<^esub>" and tan: "\ta\\<^bsub>t\<^esub> ! n = NewThread t (C, M, a) H" hence nt: "NewThread t (C, M, a) H \ set \ta\\<^bsub>t\<^esub>" unfolding set_conv_nth by(auto intro!: exI) from red1_new_threadD[OF red nt] obtain ad M' vs va T C' Ts' Tr' D' where rede: "P,t' \ \ad\M'(vs),hp s\ -ta\ext \va,hp s'\" and ad: "typeof_addr (hp s) ad = \T\" by blast from red_ext_new_thread_heap[OF rede nt] have [simp]: "hp s' = H" by simp from red_external_new_thread_sees[OF wf rede nt] obtain T body D where Ha: "typeof_addr H a = \Class_type C\" and sees: "P \ C sees M:[]\T=\body\ in D" by auto have sees': "compP2 P \ C sees M:[]\T=\(max_stack body, max_vars body, compE2 body @ [Return], compxE2 body 0 0)\ in D" using sees unfolding compP2_def compMb2_def Let_def by(auto dest: sees_method_compP) from red_external_new_thread_exists_thread_object[unfolded compP2_def, simplified, OF rede nt] hconf Ha sees have "compTP P \ t:(None, H, [([], Addr a # replicate (max_vars body) undefined_value, D, M, 0)]) \" by(auto intro: new_thread_conf_compTP) hence "bisim1_list1 t H ({0:Class D=None; body}, Addr a # replicate (max_vars body) undefined_value) [] None [([], Addr a # replicate (max_vars body) undefined_value, D, M, 0)]" proof from sees show "P \ D sees M:[]\T=\body\ in D" by(rule sees_method_idemp) show "P,blocks1 0 [Class D] body,H \ ({0:Class D=None; body}, Addr a # replicate (max_vars body) undefined_value) \ ([], Addr a # replicate (max_vars body) undefined_value, 0, None)" by(auto intro: bisim1_refl) qed simp_all hence "bisim1_list1 t H ({0:Class (fst (method P C M))=None; the (snd (snd (snd (method P C M))))}, Addr a # replicate (max_vars (the (snd (snd (snd (method P C M)))))) undefined_value) [] None [([], Addr a # replicate (fst (snd (the (snd (snd (snd (method (compP2 P) C M))))))) undefined_value, fst (method (compP2 P) C M), M, 0)]" using sees sees' by simp } thus ?thesis apply(auto simp add: ta_bisim_def intro!: list_all2_all_nthI) apply(case_tac "\ta\\<^bsub>t\<^esub> ! n") apply(auto simp add: extNTA2JVM_def extNTA2J1_def) done qed end sublocale J1_JVM_conf_read < heap_conf?: J1_JVM_heap_conf by(unfold_locales) sublocale J1_JVM_conf_read < heap?: J1_heap apply(rule J1_heap.intro) apply(subst compP_heap[symmetric, where f="\_ _ _ _. compMb2", folded compP2_def]) apply(unfold_locales) done end diff --git a/thys/JinjaThreads/Execute/Random_Scheduler.thy b/thys/JinjaThreads/Execute/Random_Scheduler.thy --- a/thys/JinjaThreads/Execute/Random_Scheduler.thy +++ b/thys/JinjaThreads/Execute/Random_Scheduler.thy @@ -1,206 +1,206 @@ (* Title: JinjaThreads/Execute/Random_Scheduler.thy Author: Andreas Lochbihler *) section \Random scheduler\ theory Random_Scheduler imports Scheduler begin type_synonym random_scheduler = Random.seed abbreviation (input) random_scheduler_invar :: "random_scheduler \ 't set \ bool" where "random_scheduler_invar \ \_ _. True" locale random_scheduler_base = scheduler_ext_base final r convert_RA thr_\ thr_invar thr_lookup thr_update thr_iterate ws_\ ws_invar ws_lookup ws_update ws_sel is_\ is_invar is_memb is_ins is_delete thr'_\ thr'_invar thr'_empty thr'_ins_dj for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and "output" :: "random_scheduler \ 't \ ('l,'t,'x,'m,'w,'o) thread_action \ 'q option" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and thr_lookup :: "'t \ 'm_t \ ('x \ 'l released_locks)" and thr_update :: "'t \ 'x \ 'l released_locks \ 'm_t \ 'm_t" and thr_iterate :: "'m_t \ ('t \ ('x \ 'l released_locks), 's_t) set_iterator" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and ws_lookup :: "'t \ 'm_w \ 'w wait_set_status" and ws_update :: "'t \ 'w wait_set_status \ 'm_w \ 'm_w" and ws_delete :: "'t \ 'm_w \ 'm_w" and ws_iterate :: "'m_w \ ('t \ 'w wait_set_status, 'm_w) set_iterator" and ws_sel :: "'m_w \ ('t \ 'w wait_set_status \ bool) \ ('t \ 'w wait_set_status)" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" and is_memb :: "'t \ 's_i \ bool" and is_ins :: "'t \ 's_i \ 's_i" and is_delete :: "'t \ 's_i \ 's_i" and thr'_\ :: "'s_t \ 't set" and thr'_invar :: "'s_t \ bool" and thr'_empty :: "unit \ 's_t" and thr'_ins_dj :: "'t \ 's_t \ 's_t" + fixes thr'_to_list :: "'s_t \ 't list" begin definition next_thread :: "random_scheduler \ 's_t \ ('t \ random_scheduler) option" where "next_thread seed active = (let ts = thr'_to_list active in if ts = [] then None else Some (Random.select (thr'_to_list active) seed))" definition random_scheduler :: "('l,'t,'x,'m,'w,'o,'m_t,'m_w,'s_i,random_scheduler) scheduler" where "random_scheduler seed s = (do { (t, seed') \ next_thread seed (active_threads s); step_thread (\ta. seed') s t })" end locale random_scheduler = random_scheduler_base final r convert_RA "output" thr_\ thr_invar thr_lookup thr_update thr_iterate ws_\ ws_invar ws_lookup ws_update ws_delete ws_iterate ws_sel is_\ is_invar is_memb is_ins is_delete thr'_\ thr'_invar thr'_empty thr'_ins_dj thr'_to_list + scheduler_ext_aux final r convert_RA thr_\ thr_invar thr_lookup thr_update thr_iterate ws_\ ws_invar ws_lookup ws_update ws_sel is_\ is_invar is_memb is_ins is_delete thr'_\ thr'_invar thr'_empty thr'_ins_dj + ws: map_delete ws_\ ws_invar ws_delete + ws: map_iteratei ws_\ ws_invar ws_iterate + thr': set_to_list thr'_\ thr'_invar thr'_to_list for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and "output" :: "random_scheduler \ 't \ ('l,'t,'x,'m,'w,'o) thread_action \ 'q option" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and thr_lookup :: "'t \ 'm_t \ ('x \ 'l released_locks)" and thr_update :: "'t \ 'x \ 'l released_locks \ 'm_t \ 'm_t" and thr_iterate :: "'m_t \ ('t \ ('x \ 'l released_locks), 's_t) set_iterator" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and ws_lookup :: "'t \ 'm_w \ 'w wait_set_status" and ws_update :: "'t \ 'w wait_set_status \ 'm_w \ 'm_w" and ws_delete :: "'t \ 'm_w \ 'm_w" and ws_iterate :: "'m_w \ ('t \ 'w wait_set_status, 'm_w) set_iterator" and ws_sel :: "'m_w \ ('t \ 'w wait_set_status \ bool) \ ('t \ 'w wait_set_status)" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" and is_memb :: "'t \ 's_i \ bool" and is_ins :: "'t \ 's_i \ 's_i" and is_delete :: "'t \ 's_i \ 's_i" and thr'_\ :: "'s_t \ 't set" and thr'_invar :: "'s_t \ bool" and thr'_empty :: "unit \ 's_t" and thr'_ins_dj :: "'t \ 's_t \ 's_t" and thr'_to_list :: "'s_t \ 't list" begin lemma next_thread_eq_None_iff: assumes "thr'_invar active" "random_scheduler_invar seed T" shows "next_thread seed active = None \ thr'_\ active = {}" using thr'.to_list_correct[OF assms(1)] by(auto simp add: next_thread_def neq_Nil_conv) lemma next_thread_eq_SomeD: assumes "next_thread seed active = Some (t, seed')" and "thr'_invar active" "random_scheduler_invar seed T" shows "t \ thr'_\ active" using assms by(auto simp add: next_thread_def thr'.to_list_correct split: if_split_asm dest: select[of _ seed]) lemma random_scheduler_spec: assumes det: "\.deterministic I" shows "scheduler_spec final r random_scheduler random_scheduler_invar thr_\ thr_invar ws_\ ws_invar is_\ is_invar I" proof fix \ s assume rr: "random_scheduler \ s = None" and invar: "random_scheduler_invar \ (dom (thr_\ (thr s)))" "state_invar s" "state_\ s \ I" from invar(2) have "thr'_invar (active_threads s)" by(rule active_threads_correct) thus "\.active_threads (state_\ s) = {}" using rr invar by(auto simp add: random_scheduler_def Option_bind_eq_None_conv next_thread_eq_None_iff step_thread_eq_None_conv[OF det] dest: next_thread_eq_SomeD) next fix \ s t \' assume rr: "random_scheduler \ s = \(t, None, \')\" and invar: "random_scheduler_invar \ (dom (thr_\ (thr s)))" "state_invar s" "state_\ s \ I" thus "\x ln n. thr_\ (thr s) t = \(x, ln)\ \ 0 < ln $ n \ \ waiting (ws_\ (wset s) t) \ may_acquire_all (locks s) t ln" - by(fastforce simp add: random_scheduler_def Option_bind_eq_Some_conv dest: step_thread_Some_NoneD[OF det]) + by(fastforce simp add: random_scheduler_def bind_eq_Some_conv dest: step_thread_Some_NoneD[OF det]) next fix \ s t ta x' m' \' assume rr: "random_scheduler \ s = \(t, \(ta, x', m')\, \')\" and invar: "random_scheduler_invar \ (dom (thr_\ (thr s)))" "state_invar s" "state_\ s \ I" thus "\x. thr_\ (thr s) t = \(x, no_wait_locks)\ \ Predicate.eval (r t (x, shr s)) (ta, x', m') \ \.actions_ok (state_\ s) t ta" - by(auto simp add: random_scheduler_def Option_bind_eq_Some_conv dest: step_thread_Some_SomeD[OF det]) + by(auto simp add: random_scheduler_def bind_eq_Some_conv dest: step_thread_Some_SomeD[OF det]) qed simp_all end sublocale random_scheduler_base < scheduler_base final r convert_RA "random_scheduler" "output" "pick_wakeup_via_sel (\s P. ws_sel s (\(k,v). P k v))" random_scheduler_invar thr_\ thr_invar thr_lookup thr_update ws_\ ws_invar ws_lookup ws_update ws_delete ws_iterate is_\ is_invar is_memb is_ins is_delete for n0 . sublocale random_scheduler < pick_wakeup_spec final r convert_RA "pick_wakeup_via_sel (\s P. ws_sel s (\(k,v). P k v))" random_scheduler_invar thr_\ thr_invar ws_\ ws_invar is_\ is_invar by(rule pick_wakeup_spec_via_sel)(unfold_locales) context random_scheduler begin lemma random_scheduler_scheduler: assumes det: "\.deterministic I" shows "scheduler final r convert_RA random_scheduler (pick_wakeup_via_sel (\s P. ws_sel s (\(k,v). P k v))) random_scheduler_invar thr_\ thr_invar thr_lookup thr_update ws_\ ws_invar ws_lookup ws_update ws_delete ws_iterate is_\ is_invar is_memb is_ins is_delete I" proof - interpret scheduler_spec final r convert_RA random_scheduler random_scheduler_invar thr_\ thr_invar ws_\ ws_invar is_\ is_invar I using det by(rule random_scheduler_spec) show ?thesis by(unfold_locales)(rule \.deterministic_invariant3p[OF det]) qed end subsection \Code generator setup\ lemmas [code] = random_scheduler_base.next_thread_def random_scheduler_base.random_scheduler_def end diff --git a/thys/JinjaThreads/Execute/Scheduler.thy b/thys/JinjaThreads/Execute/Scheduler.thy --- a/thys/JinjaThreads/Execute/Scheduler.thy +++ b/thys/JinjaThreads/Execute/Scheduler.thy @@ -1,1260 +1,1260 @@ (* Title: JinjaThreads/Execute/Scheduler.thy Author: Andreas Lochbihler *) section \Abstract scheduler\ theory Scheduler imports State_Refinement "../Framework/FWProgressAux" "../Framework/FWLTS" (*"../../Collections/spec/SetSpec" "../../Collections/spec/MapSpec" "../../Collections/spec/ListSpec"*) "../Basic/JT_ICF" begin text \ Define an unfold operation that puts everything into one function to avoid duplicate evaluation. \ definition unfold_tllist' :: "('a \ 'b \ 'a + 'c) \ 'a \ ('b, 'c) tllist" where [code del]: "unfold_tllist' f = unfold_tllist (\a. \c. f a = Inr c) (projr \ f) (fst \ projl \ f) (snd \ projl \ f)" lemma unfold_tllist' [code]: "unfold_tllist' f a = (case f a of Inr c \ TNil c | Inl (b, a') \ TCons b (unfold_tllist' f a'))" by(rule tllist.expand)(auto simp add: unfold_tllist'_def split: sum.split_asm) type_synonym ('l,'t,'x,'m,'w,'o,'m_t,'m_w,'s_i,'s) scheduler = "'s \ ('l,'t,'m,'m_t,'m_w,'s_i) state_refine \ ('t \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) option \ 's) option" locale scheduler_spec_base = state_refine_base final r convert_RA thr_\ thr_invar ws_\ ws_invar is_\ is_invar for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and schedule :: "('l,'t,'x,'m,'w,'o,'m_t,'m_w,'s_i,'s) scheduler" and \_invar :: "'s \ 't set \ bool" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" locale scheduler_spec = scheduler_spec_base final r convert_RA schedule \_invar thr_\ thr_invar ws_\ ws_invar is_\ is_invar for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and schedule :: "('l,'t,'x,'m,'w,'o,'m_t,'m_w,'s_i,'s) scheduler" and \_invar :: "'s \ 't set \ bool" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" + fixes invariant :: "('l,'t,'x,'m,'w) state set" assumes schedule_NoneD: "\ schedule \ s = None; state_invar s; \_invar \ (dom (thr_\ (thr s))); state_\ s \ invariant \ \ \.active_threads (state_\ s) = {}" and schedule_Some_NoneD: "\ schedule \ s = \(t, None, \')\; state_invar s; \_invar \ (dom (thr_\ (thr s))); state_\ s \ invariant \ \ \x ln n. thr_\ (thr s) t = \(x, ln)\ \ ln $ n > 0 \ \ waiting (ws_\ (wset s) t) \ may_acquire_all (locks s) t ln" and schedule_Some_SomeD: "\ schedule \ s = \(t, \(ta, x', m')\, \')\; state_invar s; \_invar \ (dom (thr_\ (thr s))); state_\ s \ invariant \ \ \x. thr_\ (thr s) t = \(x, no_wait_locks)\ \ Predicate.eval (r t (x, shr s)) (ta, x', m') \ \.actions_ok (state_\ s) t ta" and schedule_invar_None: "\ schedule \ s = \(t, None, \')\; state_invar s; \_invar \ (dom (thr_\ (thr s))); state_\ s \ invariant \ \ \_invar \' (dom (thr_\ (thr s)))" and schedule_invar_Some: "\ schedule \ s = \(t, \(ta, x', m')\, \')\; state_invar s; \_invar \ (dom (thr_\ (thr s))); state_\ s \ invariant \ \ \_invar \' (dom (thr_\ (thr s)) \ {t. \x m. NewThread t x m \ set \ta\\<^bsub>t\<^esub>})" locale pick_wakeup_spec_base = state_refine_base final r convert_RA thr_\ thr_invar ws_\ ws_invar is_\ is_invar for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and pick_wakeup :: "'s \ 't \ 'w \ 'm_w \ 't option" and \_invar :: "'s \ 't set \ bool" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" locale pick_wakeup_spec = pick_wakeup_spec_base final r convert_RA pick_wakeup \_invar thr_\ thr_invar ws_\ ws_invar is_\ is_invar for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and pick_wakeup :: "'s \ 't \ 'w \ 'm_w \ 't option" and \_invar :: "'s \ 't set \ bool" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" + assumes pick_wakeup_NoneD: "\ pick_wakeup \ t w ws = None; ws_invar ws; \_invar \ T; dom (ws_\ ws) \ T; t \ T \ \ InWS w \ ran (ws_\ ws)" and pick_wakeup_SomeD: "\ pick_wakeup \ t w ws = \t'\; ws_invar ws; \_invar \ T; dom (ws_\ ws) \ T; t \ T \ \ ws_\ ws t' = \InWS w\" locale scheduler_base_aux = state_refine_base final r convert_RA thr_\ thr_invar ws_\ ws_invar is_\ is_invar for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and thr_lookup :: "'t \ 'm_t \ ('x \ 'l released_locks)" and thr_update :: "'t \ 'x \ 'l released_locks \ 'm_t \ 'm_t" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and ws_lookup :: "'t \ 'm_w \ 'w wait_set_status" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" and is_memb :: "'t \ 's_i \ bool" and is_ins :: "'t \ 's_i \ 's_i" and is_delete :: "'t \ 's_i \ 's_i" begin definition free_thread_id :: "'m_t \ 't \ bool" where "free_thread_id ts t \ thr_lookup t ts = None" fun redT_updT :: "'m_t \ ('t,'x,'m) new_thread_action \ 'm_t" where "redT_updT ts (NewThread t' x m) = thr_update t' (x, no_wait_locks) ts" | "redT_updT ts _ = ts" definition redT_updTs :: "'m_t \ ('t,'x,'m) new_thread_action list \ 'm_t" where "redT_updTs = foldl redT_updT" primrec thread_ok :: "'m_t \ ('t,'x,'m) new_thread_action \ bool" where "thread_ok ts (NewThread t x m) = free_thread_id ts t" | "thread_ok ts (ThreadExists t b) = (b \ free_thread_id ts t)" text \ We use @{term "redT_updT"} in \thread_ok\ instead of @{term "redT_updT'"} like in theory @{theory JinjaThreads.FWThread}. This fixes @{typ "'x"} in the @{typ "('t,'x,'m) new_thread_action list"} type, but avoids @{term "undefined"}, which raises an exception during execution in the generated code. \ primrec thread_oks :: "'m_t \ ('t,'x,'m) new_thread_action list \ bool" where "thread_oks ts [] = True" | "thread_oks ts (ta#tas) = (thread_ok ts ta \ thread_oks (redT_updT ts ta) tas)" definition wset_actions_ok :: "'m_w \ 't \ ('t,'w) wait_set_action list \ bool" where "wset_actions_ok ws t was \ ws_lookup t ws = (if Notified \ set was then \PostWS WSNotified\ else if WokenUp \ set was then \PostWS WSWokenUp\ else None)" primrec cond_action_ok :: "('l,'t,'m,'m_t,'m_w,'s_i) state_refine \ 't \ 't conditional_action \ bool" where "\ln. cond_action_ok s t (Join T) = (case thr_lookup T (thr s) of None \ True | \(x, ln)\ \ t \ T \ final x \ ln = no_wait_locks \ ws_lookup T (wset s) = None)" | "cond_action_ok s t Yield = True" definition cond_action_oks :: "('l,'t,'m,'m_t,'m_w,'s_i) state_refine \ 't \ 't conditional_action list \ bool" where "cond_action_oks s t cts = list_all (cond_action_ok s t) cts" primrec redT_updI :: "'s_i \ 't interrupt_action \ 's_i" where "redT_updI is (Interrupt t) = is_ins t is" | "redT_updI is (ClearInterrupt t) = is_delete t is" | "redT_updI is (IsInterrupted t b) = is" primrec redT_updIs :: "'s_i \ 't interrupt_action list \ 's_i" where "redT_updIs is [] = is" | "redT_updIs is (ia # ias) = redT_updIs (redT_updI is ia) ias" primrec interrupt_action_ok :: "'s_i \ 't interrupt_action \ bool" where "interrupt_action_ok is (Interrupt t) = True" | "interrupt_action_ok is (ClearInterrupt t) = True" | "interrupt_action_ok is (IsInterrupted t b) = (b = (is_memb t is))" primrec interrupt_actions_ok :: "'s_i \ 't interrupt_action list \ bool" where "interrupt_actions_ok is [] = True" | "interrupt_actions_ok is (ia # ias) \ interrupt_action_ok is ia \ interrupt_actions_ok (redT_updI is ia) ias" definition actions_ok :: "('l,'t,'m,'m_t,'m_w,'s_i) state_refine \ 't \ ('l,'t,'x,'m,'w,'o') thread_action \ bool" where "actions_ok s t ta \ lock_ok_las (locks s) t \ta\\<^bsub>l\<^esub> \ thread_oks (thr s) \ta\\<^bsub>t\<^esub> \ cond_action_oks s t \ta\\<^bsub>c\<^esub> \ wset_actions_ok (wset s) t \ta\\<^bsub>w\<^esub> \ interrupt_actions_ok (interrupts s) \ta\\<^bsub>i\<^esub>" end locale scheduler_base = scheduler_base_aux final r convert_RA thr_\ thr_invar thr_lookup thr_update ws_\ ws_invar ws_lookup is_\ is_invar is_memb is_ins is_delete + scheduler_spec_base final r convert_RA schedule \_invar thr_\ thr_invar ws_\ ws_invar is_\ is_invar + pick_wakeup_spec_base final r convert_RA pick_wakeup \_invar thr_\ thr_invar ws_\ ws_invar is_\ is_invar for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and schedule :: "('l,'t,'x,'m,'w,'o,'m_t,'m_w,'s_i,'s) scheduler" and "output" :: "'s \ 't \ ('l,'t,'x,'m,'w,'o) thread_action \ 'q option" and pick_wakeup :: "'s \ 't \ 'w \ 'm_w \ 't option" and \_invar :: "'s \ 't set \ bool" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and thr_lookup :: "'t \ 'm_t \ ('x \ 'l released_locks)" and thr_update :: "'t \ 'x \ 'l released_locks \ 'm_t \ 'm_t" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and ws_lookup :: "'t \ 'm_w \ 'w wait_set_status" and ws_update :: "'t \ 'w wait_set_status \ 'm_w \ 'm_w" and ws_delete :: "'t \ 'm_w \ 'm_w" and ws_iterate :: "'m_w \ ('t \ 'w wait_set_status, 'm_w) set_iterator" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" and is_memb :: "'t \ 's_i \ bool" and is_ins :: "'t \ 's_i \ 's_i" and is_delete :: "'t \ 's_i \ 's_i" begin primrec exec_updW :: "'s \ 't \ 'm_w \ ('t,'w) wait_set_action \ 'm_w" where "exec_updW \ t ws (Notify w) = (case pick_wakeup \ t w ws of None \ ws | Some t \ ws_update t (PostWS WSNotified) ws)" | "exec_updW \ t ws (NotifyAll w) = ws_iterate ws (\_. True) (\(t, w') ws'. if w' = InWS w then ws_update t (PostWS WSNotified) ws' else ws') ws" | "exec_updW \ t ws (Suspend w) = ws_update t (InWS w) ws" | "exec_updW \ t ws (WakeUp t') = (case ws_lookup t' ws of \InWS w\ \ ws_update t' (PostWS WSWokenUp) ws | _ \ ws)" | "exec_updW \ t ws Notified = ws_delete t ws" | "exec_updW \ t ws WokenUp = ws_delete t ws" definition exec_updWs :: "'s \ 't \ 'm_w \ ('t,'w) wait_set_action list \ 'm_w" where "exec_updWs \ t = foldl (exec_updW \ t)" definition exec_upd :: "'s \ ('l,'t,'m,'m_t,'m_w,'s_i) state_refine \ 't \ ('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm \ ('l,'t,'m,'m_t,'m_w,'s_i) state_refine" where [simp]: "exec_upd \ s t ta x' m' = (redT_updLs (locks s) t \ta\\<^bsub>l\<^esub>, (thr_update t (x', redT_updLns (locks s) t (snd (the (thr_lookup t (thr s)))) \ta\\<^bsub>l\<^esub>) (redT_updTs (thr s) \ta\\<^bsub>t\<^esub>), m'), exec_updWs \ t (wset s) \ta\\<^bsub>w\<^esub>, redT_updIs (interrupts s) \ta\\<^bsub>i\<^esub>)" definition execT :: "'s \ ('l,'t,'m,'m_t,'m_w,'s_i) state_refine \ ('s \ 't \ ('l,'t,'x,'m,'w,'o) thread_action \ ('l,'t,'m,'m_t,'m_w,'s_i) state_refine) option" where "execT \ s = (do { (t, tax'm', \') \ schedule \ s; case tax'm' of None \ (let (x, ln) = the (thr_lookup t (thr s)); ta = (K$ [], [], [], [], [], convert_RA ln); s' = (acquire_all (locks s) t ln, (thr_update t (x, no_wait_locks) (thr s), shr s), wset s, interrupts s) in \(\', t, ta, s')\) | \(ta, x', m')\ \ \(\', t, ta, exec_upd \ s t ta x' m')\ })" primrec exec_step :: "'s \ ('l,'t,'m,'m_t,'m_w,'s_i) state_refine \ ('s \ 't \ ('l,'t,'x,'m,'w,'o) thread_action) \ 's \ ('l,'t,'m,'m_t,'m_w,'s_i) state_refine + ('l,'t,'m,'m_t,'m_w,'s_i) state_refine" where "exec_step (\, s) = (case execT \ s of None \ Inr s | Some (\', t, ta, s') \ Inl ((\, t, ta), \', s'))" declare exec_step.simps [simp del] definition exec_aux :: "'s \ ('l,'t,'m,'m_t,'m_w,'s_i) state_refine \ ('s \ 't \ ('l,'t,'x,'m,'w,'o) thread_action, ('l,'t,'m,'m_t,'m_w,'s_i) state_refine) tllist" where "exec_aux \s = unfold_tllist' exec_step \s" definition exec :: "'s \ ('l,'t,'m,'m_t,'m_w,'s_i) state_refine \ ('q, ('l,'t,'m,'m_t,'m_w,'s_i) state_refine) tllist" where "exec \ s = tmap the id (tfilter undefined (\q. q \ None) (tmap (\(\, t, ta). output \ t ta) id (exec_aux (\, s))))" end text \ Implement \pick_wakeup\ by \map_sel'\ \ definition pick_wakeup_via_sel :: "('m_w \ ('t \ 'w wait_set_status \ bool) \ 't \ 'w wait_set_status) \ 's \ 't \ 'w \ 'm_w \ 't option" where "pick_wakeup_via_sel ws_sel \ t w ws = map_option fst (ws_sel ws (\t w'. w' = InWS w))" lemma pick_wakeup_spec_via_sel: assumes sel: "map_sel' ws_\ ws_invar ws_sel" shows "pick_wakeup_spec (pick_wakeup_via_sel (\s P. ws_sel s (\(k,v). P k v))) \_invar ws_\ ws_invar" proof - interpret ws: map_sel' ws_\ ws_invar ws_sel by(rule sel) show ?thesis by(unfold_locales)(auto simp add: pick_wakeup_via_sel_def ran_def dest: ws.sel'_noneD ws.sel'_SomeD) qed locale scheduler_ext_base = scheduler_base_aux final r convert_RA thr_\ thr_invar thr_lookup thr_update ws_\ ws_invar ws_lookup is_\ is_invar is_memb is_ins is_delete for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and thr_lookup :: "'t \ 'm_t \ ('x \ 'l released_locks)" and thr_update :: "'t \ 'x \ 'l released_locks \ 'm_t \ 'm_t" and thr_iterate :: "'m_t \ ('t \ ('x \ 'l released_locks), 's_t) set_iterator" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and ws_lookup :: "'t \ 'm_w \ 'w wait_set_status" and ws_update :: "'t \ 'w wait_set_status \ 'm_w \ 'm_w" and ws_sel :: "'m_w \ ('t \ 'w wait_set_status \ bool) \ ('t \ 'w wait_set_status)" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" and is_memb :: "'t \ 's_i \ bool" and is_ins :: "'t \ 's_i \ 's_i" and is_delete :: "'t \ 's_i \ 's_i" + fixes thr'_\ :: "'s_t \ 't set" and thr'_invar :: "'s_t \ bool" and thr'_empty :: "unit \ 's_t" and thr'_ins_dj :: "'t \ 's_t \ 's_t" begin abbreviation pick_wakeup :: "'s \ 't \ 'w \ 'm_w \ 't option" where "pick_wakeup \ pick_wakeup_via_sel (\s P. ws_sel s (\(k,v). P k v))" fun active_threads :: "('l,'t,'m,'m_t,'m_w,'s_i) state_refine \ 's_t" where "active_threads (ls, (ts, m), ws, is) = thr_iterate ts (\_. True) (\(t, (x, ln)) ts'. if ln = no_wait_locks then if Predicate.holds (do { (ta, _) \ r t (x, m); Predicate.if_pred (actions_ok (ls, (ts, m), ws, is) t ta) }) then thr'_ins_dj t ts' else ts' else if \ waiting (ws_lookup t ws) \ may_acquire_all ls t ln then thr'_ins_dj t ts' else ts') (thr'_empty ())" end locale scheduler_aux = scheduler_base_aux final r convert_RA thr_\ thr_invar thr_lookup thr_update ws_\ ws_invar ws_lookup is_\ is_invar is_memb is_ins is_delete + thr: finite_map thr_\ thr_invar + thr: map_lookup thr_\ thr_invar thr_lookup + thr: map_update thr_\ thr_invar thr_update + ws: map ws_\ ws_invar + ws: map_lookup ws_\ ws_invar ws_lookup + "is": set is_\ is_invar + "is": set_memb is_\ is_invar is_memb + "is": set_ins is_\ is_invar is_ins + "is": set_delete is_\ is_invar is_delete for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and thr_lookup :: "'t \ 'm_t \ ('x \ 'l released_locks)" and thr_update :: "'t \ 'x \ 'l released_locks \ 'm_t \ 'm_t" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and ws_lookup :: "'t \ 'm_w \ 'w wait_set_status" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" and is_memb :: "'t \ 's_i \ bool" and is_ins :: "'t \ 's_i \ 's_i" and is_delete :: "'t \ 's_i \ 's_i" begin lemma free_thread_id_correct [simp]: "thr_invar ts \ free_thread_id ts = FWThread.free_thread_id (thr_\ ts)" by(auto simp add: free_thread_id_def fun_eq_iff thr.lookup_correct intro: free_thread_id.intros) lemma redT_updT_correct [simp]: assumes "thr_invar ts" shows "thr_\ (redT_updT ts nta) = FWThread.redT_updT (thr_\ ts) nta" and "thr_invar (redT_updT ts nta)" by(case_tac [!] nta)(simp_all add: thr.update_correct assms) lemma redT_updTs_correct [simp]: assumes "thr_invar ts" shows "thr_\ (redT_updTs ts ntas) = FWThread.redT_updTs (thr_\ ts) ntas" and "thr_invar (redT_updTs ts ntas)" using assms by(induct ntas arbitrary: ts)(simp_all add: redT_updTs_def) lemma thread_ok_correct [simp]: "thr_invar ts \ thread_ok ts nta \ FWThread.thread_ok (thr_\ ts) nta" by(cases nta) simp_all lemma thread_oks_correct [simp]: "thr_invar ts \ thread_oks ts ntas \ FWThread.thread_oks (thr_\ ts) ntas" by(induct ntas arbitrary: ts) simp_all lemma wset_actions_ok_correct [simp]: "ws_invar ws \ wset_actions_ok ws t was \ FWWait.wset_actions_ok (ws_\ ws) t was" by(simp add: wset_actions_ok_def FWWait.wset_actions_ok_def ws.lookup_correct) lemma cond_action_ok_correct [simp]: "state_invar s \ cond_action_ok s t cta \ \.cond_action_ok (state_\ s) t cta" by(cases s,cases cta)(auto simp add: thr.lookup_correct ws.lookup_correct) lemma cond_action_oks_correct [simp]: assumes "state_invar s" shows "cond_action_oks s t ctas \ \.cond_action_oks (state_\ s) t ctas" by(induct ctas)(simp_all add: cond_action_oks_def assms) lemma redT_updI_correct [simp]: assumes "is_invar is" shows "is_\ (redT_updI is ia) = FWInterrupt.redT_updI (is_\ is) ia" and "is_invar (redT_updI is ia)" using assms by(case_tac [!] ia)(auto simp add: is.ins_correct is.delete_correct) lemma redT_updIs_correct [simp]: assumes "is_invar is" shows "is_\ (redT_updIs is ias) = FWInterrupt.redT_updIs (is_\ is) ias" and "is_invar (redT_updIs is ias)" using assms by(induct ias arbitrary: "is")(auto) lemma interrupt_action_ok_correct [simp]: "is_invar is \ interrupt_action_ok is ia \ FWInterrupt.interrupt_action_ok (is_\ is) ia" by(cases ia)(auto simp add: is.memb_correct) lemma interrupt_actions_ok_correct [simp]: "is_invar is \ interrupt_actions_ok is ias \ FWInterrupt.interrupt_actions_ok (is_\ is) ias" by(induct ias arbitrary:"is") simp_all lemma actions_ok_correct [simp]: "state_invar s \ actions_ok s t ta \ \.actions_ok (state_\ s) t ta" by(auto simp add: actions_ok_def) end locale scheduler = scheduler_base final r convert_RA schedule "output" pick_wakeup \_invar thr_\ thr_invar thr_lookup thr_update ws_\ ws_invar ws_lookup ws_update ws_delete ws_iterate is_\ is_invar is_memb is_ins is_delete + scheduler_aux final r convert_RA thr_\ thr_invar thr_lookup thr_update ws_\ ws_invar ws_lookup is_\ is_invar is_memb is_ins is_delete + scheduler_spec final r convert_RA schedule \_invar thr_\ thr_invar ws_\ ws_invar is_\ is_invar invariant + pick_wakeup_spec final r convert_RA pick_wakeup \_invar thr_\ thr_invar ws_\ ws_invar is_\ is_invar + ws: map_update ws_\ ws_invar ws_update + ws: map_delete ws_\ ws_invar ws_delete + ws: map_iteratei ws_\ ws_invar ws_iterate for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and schedule :: "('l,'t,'x,'m,'w,'o,'m_t,'m_w,'s_i,'s) scheduler" and "output" :: "'s \ 't \ ('l,'t,'x,'m,'w,'o) thread_action \ 'q option" and pick_wakeup :: "'s \ 't \ 'w \ 'm_w \ 't option" and \_invar :: "'s \ 't set \ bool" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and thr_lookup :: "'t \ 'm_t \ ('x \ 'l released_locks)" and thr_update :: "'t \ 'x \ 'l released_locks \ 'm_t \ 'm_t" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and ws_lookup :: "'t \ 'm_w \ 'w wait_set_status" and ws_update :: "'t \ 'w wait_set_status \ 'm_w \ 'm_w" and ws_delete :: "'t \ 'm_w \ 'm_w" and ws_iterate :: "'m_w \ ('t \ 'w wait_set_status, 'm_w) set_iterator" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" and is_memb :: "'t \ 's_i \ bool" and is_ins :: "'t \ 's_i \ 's_i" and is_delete :: "'t \ 's_i \ 's_i" and invariant :: "('l,'t,'x,'m,'w) state set" + assumes invariant: "invariant3p \.redT invariant" begin lemma exec_updW_correct: assumes invar: "ws_invar ws" "\_invar \ T" "dom (ws_\ ws) \ T" "t \ T" shows "redT_updW t (ws_\ ws) wa (ws_\ (exec_updW \ t ws wa))" (is "?thesis1") and "ws_invar (exec_updW \ t ws wa)" (is "?thesis2") proof - from invar have "?thesis1 \ ?thesis2" proof(cases wa) case [simp]: (Notify w) show ?thesis proof(cases "pick_wakeup \ t w ws") case (Some t') hence "ws_\ ws t' = \InWS w\" using invar by(rule pick_wakeup_SomeD) with Some show ?thesis using invar by(auto simp add: ws.update_correct) next case None hence "InWS w \ ran (ws_\ ws)" using invar by(rule pick_wakeup_NoneD) with None show ?thesis using invar by(auto simp add: ran_def) qed next case [simp]: (NotifyAll w) let ?f = "\(t, w') ws'. if w' = InWS w then ws_update t (PostWS WSNotified) ws' else ws'" let ?I = "\T ws'. (\k. if k\T \ ws_\ ws k = \InWS w\ then ws_\ ws' k = \PostWS WSNotified\ else ws_\ ws' k = ws_\ ws k) \ ws_invar ws'" from invar have "?I (dom (ws_\ ws)) ws" by(auto simp add: ws.lookup_correct) with \ws_invar ws\ have "?I {} (ws_iterate ws (\_. True) ?f ws)" proof(rule ws.iterate_rule_P[where I="?I"]) fix t w' T ws' assume t: "t \ T" and w': "ws_\ ws t = \w'\" and T: "T \ dom (ws_\ ws)" and I: "?I T ws'" { fix t' assume "t' \ T - {t}" "ws_\ ws t' = \InWS w\" with t I w' invar have "ws_\ (?f (t, w') ws') t' = \PostWS WSNotified\" by(auto)(simp_all add: ws.update_correct) } moreover { fix t' assume "t' \ T - {t} \ ws_\ ws t' \ \InWS w\" with t I w' invar have "ws_\ (?f (t,w') ws') t' = ws_\ ws t'" by(auto simp add: ws.update_correct) } moreover have "ws_invar (?f (t, w') ws')" using I by(simp add: ws.update_correct) ultimately show "?I (T - {t}) (?f (t, w') ws')" by safe simp qed hence "ws_\ (ws_iterate ws (\_. True) ?f ws) = (\t. if ws_\ ws t = \InWS w\ then \PostWS WSNotified\ else ws_\ ws t)" and "ws_invar (ws_iterate ws (\_. True) ?f ws)" by(simp_all add: fun_eq_iff) thus ?thesis by simp next case WakeUp thus ?thesis using assms by(auto simp add: ws.lookup_correct ws.update_correct split: wait_set_status.split) qed(simp_all add: ws.update_correct ws.delete_correct map_upd_eq_restrict) thus ?thesis1 ?thesis2 by simp_all qed lemma exec_updWs_correct: assumes "ws_invar ws" "\_invar \ T" "dom (ws_\ ws) \ T" "t \ T" shows "redT_updWs t (ws_\ ws) was (ws_\ (exec_updWs \ t ws was))" (is "?thesis1") and "ws_invar (exec_updWs \ t ws was)" (is "?thesis2") proof - from \ws_invar ws\ \dom (ws_\ ws) \ T\ have "?thesis1 \ ?thesis2" proof(induct was arbitrary: ws) case Nil thus ?case by(auto simp add: exec_updWs_def redT_updWs_def) next case (Cons wa was) let ?ws' = "exec_updW \ t ws wa" from \ws_invar ws\ \\_invar \ T\ \dom (ws_\ ws) \ T\ \t \ T\ have invar': "ws_invar ?ws'" and red: "redT_updW t (ws_\ ws) wa (ws_\ ?ws')" by(rule exec_updW_correct)+ have "dom (ws_\ ?ws') \ T" proof fix t' assume "t' \ dom (ws_\ ?ws')" with red have "t' \ dom (ws_\ ws) \ t = t'" by(auto dest!: redT_updW_Some_otherD split: wait_set_status.split_asm) with \dom (ws_\ ws) \ T\ \t \ T\ show "t' \ T" by auto qed with invar' have "redT_updWs t (ws_\ ?ws') was (ws_\ (exec_updWs \ t ?ws' was)) \ ws_invar (exec_updWs \ t ?ws' was)" by(rule Cons.hyps) thus ?case using red by(auto simp add: exec_updWs_def redT_updWs_def intro: rtrancl3p_step_converse) qed thus ?thesis1 ?thesis2 by simp_all qed lemma exec_upd_correct: assumes "state_invar s" "\_invar \ (dom (thr_\ (thr s)))" "t \ (dom (thr_\ (thr s)))" and "wset_thread_ok (ws_\ (wset s)) (thr_\ (thr s))" shows "redT_upd (state_\ s) t ta x' m' (state_\ (exec_upd \ s t ta x' m'))" and "state_invar (exec_upd \ s t ta x' m')" using assms unfolding wset_thread_ok_conv_dom by(auto simp add: thr.update_correct thr.lookup_correct intro: exec_updWs_correct) lemma execT_None: assumes invar: "state_invar s" "\_invar \ (dom (thr_\ (thr s)))" "state_\ s \ invariant" and exec: "execT \ s = None" shows "\.active_threads (state_\ s) = {}" using assms by(cases "schedule \ s")(fastforce simp add: execT_def thr.lookup_correct dest: schedule_Some_NoneD schedule_NoneD)+ lemma execT_Some: assumes invar: "state_invar s" "\_invar \ (dom (thr_\ (thr s)))" "state_\ s \ invariant" and wstok: "wset_thread_ok (ws_\ (wset s)) (thr_\ (thr s))" and exec: "execT \ s = \(\', t, ta, s')\" shows "\.redT (state_\ s) (t, ta) (state_\ s')" (is "?thesis1") and "state_invar s'" (is "?thesis2") and "\_invar \' (dom (thr_\ (thr s')))" (is "?thesis3") proof - note [simp del] = redT_upd_simps exec_upd_def have "?thesis1 \ ?thesis2 \ ?thesis3" proof(cases "fst (snd (the (schedule \ s)))") case None with exec invar have schedule: "schedule \ s = \(t, None, \')\" and ta: "ta = (K$ [], [], [], [], [], convert_RA (snd (the (thr_\ (thr s) t))))" and s': "s' = (acquire_all (locks s) t (snd (the (thr_\ (thr s) t))), (thr_update t (fst (the (thr_\ (thr s) t)), no_wait_locks) (thr s), shr s), wset s, interrupts s)" - by(auto simp add: execT_def Option_bind_eq_Some_conv thr.lookup_correct split_beta split del: option.split_asm) + by(auto simp add: execT_def bind_eq_Some_conv thr.lookup_correct split_beta split del: option.split_asm) from schedule_Some_NoneD[OF schedule invar] obtain x ln n where t: "thr_\ (thr s) t = \(x, ln)\" and "0 < ln $ n" "\ waiting (ws_\ (wset s) t)" "may_acquire_all (locks s) t ln" by blast hence ?thesis1 using ta s' invar by(auto intro: \.redT.redT_acquire simp add: thr.update_correct) moreover from invar s' have "?thesis2" by(simp add: thr.update_correct) moreover from t s' invar have "dom (thr_\ (thr s')) = dom (thr_\ (thr s))" by(auto simp add: thr.update_correct) hence "?thesis3" using invar schedule by(auto intro: schedule_invar_None) ultimately show ?thesis by simp next case (Some taxm) with exec invar obtain x' m' where schedule: "schedule \ s = \(t, \(ta, x', m')\, \')\" and s': "s' = exec_upd \ s t ta x' m'" - by(cases taxm)(fastforce simp add: execT_def Option_bind_eq_Some_conv split del: option.split_asm) + by(cases taxm)(fastforce simp add: execT_def bind_eq_Some_conv split del: option.split_asm) from schedule_Some_SomeD[OF schedule invar] obtain x where t: "thr_\ (thr s) t = \(x, no_wait_locks)\" and "Predicate.eval (r t (x, shr s)) (ta, x', m')" and aok: "\.actions_ok (state_\ s) t ta" by blast with s' have ?thesis1 using invar wstok by(fastforce intro: \.redT.intros exec_upd_correct) moreover from invar s' t wstok have ?thesis2 by(auto intro: exec_upd_correct) moreover { from schedule invar have "\_invar \' (dom (thr_\ (thr s)) \ {t. \x m. NewThread t x m \ set \ta\\<^bsub>t\<^esub>})" by(rule schedule_invar_Some) also have "dom (thr_\ (thr s)) \ {t. \x m. NewThread t x m \ set \ta\\<^bsub>t\<^esub>} = dom (thr_\ (thr s'))" using invar s' aok t by(auto simp add: exec_upd_def thr.lookup_correct thr.update_correct simp del: split_paired_Ex)(fastforce dest: redT_updTs_new_thread intro: redT_updTs_Some1 redT_updTs_new_thread_ts simp del: split_paired_Ex)+ finally have "\_invar \' (dom (thr_\ (thr s')))" . } ultimately show ?thesis by simp qed thus ?thesis1 ?thesis2 ?thesis3 by simp_all qed lemma exec_step_into_redT: assumes invar: "state_invar s" "\_invar \ (dom (thr_\ (thr s)))" "state_\ s \ invariant" and wstok: "wset_thread_ok (ws_\ (wset s)) (thr_\ (thr s))" and exec: "exec_step (\, s) = Inl ((\'', t, ta), \', s')" shows "\.redT (state_\ s) (t, ta) (state_\ s')" "\'' = \" and "state_invar s'" "\_invar \' (dom (thr_\ (thr s')))" "state_\ s' \ invariant" proof - from exec have execT: "execT \ s = \(\', t, ta, s')\" and q: "\'' = \" by(auto simp add: exec_step.simps split_beta) from invar wstok execT show red: "\.redT (state_\ s) (t, ta) (state_\ s')" and invar': "state_invar s'" "\_invar \' (dom (thr_\ (thr s')))" "\'' = \" by(rule execT_Some)+(rule q) from invariant red \state_\ s \ invariant\ show "state_\ s' \ invariant" by(rule invariant3pD) qed lemma exec_step_InrD: assumes "state_invar s" "\_invar \ (dom (thr_\ (thr s)))" "state_\ s \ invariant" and "exec_step (\, s) = Inr s'" shows "\.active_threads (state_\ s) = {}" and "s' = s" using assms by(auto simp add: exec_step_def dest: execT_None) lemma (in multithreaded_base) red_in_active_threads: assumes "s -t\ta\ s'" shows "t \ active_threads s" using assms by cases(auto intro: active_threads.intros) lemma exec_aux_into_Runs: assumes "state_invar s" "\_invar \ (dom (thr_\ (thr s)))" "state_\ s \ invariant" and "wset_thread_ok (ws_\ (wset s)) (thr_\ (thr s))" shows "\.mthr.Runs (state_\ s) (lmap snd (llist_of_tllist (exec_aux (\, s))))" (is ?thesis1) and "tfinite (exec_aux (\, s)) \ state_invar (terminal (exec_aux (\, s)))" (is "_ \ ?thesis2") proof - from assms show ?thesis1 proof(coinduction arbitrary: \ s) case (Runs \ s) note invar = \state_invar s\ \\_invar \ (dom (thr_\ (thr s)))\ \state_\ s \ invariant\ and wstok = \wset_thread_ok (ws_\ (wset s)) (thr_\ (thr s))\ show ?case proof(cases "exec_aux (\, s)") case (TNil s') hence "\.active_threads (state_\ s) = {}" "s' = s" by(auto simp add: exec_aux_def unfold_tllist' split: sum.split_asm dest: exec_step_InrD[OF invar]) hence ?Stuck using TNil by(auto dest: \.red_in_active_threads) thus ?thesis .. next case (TCons \tta ttls') then obtain t ta \' s' \'' where [simp]: "\tta = (\'', t, ta)" and [simp]: "ttls' = exec_aux (\', s')" and step: "exec_step (\, s) = Inl ((\'', t, ta), \', s')" unfolding exec_aux_def by(subst (asm) (2) unfold_tllist')(fastforce split: sum.split_asm) from invar wstok step have redT: "\.redT (state_\ s) (t, ta) (state_\ s')" and [simp]: "\'' = \" and invar': "state_invar s'" "\_invar \' (dom (thr_\ (thr s')))" "state_\ s' \ invariant" by(rule exec_step_into_redT)+ from wstok \.redT_preserves_wset_thread_ok[OF redT] have "wset_thread_ok (ws_\ (wset s')) (thr_\ (thr s'))" by simp with invar' redT TCons have ?Step by(auto simp del: split_paired_Ex) thus ?thesis .. qed qed next assume "tfinite (exec_aux (\, s))" thus "?thesis2" using assms proof(induct "exec_aux (\, s)" arbitrary: \ s rule: tfinite_induct) case TNil thus ?case by(auto simp add: exec_aux_def unfold_tllist' split_beta split: sum.split_asm dest: exec_step_InrD) next case (TCons \tta ttls) from \TCons \tta ttls = exec_aux (\, s)\ obtain \'' t ta \' s' where [simp]: "\tta = (\'', t, ta)" and ttls: "ttls = exec_aux (\', s')" and step: "exec_step (\, s) = Inl ((\'', t, ta), \', s')" unfolding exec_aux_def by(subst (asm) (2) unfold_tllist')(fastforce split: sum.split_asm) note ttls moreover from \state_invar s\ \\_invar \ (dom (thr_\ (thr s)))\ \state_\ s \ invariant\ \wset_thread_ok (ws_\ (wset s)) (thr_\ (thr s))\ step have [simp]: "\'' = \" and invar': "state_invar s'" "\_invar \' (dom (thr_\ (thr s')))" "state_\ s' \ invariant" and redT: "\.redT (state_\ s) (t, ta) (state_\ s')" by(rule exec_step_into_redT)+ note invar' moreover from \.redT_preserves_wset_thread_ok[OF redT] \wset_thread_ok (ws_\ (wset s)) (thr_\ (thr s))\ have "wset_thread_ok (ws_\ (wset s')) (thr_\ (thr s'))" by simp ultimately have "state_invar (terminal (exec_aux (\', s')))" by(rule TCons) with \TCons \tta ttls = exec_aux (\, s)\[symmetric] show ?case unfolding ttls by simp qed qed end locale scheduler_ext_aux = scheduler_ext_base final r convert_RA thr_\ thr_invar thr_lookup thr_update thr_iterate ws_\ ws_invar ws_lookup ws_update ws_sel is_\ is_invar is_memb is_ins is_delete thr'_\ thr'_invar thr'_empty thr'_ins_dj + scheduler_aux final r convert_RA thr_\ thr_invar thr_lookup thr_update ws_\ ws_invar ws_lookup is_\ is_invar is_memb is_ins is_delete + thr: map_iteratei thr_\ thr_invar thr_iterate + ws: map_update ws_\ ws_invar ws_update + ws: map_sel' ws_\ ws_invar ws_sel + thr': finite_set thr'_\ thr'_invar + thr': set_empty thr'_\ thr'_invar thr'_empty + thr': set_ins_dj thr'_\ thr'_invar thr'_ins_dj for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and thr_lookup :: "'t \ 'm_t \ ('x \ 'l released_locks)" and thr_update :: "'t \ 'x \ 'l released_locks \ 'm_t \ 'm_t" and thr_iterate :: "'m_t \ ('t \ ('x \ 'l released_locks), 's_t) set_iterator" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and ws_lookup :: "'t \ 'm_w \ 'w wait_set_status" and ws_update :: "'t \ 'w wait_set_status \ 'm_w \ 'm_w" and ws_sel :: "'m_w \ (('t \ 'w wait_set_status) \ bool) \ ('t \ 'w wait_set_status)" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" and is_memb :: "'t \ 's_i \ bool" and is_ins :: "'t \ 's_i \ 's_i" and is_delete :: "'t \ 's_i \ 's_i" and thr'_\ :: "'s_t \ 't set" and thr'_invar :: "'s_t \ bool" and thr'_empty :: "unit \ 's_t" and thr'_ins_dj :: "'t \ 's_t \ 's_t" begin lemma active_threads_correct [simp]: assumes "state_invar s" shows "thr'_\ (active_threads s) = \.active_threads (state_\ s)" (is "?thesis1") and "thr'_invar (active_threads s)" (is "?thesis2") proof - obtain ls ts m ws "is" where s: "s = (ls, (ts, m), ws, is)" by(cases s) fastforce let ?f = "\(t, (x, ln)) TS. if ln = no_wait_locks then if Predicate.holds (do { (ta, _) \ r t (x, m); Predicate.if_pred (actions_ok (ls, (ts, m), ws, is) t ta) }) then thr'_ins_dj t TS else TS else if \ waiting (ws_lookup t ws) \ may_acquire_all ls t ln then thr'_ins_dj t TS else TS" let ?I = "\T TS. thr'_invar TS \ thr'_\ TS \ dom (thr_\ ts) - T \ (\t. t \ T \ t \ thr'_\ TS \ t \ \.active_threads (state_\ s))" from assms s have "thr_invar ts" by simp moreover have "?I (dom (thr_\ ts)) (thr'_empty ())" by(auto simp add: thr'.empty_correct s elim: \.active_threads.cases) ultimately have "?I {} (thr_iterate ts (\_. True) ?f (thr'_empty ()))" proof(rule thr.iterate_rule_P[where I="?I"]) fix t xln T TS assume tT: "t \ T" and tst: "thr_\ ts t = \xln\" and Tdom: "T \ dom (thr_\ ts)" and I: "?I T TS" obtain x ln where xln: "xln = (x, ln)" by(cases xln) from tT I have t: "t \ thr'_\ TS" by blast from I have invar: "thr'_invar TS" .. hence "thr'_invar (?f (t, xln) TS)" using t unfolding xln by(auto simp add: thr'.ins_dj_correct) moreover from I have "thr'_\ TS \ dom (thr_\ ts) - T" by blast hence "thr'_\ (?f (t, xln) TS) \ dom (thr_\ ts) - (T - {t})" using invar tst t by(auto simp add: xln thr'.ins_dj_correct) moreover { fix t' assume t': "t' \ T - {t}" have "t' \ thr'_\ (?f (t, xln) TS) \ t' \ \.active_threads (state_\ s)" (is "?lhs \ ?rhs") proof(cases "t' = t") case True show ?thesis proof assume ?lhs with True xln invar tst \state_invar s\ t show ?rhs by(fastforce simp add: holds_eq thr'.ins_dj_correct s split_beta ws.lookup_correct split: if_split_asm elim!: bindE if_predE intro: \.active_threads.intros) next assume ?rhs with True xln invar tst \state_invar s\ t show ?lhs by(fastforce elim!: \.active_threads.cases simp add: holds_eq s thr'.ins_dj_correct ws.lookup_correct elim!: bindE if_predE intro: bindI if_predI) qed next case False with t' have "t' \ T" by simp with I have "t' \ thr'_\ TS \ t' \ \.active_threads (state_\ s)" by blast thus ?thesis using xln False invar t by(auto simp add: thr'.ins_dj_correct) qed } ultimately show "?I (T - {t}) (?f (t, xln) TS)" by blast qed thus "?thesis1" "?thesis2" by(auto simp add: s) qed end locale scheduler_ext = scheduler_ext_aux final r convert_RA thr_\ thr_invar thr_lookup thr_update thr_iterate ws_\ ws_invar ws_lookup ws_update ws_sel is_\ is_invar is_memb is_ins is_delete thr'_\ thr'_invar thr'_empty thr'_ins_dj + scheduler_spec final r convert_RA schedule \_invar thr_\ thr_invar ws_\ ws_invar is_\ is_invar invariant + ws: map_delete ws_\ ws_invar ws_delete + ws: map_iteratei ws_\ ws_invar ws_iterate for final :: "'x \ bool" and r :: "'t \ ('x \ 'm) \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) Predicate.pred" and convert_RA :: "'l released_locks \ 'o list" and schedule :: "('l,'t,'x,'m,'w,'o,'m_t,'m_w,'s_i,'s) scheduler" and "output" :: "'s \ 't \ ('l,'t,'x,'m,'w,'o) thread_action \ 'q option" and \_invar :: "'s \ 't set \ bool" and thr_\ :: "'m_t \ ('l,'t,'x) thread_info" and thr_invar :: "'m_t \ bool" and thr_lookup :: "'t \ 'm_t \ ('x \ 'l released_locks)" and thr_update :: "'t \ 'x \ 'l released_locks \ 'm_t \ 'm_t" and thr_iterate :: "'m_t \ ('t \ ('x \ 'l released_locks), 's_t) set_iterator" and ws_\ :: "'m_w \ ('w,'t) wait_sets" and ws_invar :: "'m_w \ bool" and ws_empty :: "unit \ 'm_w" and ws_lookup :: "'t \ 'm_w \ 'w wait_set_status" and ws_update :: "'t \ 'w wait_set_status \ 'm_w \ 'm_w" and ws_delete :: "'t \ 'm_w \ 'm_w" and ws_iterate :: "'m_w \ ('t \ 'w wait_set_status, 'm_w) set_iterator" and ws_sel :: "'m_w \ ('t \ 'w wait_set_status \ bool) \ ('t \ 'w wait_set_status)" and is_\ :: "'s_i \ 't interrupts" and is_invar :: "'s_i \ bool" and is_memb :: "'t \ 's_i \ bool" and is_ins :: "'t \ 's_i \ 's_i" and is_delete :: "'t \ 's_i \ 's_i" and thr'_\ :: "'s_t \ 't set" and thr'_invar :: "'s_t \ bool" and thr'_empty :: "unit \ 's_t" and thr'_ins_dj :: "'t \ 's_t \ 's_t" and invariant :: "('l,'t,'x,'m,'w) state set" + assumes invariant: "invariant3p \.redT invariant" sublocale scheduler_ext < pick_wakeup_spec final r convert_RA pick_wakeup \_invar thr_\ thr_invar ws_\ ws_invar by(rule pick_wakeup_spec_via_sel)(unfold_locales) sublocale scheduler_ext < scheduler final r convert_RA schedule "output" "pick_wakeup" \_invar thr_\ thr_invar thr_lookup thr_update ws_\ ws_invar ws_lookup ws_update ws_delete ws_iterate is_\ is_invar is_memb is_ins is_delete invariant by(unfold_locales)(rule invariant) subsection \Schedulers for deterministic small-step semantics\ text \ The default code equations for @{term Predicate.the} impose the type class constraint \eq\ on the predicate elements. For the semantics, which contains the heap, there might be no such instance, so we use new constants for which other code equations can be used. These do not add the type class constraint, but may fail more often with non-uniqueness exception. \ definition singleton2 where [simp]: "singleton2 = Predicate.singleton" definition the_only2 where [simp]: "the_only2 = Predicate.the_only" definition the2 where [simp]: "the2 = Predicate.the" context multithreaded_base begin definition step_thread :: "(('l,'t,'x,'m,'w,'o) thread_action \ 's) \ ('l,'t,'x,'m,'w) state \ 't \ ('t \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) option \ 's) option" where "\ln. step_thread update_state s t = (case thr s t of \(x, ln)\ \ if ln = no_wait_locks then if \ta x' m'. t \ (x, shr s) -ta\ (x', m') \ actions_ok s t ta then let (ta, x', m') = THE (ta, x', m'). t \ (x, shr s) -ta\ (x', m') \ actions_ok s t ta in \(t, \(ta, x', m')\, update_state ta)\ else None else if may_acquire_all (locks s) t ln \ \ waiting (wset s t) then \(t, None, update_state (K$ [], [], [], [], [], convert_RA ln))\ else None | None \ None)" lemma step_thread_NoneD: "step_thread update_state s t = None \ t \ active_threads s" unfolding step_thread_def by(fastforce simp add: split_beta elim!: active_threads.cases split: if_split_asm) lemma inactive_step_thread_eq_NoneI: "t \ active_threads s \ step_thread update_state s t = None" unfolding step_thread_def by(fastforce simp add: split_beta split: if_split_asm intro: active_threads.intros) lemma step_thread_eq_None_conv: "step_thread update_state s t = None \ t \ active_threads s" by(blast dest: step_thread_NoneD intro: inactive_step_thread_eq_NoneI) lemma step_thread_eq_Some_activeD: "step_thread update_state s t = \(t', taxm\')\ \ t' = t \ t \ active_threads s" unfolding step_thread_def by(fastforce split: if_split_asm simp add: split_beta intro: active_threads.intros) declare actions_ok_iff [simp del] declare actions_ok.cases [rule del] lemma step_thread_Some_NoneD: "step_thread update_state s t' = \(t, None, \')\ \ \x ln n. thr s t = \(x, ln)\ \ ln $ n > 0 \ \ waiting (wset s t) \ may_acquire_all (locks s) t ln \ \' = update_state (K$ [], [], [], [], [], convert_RA ln)" unfolding step_thread_def by(auto split: if_split_asm simp add: split_beta elim!: neq_no_wait_locksE) lemma step_thread_Some_SomeD: "\ deterministic I; step_thread update_state s t' = \(t, \(ta, x', m')\, \')\; s \ I \ \ \x. thr s t = \(x, no_wait_locks)\ \ t \ \x, shr s\ -ta\ \x', m'\ \ actions_ok s t ta \ \' = update_state ta" unfolding step_thread_def by(auto simp add: split_beta deterministic_THE split: if_split_asm) end context scheduler_base_aux begin definition step_thread :: "(('l,'t,'x,'m,'w,'o) thread_action \ 's) \ ('l,'t,'m,'m_t,'m_w,'s_i) state_refine \ 't \ ('t \ (('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm) option \ 's) option" where "\ln. step_thread update_state s t = (case thr_lookup t (thr s) of \(x, ln)\ \ if ln = no_wait_locks then let reds = do { (ta, x', m') \ r t (x, shr s); if actions_ok s t ta then Predicate.single (ta, x', m') else bot } in if Predicate.holds (reds \ (\_. Predicate.single ())) then let (ta, x', m') = the2 reds in \(t, \(ta, x', m')\, update_state ta)\ else None else if may_acquire_all (locks s) t ln \ \ waiting (ws_lookup t (wset s)) then \(t, None, update_state (K$ [], [], [], [], [],convert_RA ln))\ else None | None \ None)" end context scheduler_aux begin lemma deterministic_THE2: assumes "\.deterministic I" and tst: "thr_\ (thr s) t = \(x, no_wait_locks)\" and red: "Predicate.eval (r t (x, shr s)) (ta, x', m')" and aok: "\.actions_ok (state_\ s) t ta" and I: "state_\ s \ I" shows "Predicate.the (r t (x, shr s) \ (\(ta, x', m'). if \.actions_ok (state_\ s) t ta then Predicate.single (ta, x', m') else bot)) = (ta, x', m')" proof - show ?thesis unfolding the_def apply(rule the_equality) apply(rule bindI[OF red]) apply(simp add: singleI aok) apply(erule bindE) apply(clarsimp split: if_split_asm) apply(drule (1) \.deterministicD[OF \\.deterministic I\, where s="state_\ s", simplified, OF red _ tst aok]) apply(rule I) apply simp done qed lemma step_thread_correct: assumes det: "\.deterministic I" and invar: "\_invar \ (dom (thr_\ (thr s)))" "state_invar s" "state_\ s \ I" shows "map_option (apsnd (apsnd \_\)) (step_thread update_state s t) = \.step_thread (\_\ \ update_state) (state_\ s) t" (is ?thesis1) and "(\ta. FWThread.thread_oks (thr_\ (thr s)) \ta\\<^bsub>t\<^esub> \ \_invar (update_state ta) (dom (thr_\ (thr s)) \ {t. \x m. NewThread t x m \ set \ta\\<^bsub>t\<^esub>})) \ case_option True (\(t, taxm, \). \_invar \ (case taxm of None \ dom (thr_\ (thr s)) | Some (ta, x', m') \ dom (thr_\ (thr s)) \ {t. \x m. NewThread t x m \ set \ta\\<^bsub>t\<^esub>})) (step_thread update_state s t)" (is "(\ta. ?tso ta \ ?inv ta) \ ?thesis2") proof - have "?thesis1 \ ((\ta. ?tso ta \ ?inv ta) \ ?thesis2)" proof(cases "step_thread update_state s t") case None with invar show ?thesis apply (auto simp add: thr.lookup_correct \.step_thread_def step_thread_def ws.lookup_correct split_beta holds_eq split: if_split_asm cong del: image_cong_simp) apply metis apply metis done next case (Some a) then obtain t' taxm \' where rrs: "step_thread update_state s t = \(t', taxm, \')\" by(cases a) auto show ?thesis proof(cases "taxm") case None with rrs invar have ?thesis1 by(auto simp add: thr.lookup_correct ws.lookup_correct \.step_thread_def step_thread_def split_beta split: if_split_asm) moreover { let ?ta = "(K$ [], [], [], [], [], convert_RA (snd (the (thr_lookup t (thr s)))))" assume "?tso ?ta \ ?inv ?ta" hence ?thesis2 using None rrs by(auto simp add: thr.lookup_correct ws.lookup_correct \.step_thread_def step_thread_def split_beta split: if_split_asm) } ultimately show ?thesis by blast next case (Some a) with rrs obtain ta x' m' where rrs: "step_thread update_state s t = \(t', \(ta, x', m')\, \')\" by(cases a) fastforce with invar have ?thesis1 by (auto simp add: thr.lookup_correct ws.lookup_correct \.step_thread_def step_thread_def split_beta \.deterministic_THE [OF det, where s="state_\ s", simplified] deterministic_THE2[OF det] holds_eq split: if_split_asm cong del: image_cong_simp) blast+ moreover { assume "?tso ta \ ?inv ta" hence ?thesis2 using rrs invar by(auto simp add: thr.lookup_correct ws.lookup_correct \.step_thread_def step_thread_def split_beta \.deterministic_THE[OF det, where s="state_\ s", simplified] deterministic_THE2[OF det] holds_eq split: if_split_asm)(auto simp add: \.actions_ok_iff) } ultimately show ?thesis by blast qed qed thus ?thesis1 "(\ta. ?tso ta \ ?inv ta) \ ?thesis2" by blast+ qed lemma step_thread_eq_None_conv: assumes det: "\.deterministic I" and invar: "state_invar s" "state_\ s \ I" shows "step_thread update_state s t = None \ t \ \.active_threads (state_\ s)" using assms step_thread_correct(1)[OF det _ invar(1), of "\_ _. True", of id update_state t] by(simp add: map_option.id \.step_thread_eq_None_conv) lemma step_thread_Some_NoneD: assumes det: "\.deterministic I" and step: "step_thread update_state s t' = \(t, None, \')\" and invar: "state_invar s" "state_\ s \ I" shows "\x ln n. thr_\ (thr s) t = \(x, ln)\ \ ln $ n > 0 \ \ waiting (ws_\ (wset s) t) \ may_acquire_all (locks s) t ln \ \' = update_state (K$ [], [], [], [], [], convert_RA ln)" using assms step_thread_correct(1)[OF det _ invar(1), of "\_ _. True", of id update_state t'] by(fastforce simp add: map_option.id dest: \.step_thread_Some_NoneD[OF sym]) lemma step_thread_Some_SomeD: assumes det: "\.deterministic I" and step: "step_thread update_state s t' = \(t, \(ta, x', m')\, \')\" and invar: "state_invar s" "state_\ s \ I" shows "\x. thr_\ (thr s) t = \(x, no_wait_locks)\ \ Predicate.eval (r t (x, shr s)) (ta, x', m') \ actions_ok s t ta \ \' = update_state ta" using assms step_thread_correct(1)[OF det _ invar(1), of "\_ _. True", of id update_state t'] by(auto simp add: map_option.id dest: \.step_thread_Some_SomeD[OF det sym]) end subsection \Code Generator setup\ lemmas [code] = scheduler_base_aux.free_thread_id_def scheduler_base_aux.redT_updT.simps scheduler_base_aux.redT_updTs_def scheduler_base_aux.thread_ok.simps scheduler_base_aux.thread_oks.simps scheduler_base_aux.wset_actions_ok_def scheduler_base_aux.cond_action_ok.simps scheduler_base_aux.cond_action_oks_def scheduler_base_aux.redT_updI.simps scheduler_base_aux.redT_updIs.simps scheduler_base_aux.interrupt_action_ok.simps scheduler_base_aux.interrupt_actions_ok.simps scheduler_base_aux.actions_ok_def scheduler_base_aux.step_thread_def lemmas [code] = scheduler_base.exec_updW.simps scheduler_base.exec_updWs_def scheduler_base.exec_upd_def scheduler_base.execT_def scheduler_base.exec_step.simps scheduler_base.exec_aux_def scheduler_base.exec_def lemmas [code] = scheduler_ext_base.active_threads.simps lemma singleton2_code [code]: "singleton2 dfault (Predicate.Seq f) = (case f () of Predicate.Empty \ dfault () | Predicate.Insert x P \ if Predicate.is_empty P then x else Code.abort (STR ''singleton2 not unique'') (\_. singleton2 dfault (Predicate.Seq f)) | Predicate.Join P xq \ if Predicate.is_empty P then the_only2 dfault xq else if Predicate.null xq then singleton2 dfault P else Code.abort (STR ''singleton2 not unique'') (\_. singleton2 dfault (Predicate.Seq f)))" unfolding singleton2_def the_only2_def by(auto simp only: singleton_code Code.abort_def split: seq.split if_split) lemma the_only2_code [code]: "the_only2 dfault Predicate.Empty = Code.abort (STR ''the_only2 empty'') dfault" "the_only2 dfault (Predicate.Insert x P) = (if Predicate.is_empty P then x else Code.abort (STR ''the_only2 not unique'') (\_. the_only2 dfault (Predicate.Insert x P)))" "the_only2 dfault (Predicate.Join P xq) = (if Predicate.is_empty P then the_only2 dfault xq else if Predicate.null xq then singleton2 dfault P else Code.abort (STR ''the_only2 not unique'') (\_. the_only2 dfault (Predicate.Join P xq)))" unfolding singleton2_def the_only2_def by simp_all lemma the2_eq [code]: "the2 A = singleton2 (\x. Code.abort (STR ''not_unique'') (\_. the2 A)) A" unfolding the2_def singleton2_def by(rule the_eq) end diff --git a/thys/JinjaThreads/Framework/Bisimulation.thy b/thys/JinjaThreads/Framework/Bisimulation.thy --- a/thys/JinjaThreads/Framework/Bisimulation.thy +++ b/thys/JinjaThreads/Framework/Bisimulation.thy @@ -1,1296 +1,1296 @@ (* Title: JinjaThreads/Framework/Bisimulation.thy Author: Andreas Lochbihler *) section \Various notions of bisimulation\ theory Bisimulation imports LTS begin type_synonym ('a, 'b) bisim = "'a \ 'b \ bool" subsection \Strong bisimulation\ locale bisimulation_base = r1: trsys trsys1 + r2: trsys trsys2 for trsys1 :: "('s1, 'tl1) trsys" ("_/ -1-_\/ _" [50,0,50] 60) and trsys2 :: "('s2, 'tl2) trsys" ("_/ -2-_\/ _" [50,0,50] 60) + fixes bisim :: "('s1, 's2) bisim" ("_/ \ _" [50, 50] 60) and tlsim :: "('tl1, 'tl2) bisim" ("_/ \ _" [50, 50] 60) begin notation r1.Trsys ("_/ -1-_\*/ _" [50,0,50] 60) and r2.Trsys ("_/ -2-_\*/ _" [50,0,50] 60) notation r1.inf_step ("_ -1-_\* \" [50, 0] 80) and r2.inf_step ("_ -2-_\* \" [50, 0] 80) notation r1.inf_step_table ("_ -1-_\*t \" [50, 0] 80) and r2.inf_step_table ("_ -2-_\*t \" [50, 0] 80) abbreviation Tlsim :: "('tl1 list, 'tl2 list) bisim" ("_/ [\] _" [50, 50] 60) where "Tlsim tl1 tl2 \ list_all2 tlsim tl1 tl2" abbreviation Tlsiml :: "('tl1 llist, 'tl2 llist) bisim" ("_/ [[\]] _" [50, 50] 60) where "Tlsiml tl1 tl2 \ llist_all2 tlsim tl1 tl2" end locale bisimulation = bisimulation_base + constrains trsys1 :: "('s1, 'tl1) trsys" and trsys2 :: "('s2, 'tl2) trsys" and bisim :: "('s1, 's2) bisim" and tlsim :: "('tl1, 'tl2) bisim" assumes simulation1: "\ s1 \ s2; s1 -1-tl1\ s1' \ \ \s2' tl2. s2 -2-tl2\ s2' \ s1' \ s2' \ tl1 \ tl2" and simulation2: "\ s1 \ s2; s2 -2-tl2\ s2' \ \ \s1' tl1. s1 -1-tl1\ s1' \ s1' \ s2' \ tl1 \ tl2" begin lemma bisimulation_flip: "bisimulation trsys2 trsys1 (flip bisim) (flip tlsim)" by(unfold_locales)(unfold flip_simps,(blast intro: simulation1 simulation2)+) end lemma bisimulation_flip_simps [flip_simps]: "bisimulation trsys2 trsys1 (flip bisim) (flip tlsim) = bisimulation trsys1 trsys2 bisim tlsim" by(auto dest: bisimulation.bisimulation_flip simp only: flip_flip) context bisimulation begin lemma simulation1_rtrancl: "\s1 -1-tls1\* s1'; s1 \ s2\ \ \s2' tls2. s2 -2-tls2\* s2' \ s1' \ s2' \ tls1 [\] tls2" proof(induct rule: rtrancl3p.induct) case rtrancl3p_refl thus ?case by(auto intro: rtrancl3p.rtrancl3p_refl) next case (rtrancl3p_step s1 tls1 s1' tl1 s1'') from \s1 \ s2 \ \s2' tls2. s2 -2-tls2\* s2' \ s1' \ s2' \ tls1 [\] tls2\ \s1 \ s2\ obtain s2' tls2 where "s2 -2-tls2\* s2'" "s1' \ s2'" "tls1 [\] tls2" by blast moreover from \s1' -1-tl1\ s1''\ \s1' \ s2'\ obtain s2'' tl2 where "s2' -2-tl2\ s2''" "s1'' \ s2''" "tl1 \ tl2" by(auto dest: simulation1) ultimately have "s2 -2-tls2 @ [tl2]\* s2''" "tls1 @ [tl1] [\] tls2 @ [tl2]" by(auto intro: rtrancl3p.rtrancl3p_step list_all2_appendI) with \s1'' \ s2''\ show ?case by(blast) qed lemma simulation2_rtrancl: "\s2 -2-tls2\* s2'; s1 \ s2\ \ \s1' tls1. s1 -1-tls1\* s1' \ s1' \ s2' \ tls1 [\] tls2" using bisimulation.simulation1_rtrancl[OF bisimulation_flip] unfolding flip_simps . lemma simulation1_inf_step: assumes red1: "s1 -1-tls1\* \" and bisim: "s1 \ s2" shows "\tls2. s2 -2-tls2\* \ \ tls1 [[\]] tls2" proof - from r1.inf_step_imp_inf_step_table[OF red1] obtain stls1 where red1': "s1 -1-stls1\*t \" and tls1: "tls1 = lmap (fst \ snd) stls1" by blast define tl1_to_tl2 where "tl1_to_tl2 s2 stls1 = unfold_llist (\(s2, stls1). lnull stls1) (\(s2, stls1). let (s1, tl1, s1') = lhd stls1; (tl2, s2') = SOME (tl2, s2'). trsys2 s2 tl2 s2' \ s1' \ s2' \ tl1 \ tl2 in (s2, tl2, s2')) (\(s2, stls1). let (s1, tl1, s1') = lhd stls1; (tl2, s2') = SOME (tl2, s2'). trsys2 s2 tl2 s2' \ s1' \ s2' \ tl1 \ tl2 in (s2', ltl stls1)) (s2, stls1)" for s2 :: 's2 and stls1 :: "('s1 \ 'tl1 \ 's1) llist" have tl1_to_tl2_simps [simp]: "\s2 stls1. lnull (tl1_to_tl2 s2 stls1) \ lnull stls1" "\s2 stls1. \ lnull stls1 \ lhd (tl1_to_tl2 s2 stls1) = (let (s1, tl1, s1') = lhd stls1; (tl2, s2') = SOME (tl2, s2'). trsys2 s2 tl2 s2' \ s1' \ s2' \ tl1 \ tl2 in (s2, tl2, s2'))" "\s2 stls1. \ lnull stls1 \ ltl (tl1_to_tl2 s2 stls1) = (let (s1, tl1, s1') = lhd stls1; (tl2, s2') = SOME (tl2, s2'). trsys2 s2 tl2 s2' \ s1' \ s2' \ tl1 \ tl2 in tl1_to_tl2 s2' (ltl stls1))" "\s2. tl1_to_tl2 s2 LNil = LNil" "\s2 s1 tl1 s1' stls1'. tl1_to_tl2 s2 (LCons (s1, tl1, s1') stls1') = LCons (s2, SOME (tl2, s2'). trsys2 s2 tl2 s2' \ s1' \ s2' \ tl1 \ tl2) (tl1_to_tl2 (snd (SOME (tl2, s2'). trsys2 s2 tl2 s2' \ s1' \ s2' \ tl1 \ tl2)) stls1')" by(simp_all add: tl1_to_tl2_def split_beta) have [simp]: "llength (tl1_to_tl2 s2 stls1) = llength stls1" by(coinduction arbitrary: s2 stls1 rule: enat_coinduct)(auto simp add: epred_llength split_beta) from red1' bisim have "s2 -2-tl1_to_tl2 s2 stls1\*t \" proof(coinduction arbitrary: s2 s1 stls1) case (inf_step_table s2 s1 stls1) note red1' = \s1 -1-stls1\*t \\ and bisim = \s1 \ s2\ from red1' show ?case proof(cases) case (inf_step_tableI s1' stls1' tl1) hence stls1: "stls1 = LCons (s1, tl1, s1') stls1'" and r: "s1 -1-tl1\ s1'" and reds1: "s1' -1-stls1'\*t \" by simp_all let ?tl2s2' = "SOME (tl2, s2'). s2 -2-tl2\ s2' \ s1' \ s2' \ tl1 \ tl2" let ?tl2 = "fst ?tl2s2'" let ?s2' = "snd ?tl2s2'" from simulation1[OF bisim r] obtain s2' tl2 where "s2 -2-tl2\ s2'" "s1' \ s2'" "tl1 \ tl2" by blast hence "(\(tl2, s2'). s2 -2-tl2\ s2' \ s1' \ s2' \ tl1 \ tl2) (tl2, s2')" by simp hence "(\(tl2, s2'). s2 -2-tl2\ s2' \ s1' \ s2' \ tl1 \ tl2) ?tl2s2'" by(rule someI) hence "s2 -2-?tl2\ ?s2'" "s1' \ ?s2'" "tl1 \ ?tl2" by(simp_all add: split_beta) then show ?thesis using reds1 stls1 by(fastforce intro: prod_eqI) qed qed hence "s2 -2-lmap (fst \ snd) (tl1_to_tl2 s2 stls1)\* \" by(rule r2.inf_step_table_imp_inf_step) moreover have "tls1 [[\]] lmap (fst \ snd) (tl1_to_tl2 s2 stls1)" proof(rule llist_all2_all_lnthI) show "llength tls1 = llength (lmap (fst \ snd) (tl1_to_tl2 s2 stls1))" using tls1 by simp next fix n assume "enat n < llength tls1" thus "lnth tls1 n \ lnth (lmap (fst \ snd) (tl1_to_tl2 s2 stls1)) n" using red1' bisim unfolding tls1 proof(induct n arbitrary: s1 s2 stls1 rule: nat_less_induct) case (1 n) hence IH: "\m s1 s2 stls1. \ m < n; enat m < llength (lmap (fst \ snd) stls1); s1 -1-stls1\*t \; s1 \ s2 \ \ lnth (lmap (fst \ snd) stls1) m \ lnth (lmap (fst \ snd) (tl1_to_tl2 s2 stls1)) m" by blast from \s1 -1-stls1\*t \\ show ?case proof cases case (inf_step_tableI s1' stls1' tl1) hence stls1: "stls1 = LCons (s1, tl1, s1') stls1'" and r: "s1 -1-tl1\ s1'" and reds: "s1' -1-stls1'\*t \" by simp_all let ?tl2s2' = "SOME (tl2, s2'). s2 -2-tl2\ s2' \ s1' \ s2' \ tl1 \ tl2" let ?tl2 = "fst ?tl2s2'" let ?s2' = "snd ?tl2s2'" from simulation1[OF \s1 \ s2\ r] obtain s2' tl2 where "s2 -2-tl2\ s2' \ s1' \ s2' \ tl1 \ tl2" by blast hence "(\(tl2, s2'). s2 -2-tl2\ s2' \ s1' \ s2' \ tl1 \ tl2) (tl2, s2')" by simp hence "(\(tl2, s2'). s2 -2-tl2\ s2' \ s1' \ s2' \ tl1 \ tl2) ?tl2s2'" by(rule someI) hence bisim': "s1' \ ?s2'" and tlsim: "tl1 \ ?tl2" by(simp_all add: split_beta) show ?thesis proof(cases n) case 0 with stls1 tlsim show ?thesis by simp next case (Suc m) hence "m < n" by simp moreover have "enat m < llength (lmap (fst \ snd) stls1')" using stls1 \enat n < llength (lmap (fst \ snd) stls1)\ Suc by(simp add: Suc_ile_eq) ultimately have "lnth (lmap (fst \ snd) stls1') m \ lnth (lmap (fst \ snd) (tl1_to_tl2 ?s2' stls1')) m" using reds bisim' by(rule IH) with Suc stls1 show ?thesis by(simp del: o_apply) qed qed qed qed ultimately show ?thesis by blast qed lemma simulation2_inf_step: "\ s2 -2-tls2\* \; s1 \ s2 \ \ \tls1. s1 -1-tls1\* \ \ tls1 [[\]] tls2" using bisimulation.simulation1_inf_step[OF bisimulation_flip] unfolding flip_simps . end locale bisimulation_final_base = bisimulation_base + constrains trsys1 :: "('s1, 'tl1) trsys" and trsys2 :: "('s2, 'tl2) trsys" and bisim :: "('s1, 's2) bisim" and tlsim :: "('tl1, 'tl2) bisim" fixes final1 :: "'s1 \ bool" and final2 :: "'s2 \ bool" locale bisimulation_final = bisimulation_final_base + bisimulation + constrains trsys1 :: "('s1, 'tl1) trsys" and trsys2 :: "('s2, 'tl2) trsys" and bisim :: "('s1, 's2) bisim" and tlsim :: "('tl1, 'tl2) bisim" and final1 :: "'s1 \ bool" and final2 :: "'s2 \ bool" assumes bisim_final: "s1 \ s2 \ final1 s1 \ final2 s2" begin lemma bisimulation_final_flip: "bisimulation_final trsys2 trsys1 (flip bisim) (flip tlsim) final2 final1" apply(intro_locales) apply(rule bisimulation_flip) apply(unfold_locales) by(unfold flip_simps, rule bisim_final[symmetric]) end lemma bisimulation_final_flip_simps [flip_simps]: "bisimulation_final trsys2 trsys1 (flip bisim) (flip tlsim) final2 final1 = bisimulation_final trsys1 trsys2 bisim tlsim final1 final2" by(auto dest: bisimulation_final.bisimulation_final_flip simp only: flip_flip) context bisimulation_final begin lemma final_simulation1: "\ s1 \ s2; s1 -1-tls1\* s1'; final1 s1' \ \ \s2' tls2. s2 -2-tls2\* s2' \ s1' \ s2' \ final2 s2' \ tls1 [\] tls2" by(auto dest: bisim_final dest!: simulation1_rtrancl) lemma final_simulation2: "\ s1 \ s2; s2 -2-tls2\* s2'; final2 s2' \ \ \s1' tls1. s1 -1-tls1\* s1' \ s1' \ s2' \ final1 s1' \ tls1 [\] tls2" by(auto dest: bisim_final dest!: simulation2_rtrancl) end subsection \Delay bisimulation\ locale delay_bisimulation_base = bisimulation_base + trsys1?: \trsys trsys1 \move1 + trsys2?: \trsys trsys2 \move2 for \move1 \move2 + constrains trsys1 :: "('s1, 'tl1) trsys" and trsys2 :: "('s2, 'tl2) trsys" and bisim :: "('s1, 's2) bisim" and tlsim :: "('tl1, 'tl2) bisim" and \move1 :: "('s1, 'tl1) trsys" and \move2 :: "('s2, 'tl2) trsys" begin notation trsys1.silent_move ("_/ -\1\ _" [50, 50] 60) and trsys2.silent_move ("_/ -\2\ _" [50, 50] 60) notation trsys1.silent_moves ("_/ -\1\* _" [50, 50] 60) and trsys2.silent_moves ("_/ -\2\* _" [50, 50] 60) notation trsys1.silent_movet ("_/ -\1\+ _" [50, 50] 60) and trsys2.silent_movet ("_/ -\2\+ _" [50, 50] 60) notation trsys1.\rtrancl3p ("_ -\1-_\* _" [50, 0, 50] 60) and trsys2.\rtrancl3p ("_ -\2-_\* _" [50, 0, 50] 60) notation trsys1.\inf_step ("_ -\1-_\* \" [50, 0] 80) and trsys2.\inf_step ("_ -\2-_\* \" [50, 0] 80) notation trsys1.\diverge ("_ -\1\ \" [50] 80) and trsys2.\diverge ("_ -\2\ \" [50] 80) notation trsys1.\inf_step_table ("_ -\1-_\*t \" [50, 0] 80) and trsys2.\inf_step_table ("_ -\2-_\*t \" [50, 0] 80) notation trsys1.\Runs ("_ \1 _" [50, 50] 51) and trsys2.\Runs ("_ \2 _" [50, 50] 51) lemma simulation_silent1I': assumes "\s2'. (if \1 s1' s1 then trsys2.silent_moves else trsys2.silent_movet) s2 s2' \ s1' \ s2'" shows "s1' \ s2 \ \1^++ s1' s1 \ (\s2'. s2 -\2\+ s2' \ s1' \ s2')" proof - from assms obtain s2' where red: "(if \1 s1' s1 then trsys2.silent_moves else trsys2.silent_movet) s2 s2'" and bisim: "s1' \ s2'" by blast show ?thesis proof(cases "\1 s1' s1") case True with red have "s2 -\2\* s2'" by simp thus ?thesis using bisim True by cases(blast intro: rtranclp_into_tranclp1)+ next case False with red bisim show ?thesis by auto qed qed lemma simulation_silent2I': assumes "\s1'. (if \2 s2' s2 then trsys1.silent_moves else trsys1.silent_movet) s1 s1' \ s1' \ s2'" shows "s1 \ s2' \ \2^++ s2' s2 \ (\s1'. s1 -\1\+ s1' \ s1' \ s2')" using assms by(rule delay_bisimulation_base.simulation_silent1I') end locale delay_bisimulation_obs = delay_bisimulation_base _ _ _ _ \move1 \move2 for \move1 :: "'s1 \ 'tl1 \ 's1 \ bool" and \move2 :: "'s2 \ 'tl2 \ 's2 \ bool" + assumes simulation1: "\ s1 \ s2; s1 -1-tl1\ s1'; \ \move1 s1 tl1 s1' \ \ \s2' s2'' tl2. s2 -\2\* s2' \ s2' -2-tl2\ s2'' \ \ \move2 s2' tl2 s2'' \ s1' \ s2'' \ tl1 \ tl2" and simulation2: "\ s1 \ s2; s2 -2-tl2\ s2'; \ \move2 s2 tl2 s2' \ \ \s1' s1'' tl1. s1 -\1\* s1' \ s1' -1-tl1\ s1'' \ \ \move1 s1' tl1 s1'' \ s1'' \ s2' \ tl1 \ tl2" begin lemma delay_bisimulation_obs_flip: "delay_bisimulation_obs trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1" apply(unfold_locales) apply(unfold flip_simps) by(blast intro: simulation1 simulation2)+ end lemma delay_bisimulation_obs_flip_simps [flip_simps]: "delay_bisimulation_obs trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1 = delay_bisimulation_obs trsys1 trsys2 bisim tlsim \move1 \move2" by(auto dest: delay_bisimulation_obs.delay_bisimulation_obs_flip simp only: flip_flip) locale delay_bisimulation_diverge = delay_bisimulation_obs _ _ _ _ \move1 \move2 for \move1 :: "'s1 \ 'tl1 \ 's1 \ bool" and \move2 :: "'s2 \ 'tl2 \ 's2 \ bool" + assumes simulation_silent1: "\ s1 \ s2; s1 -\1\ s1' \ \ \s2'. s2 -\2\* s2' \ s1' \ s2'" and simulation_silent2: "\ s1 \ s2; s2 -\2\ s2' \ \ \s1'. s1 -\1\* s1' \ s1' \ s2'" and \diverge_bisim_inv: "s1 \ s2 \ s1 -\1\ \ \ s2 -\2\ \" begin lemma delay_bisimulation_diverge_flip: "delay_bisimulation_diverge trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1" apply(rule delay_bisimulation_diverge.intro) apply(rule delay_bisimulation_obs_flip) apply(unfold_locales) apply(unfold flip_simps) by(blast intro: simulation_silent1 simulation_silent2 \diverge_bisim_inv[symmetric] del: iffI)+ end lemma delay_bisimulation_diverge_flip_simps [flip_simps]: "delay_bisimulation_diverge trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1 = delay_bisimulation_diverge trsys1 trsys2 bisim tlsim \move1 \move2" by(auto dest: delay_bisimulation_diverge.delay_bisimulation_diverge_flip simp only: flip_flip) context delay_bisimulation_diverge begin lemma simulation_silents1: assumes bisim: "s1 \ s2" and moves: "s1 -\1\* s1'" shows "\s2'. s2 -\2\* s2' \ s1' \ s2'" using moves bisim proof induct case base thus ?case by(blast) next case (step s1' s1'') from \s1 \ s2 \ \s2'. s2 -\2\* s2' \ s1' \ s2'\ \s1 \ s2\ obtain s2' where "s2 -\2\* s2'" "s1' \ s2'" by blast from simulation_silent1[OF \s1' \ s2'\ \s1' -\1\ s1''\] obtain s2'' where "s2' -\2\* s2''" "s1'' \ s2''" by blast from \s2 -\2\* s2'\ \s2' -\2\* s2''\ have "s2 -\2\* s2''" by(rule rtranclp_trans) with \s1'' \ s2''\ show ?case by blast qed lemma simulation_silents2: "\ s1 \ s2; s2 -\2\* s2' \ \ \s1'. s1 -\1\* s1' \ s1' \ s2'" using delay_bisimulation_diverge.simulation_silents1[OF delay_bisimulation_diverge_flip] unfolding flip_simps . lemma simulation1_\rtrancl3p: "\ s1 -\1-tls1\* s1'; s1 \ s2 \ \ \tls2 s2'. s2 -\2-tls2\* s2' \ s1' \ s2' \ tls1 [\] tls2" proof(induct arbitrary: s2 rule: trsys1.\rtrancl3p.induct) case (\rtrancl3p_refl s) thus ?case by(auto intro: \trsys.\rtrancl3p.intros) next case (\rtrancl3p_step s1 s1' tls1 s1'' tl1) from simulation1[OF \s1 \ s2\ \s1 -1-tl1\ s1'\ \\ \move1 s1 tl1 s1'\] obtain s2' s2'' tl2 where \red: "s2 -\2\* s2'" and red: "s2' -2-tl2\ s2''" and n\: "\ \move2 s2' tl2 s2''" and bisim': "s1' \ s2''" and tlsim: "tl1 \ tl2" by blast from bisim' \s1' \ s2'' \ \tls2 s2'. s2'' -\2-tls2\* s2' \ s1'' \ s2' \ tls1 [\] tls2\ obtain tls2 s2''' where IH: "s2'' -\2-tls2\* s2'''" "s1'' \ s2'''" "tls1 [\] tls2" by blast from \red have "s2 -\2-[]\* s2'" by(rule trsys2.silent_moves_into_\rtrancl3p) also from red n\ IH(1) have "s2' -\2-tl2 # tls2\* s2'''" by(rule \rtrancl3p.\rtrancl3p_step) finally show ?case using IH tlsim by fastforce next case (\rtrancl3p_\step s1 s1' tls1 s1'' tl1) from \s1 -1-tl1\ s1'\ \\move1 s1 tl1 s1'\ have "s1 -\1\ s1'" .. from simulation_silent1[OF \s1 \ s2\ this] obtain s2' where \red: "s2 -\2\* s2'" and bisim': "s1' \ s2'" by blast from \red have "s2 -\2-[]\* s2'" by(rule trsys2.silent_moves_into_\rtrancl3p) also from bisim' \s1' \ s2' \ \tls2 s2''. s2' -\2-tls2\* s2'' \ s1'' \ s2'' \ tls1 [\] tls2\ obtain tls2 s2'' where IH: "s2' -\2-tls2\* s2''" "s1'' \ s2''" "tls1 [\] tls2" by blast note \s2' -\2-tls2\* s2''\ finally show ?case using IH by auto qed lemma simulation2_\rtrancl3p: "\ s2 -\2-tls2\* s2'; s1 \ s2 \ \ \tls1 s1'. s1 -\1-tls1\* s1' \ s1' \ s2' \ tls1 [\] tls2" using delay_bisimulation_diverge.simulation1_\rtrancl3p[OF delay_bisimulation_diverge_flip] unfolding flip_simps . lemma simulation1_\inf_step: assumes \inf1: "s1 -\1-tls1\* \" and bisim: "s1 \ s2" shows "\tls2. s2 -\2-tls2\* \ \ tls1 [[\]] tls2" proof - from trsys1.\inf_step_imp_\inf_step_table[OF \inf1] obtain sstls1 where \inf1': "s1 -\1-sstls1\*t \" and tls1: "tls1 = lmap (fst \ snd \ snd) sstls1" by blast define tl1_to_tl2 where "tl1_to_tl2 s2 sstls1 = unfold_llist (\(s2, sstls1). lnull sstls1) (\(s2, sstls1). let (s1, s1', tl1, s1'') = lhd sstls1; (s2', tl2, s2'') = SOME (s2', tl2, s2''). s2 -\2\* s2' \ trsys2 s2' tl2 s2'' \ \ \move2 s2' tl2 s2'' \ s1'' \ s2'' \ tl1 \ tl2 in (s2, s2', tl2, s2'')) (\(s2, sstls1). let (s1, s1', tl1, s1'') = lhd sstls1; (s2', tl2, s2'') = SOME (s2', tl2, s2''). s2 -\2\* s2' \ trsys2 s2' tl2 s2'' \ \ \move2 s2' tl2 s2'' \ s1'' \ s2'' \ tl1 \ tl2 in (s2'', ltl sstls1)) (s2, sstls1)" for s2 :: 's2 and sstls1 :: "('s1 \ 's1 \ 'tl1 \ 's1) llist" have [simp]: "\s2 sstls1. lnull (tl1_to_tl2 s2 sstls1) \ lnull sstls1" "\s2 sstls1. \ lnull sstls1 \ lhd (tl1_to_tl2 s2 sstls1) = (let (s1, s1', tl1, s1'') = lhd sstls1; (s2', tl2, s2'') = SOME (s2', tl2, s2''). s2 -\2\* s2' \ trsys2 s2' tl2 s2'' \ \ \move2 s2' tl2 s2'' \ s1'' \ s2'' \ tl1 \ tl2 in (s2, s2', tl2, s2''))" "\s2 sstls1. \ lnull sstls1 \ ltl (tl1_to_tl2 s2 sstls1) = (let (s1, s1', tl1, s1'') = lhd sstls1; (s2', tl2, s2'') = SOME (s2', tl2, s2''). s2 -\2\* s2' \ trsys2 s2' tl2 s2'' \ \ \move2 s2' tl2 s2'' \ s1'' \ s2'' \ tl1 \ tl2 in tl1_to_tl2 s2'' (ltl sstls1))" "\s2. tl1_to_tl2 s2 LNil = LNil" "\s2 s1 s1' tl1 s1'' stls1'. tl1_to_tl2 s2 (LCons (s1, s1', tl1, s1'') stls1') = LCons (s2, SOME (s2', tl2, s2''). s2 -\2\* s2' \ trsys2 s2' tl2 s2'' \ \ \move2 s2' tl2 s2'' \ s1'' \ s2'' \ tl1 \ tl2) (tl1_to_tl2 (snd (snd (SOME (s2', tl2, s2''). s2 -\2\* s2' \ trsys2 s2' tl2 s2'' \ \ \move2 s2' tl2 s2'' \ s1'' \ s2'' \ tl1 \ tl2))) stls1')" by(simp_all add: tl1_to_tl2_def split_beta) have [simp]: "llength (tl1_to_tl2 s2 sstls1) = llength sstls1" by(coinduction arbitrary: s2 sstls1 rule: enat_coinduct)(auto simp add: epred_llength split_beta) define sstls2 where "sstls2 = tl1_to_tl2 s2 sstls1" with \inf1' bisim have "\s1 sstls1. s1 -\1-sstls1\*t \ \ sstls2 = tl1_to_tl2 s2 sstls1 \ s1 \ s2" by blast from \inf1' bisim have "s2 -\2-tl1_to_tl2 s2 sstls1\*t \" proof(coinduction arbitrary: s2 s1 sstls1) case (\inf_step_table s2 s1 sstls1) note \inf' = \s1 -\1-sstls1\*t \\ and bisim = \s1 \ s2\ from \inf' show ?case proof(cases) case (\inf_step_table_Cons s1' s1'' sstls1' tl1) hence sstls1: "sstls1 = LCons (s1, s1', tl1, s1'') sstls1'" and \s: "s1 -\1\* s1'" and r: "s1' -1-tl1\ s1''" and n\: "\ \move1 s1' tl1 s1''" and reds1: "s1'' -\1-sstls1'\*t \" by simp_all let ?P = "\(s2', tl2, s2''). s2 -\2\* s2' \ trsys2 s2' tl2 s2'' \ \ \move2 s2' tl2 s2'' \ s1'' \ s2'' \ tl1 \ tl2" let ?s2tl2s2' = "Eps ?P" let ?s2'' = "snd (snd ?s2tl2s2')" from simulation_silents1[OF \s1 \ s2\ \s] obtain s2' where "s2 -\2\* s2'" "s1' \ s2'" by blast from simulation1[OF \s1' \ s2'\ r n\] obtain s2'' s2''' tl2 where "s2' -\2\* s2''" and rest: "s2'' -2-tl2\ s2'''" "\ \move2 s2'' tl2 s2'''" "s1'' \ s2'''" "tl1 \ tl2" by blast from \s2 -\2\* s2'\ \s2' -\2\* s2''\ have "s2 -\2\* s2''" by(rule rtranclp_trans) with rest have "?P (s2'', tl2, s2''')" by simp hence "?P ?s2tl2s2'" by(rule someI) then show ?thesis using reds1 sstls1 by fastforce next case \inf_step_table_Nil hence [simp]: "sstls1 = LNil" and "s1 -\1\ \" by simp_all from \s1 -\1\ \\ \s1 \ s2\ have "s2 -\2\ \" by(simp add: \diverge_bisim_inv) thus ?thesis using sstls2_def by simp qed qed hence "s2 -\2-lmap (fst \ snd \ snd) (tl1_to_tl2 s2 sstls1)\* \" by(rule trsys2.\inf_step_table_into_\inf_step) moreover have "tls1 [[\]] lmap (fst \ snd \ snd) (tl1_to_tl2 s2 sstls1)" proof(rule llist_all2_all_lnthI) show "llength tls1 = llength (lmap (fst \ snd \ snd) (tl1_to_tl2 s2 sstls1))" using tls1 by simp next fix n assume "enat n < llength tls1" thus "lnth tls1 n \ lnth (lmap (fst \ snd \ snd) (tl1_to_tl2 s2 sstls1)) n" using \inf1' bisim unfolding tls1 proof(induct n arbitrary: s1 s2 sstls1 rule: less_induct) case (less n) note IH = \\m s1 s2 sstls1. \ m < n; enat m < llength (lmap (fst \ snd \ snd) sstls1); s1 -\1-sstls1\*t \; s1 \ s2 \ \ lnth (lmap (fst \ snd \ snd) sstls1) m \ lnth (lmap (fst \ snd \ snd) (tl1_to_tl2 s2 sstls1)) m\ from \s1 -\1-sstls1\*t \\ show ?case proof cases case (\inf_step_table_Cons s1' s1'' sstls1' tl1) hence sstls1: "sstls1 = LCons (s1, s1', tl1, s1'') sstls1'" and \s: "s1 -\1\* s1'" and r: "s1' -1-tl1\ s1''" and n\: "\ \move1 s1' tl1 s1''" and reds: "s1'' -\1-sstls1'\*t \" by simp_all let ?P = "\(s2', tl2, s2''). s2 -\2\* s2' \ trsys2 s2' tl2 s2'' \ \ \move2 s2' tl2 s2'' \ s1'' \ s2'' \ tl1 \ tl2" let ?s2tl2s2' = "Eps ?P" let ?tl2 = "fst (snd ?s2tl2s2')" let ?s2'' = "snd (snd ?s2tl2s2')" from simulation_silents1[OF \s1 \ s2\ \s] obtain s2' where "s2 -\2\* s2'" "s1' \ s2'" by blast from simulation1[OF \s1' \ s2'\ r n\] obtain s2'' s2''' tl2 where "s2' -\2\* s2''" and rest: "s2'' -2-tl2\ s2'''" "\ \move2 s2'' tl2 s2'''" "s1'' \ s2'''" "tl1 \ tl2" by blast from \s2 -\2\* s2'\ \s2' -\2\* s2''\ have "s2 -\2\* s2''" by(rule rtranclp_trans) with rest have "?P (s2'', tl2, s2''')" by auto hence "?P ?s2tl2s2'" by(rule someI) hence "s1'' \ ?s2''" "tl1 \ ?tl2" by(simp_all add: split_beta) show ?thesis proof(cases n) case 0 with sstls1 \tl1 \ ?tl2\ show ?thesis by simp next case (Suc m) hence "m < n" by simp moreover have "enat m < llength (lmap (fst \ snd \ snd) sstls1')" using sstls1 \enat n < llength (lmap (fst \ snd \ snd) sstls1)\ Suc by(simp add: Suc_ile_eq) ultimately have "lnth (lmap (fst \ snd \ snd) sstls1') m \ lnth (lmap (fst \ snd \ snd) (tl1_to_tl2 ?s2'' sstls1')) m" using reds \s1'' \ ?s2''\ by(rule IH) with Suc sstls1 show ?thesis by(simp del: o_apply) qed next case \inf_step_table_Nil with \enat n < llength (lmap (fst \ snd \ snd) sstls1)\ have False by simp thus ?thesis .. qed qed qed ultimately show ?thesis by blast qed lemma simulation2_\inf_step: "\ s2 -\2-tls2\* \; s1 \ s2 \ \ \tls1. s1 -\1-tls1\* \ \ tls1 [[\]] tls2" using delay_bisimulation_diverge.simulation1_\inf_step[OF delay_bisimulation_diverge_flip] unfolding flip_simps . lemma no_\move1_\s_to_no_\move2: assumes "s1 \ s2" and no_\moves1: "\s1'. \ s1 -\1\ s1'" shows "\s2'. s2 -\2\* s2' \ (\s2''. \ s2' -\2\ s2'') \ s1 \ s2'" proof - have "\ s1 -\1\ \" proof assume "s1 -\1\ \" then obtain s1' where "s1 -\1\ s1'" by cases with no_\moves1[of s1'] show False by contradiction qed with \s1 \ s2\ have "\ s2 -\2\ \" by(simp add: \diverge_bisim_inv) from trsys2.not_\diverge_to_no_\move[OF this] obtain s2' where "s2 -\2\* s2'" and "\s2''. \ s2' -\2\ s2''" by blast moreover from simulation_silents2[OF \s1 \ s2\ \s2 -\2\* s2'\] obtain s1' where "s1 -\1\* s1'" and "s1' \ s2'" by blast from \s1 -\1\* s1'\ no_\moves1 have "s1' = s1" by(auto elim: converse_rtranclpE) ultimately show ?thesis using \s1' \ s2'\ by blast qed lemma no_\move2_\s_to_no_\move1: "\ s1 \ s2; \s2'. \ s2 -\2\ s2' \ \ \s1'. s1 -\1\* s1' \ (\s1''. \ s1' -\1\ s1'') \ s1' \ s2" using delay_bisimulation_diverge.no_\move1_\s_to_no_\move2[OF delay_bisimulation_diverge_flip] unfolding flip_simps . lemma no_move1_to_no_move2: assumes "s1 \ s2" and no_moves1: "\tl1 s1'. \ s1 -1-tl1\ s1'" shows "\s2'. s2 -\2\* s2' \ (\tl2 s2''. \ s2' -2-tl2\ s2'') \ s1 \ s2'" proof - from no_moves1 have "\s1'. \ s1 -\1\ s1'" by(auto) from no_\move1_\s_to_no_\move2[OF \s1 \ s2\ this] obtain s2' where "s2 -\2\* s2'" and "s1 \ s2'" and no_\moves2: "\s2''. \ s2' -\2\ s2''" by blast moreover have "\tl2 s2''. \ s2' -2-tl2\ s2''" proof fix tl2 s2'' assume "s2' -2-tl2\ s2''" with no_\moves2[of s2''] have "\ \move2 s2' tl2 s2''" by(auto) from simulation2[OF \s1 \ s2'\ \s2' -2-tl2\ s2''\ this] obtain s1' s1'' tl1 where "s1 -\1\* s1'" and "s1' -1-tl1\ s1''" by blast with no_moves1 show False by(fastforce elim: converse_rtranclpE) qed ultimately show ?thesis by blast qed lemma no_move2_to_no_move1: "\ s1 \ s2; \tl2 s2'. \ s2 -2-tl2\ s2' \ \ \s1'. s1 -\1\* s1' \ (\tl1 s1''. \ s1' -1-tl1\ s1'') \ s1' \ s2" using delay_bisimulation_diverge.no_move1_to_no_move2[OF delay_bisimulation_diverge_flip] unfolding flip_simps . lemma simulation_\Runs_table1: assumes bisim: "s1 \ s2" and run1: "trsys1.\Runs_table s1 stlsss1" shows "\stlsss2. trsys2.\Runs_table s2 stlsss2 \ tllist_all2 (\(tl1, s1'') (tl2, s2''). tl1 \ tl2 \ s1'' \ s2'') (rel_option bisim) stlsss1 stlsss2" proof(intro exI conjI) let ?P = "\(s2 :: 's2) (stlsss1 :: ('tl1 \ 's1, 's1 option) tllist) (tl2, s2''). \s2'. s2 -\2\* s2' \ s2' -2-tl2\ s2'' \ \ \move2 s2' tl2 s2'' \ snd (thd stlsss1) \ s2'' \ fst (thd stlsss1) \ tl2" define tls1_to_tls2 where "tls1_to_tls2 s2 stlsss1 = unfold_tllist (\(s2, stlsss1). is_TNil stlsss1) (\(s2, stlsss1). map_option (\s1'. SOME s2'. s2 -\2\* s2' \ (\tl s2''. \ s2' -2-tl\ s2'') \ s1' \ s2') (terminal stlsss1)) (\(s2, stlsss1). let (tl2, s2'') = Eps (?P s2 stlsss1) in (tl2, s2'')) (\(s2, stlsss1). let (tl2, s2'') = Eps (?P s2 stlsss1) in (s2'', ttl stlsss1)) (s2, stlsss1)" for s2 stlsss1 have [simp]: "\s2 stlsss1. is_TNil (tls1_to_tls2 s2 stlsss1) \ is_TNil stlsss1" "\s2 stlsss1. is_TNil stlsss1 \ terminal (tls1_to_tls2 s2 stlsss1) = map_option (\s1'. SOME s2'. s2 -\2\* s2' \ (\tl s2''. \ s2' -2-tl\ s2'') \ s1' \ s2') (terminal stlsss1)" "\s2 stlsss1. \ is_TNil stlsss1 \ thd (tls1_to_tls2 s2 stlsss1) = (let (tl2, s2'') = Eps (?P s2 stlsss1) in (tl2, s2''))" "\s2 stlsss1. \ is_TNil stlsss1 \ ttl (tls1_to_tls2 s2 stlsss1) = (let (tl2, s2'') = Eps (?P s2 stlsss1) in tls1_to_tls2 s2'' (ttl stlsss1))" "\s2 os1. tls1_to_tls2 s2 (TNil os1) = TNil (map_option (\s1'. SOME s2'. s2 -\2\* s2' \ (\tl s2''. \ s2' -2-tl\ s2'') \ s1' \ s2') os1)" by(simp_all add: tls1_to_tls2_def split_beta) have [simp]: "\s2 s1 s1' tl1 s1'' stlsss1. tls1_to_tls2 s2 (TCons (tl1, s1'') stlsss1) = (let (tl2, s2'') = SOME (tl2, s2''). \s2'. s2 -\2\* s2' \ s2' -2-tl2\ s2'' \ \ \move2 s2' tl2 s2'' \ s1'' \ s2'' \ tl1 \ tl2 in TCons (tl2, s2'') (tls1_to_tls2 s2'' stlsss1))" by(rule tllist.expand)(simp_all add: split_beta) from bisim run1 show "trsys2.\Runs_table s2 (tls1_to_tls2 s2 stlsss1)" proof(coinduction arbitrary: s2 s1 stlsss1) case (\Runs_table s2 s1 stlsss1) note bisim = \s1 \ s2\ and run1 = \trsys1.\Runs_table s1 stlsss1\ from run1 show ?case proof cases case (Terminate s1') let ?P = "\s2'. s2 -\2\* s2' \ (\tl2 s2''. \ s2' -2-tl2\ s2'') \ s1' \ s2'" from simulation_silents1[OF bisim \s1 -\1\* s1'\] obtain s2' where "s2 -\2\* s2'" and "s1' \ s2'" by blast moreover from no_move1_to_no_move2[OF \s1' \ s2'\ \\tl1 s1''. \ s1' -1-tl1\ s1''\] obtain s2'' where "s2' -\2\* s2''" and "s1' \ s2''" and "\tl2 s2'''. \ s2'' -2-tl2\ s2'''" by blast ultimately have "?P s2''" by(blast intro: rtranclp_trans) hence "?P (Eps ?P)" by(rule someI) hence ?Terminate using \stlsss1 = TNil \s1'\\ by simp thus ?thesis .. next case Diverge with \diverge_bisim_inv[OF bisim] have ?Diverge by simp thus ?thesis by simp next case (Proceed s1' s1'' stlsss1' tl1) let ?P = "\(tl2, s2''). \s2'. s2 -\2\* s2' \ s2' -2-tl2\ s2'' \ \ \move2 s2' tl2 s2'' \ s1'' \ s2'' \ tl1 \ tl2" from simulation_silents1[OF bisim \s1 -\1\* s1'\] obtain s2' where "s2 -\2\* s2'" and "s1' \ s2'" by blast moreover from simulation1[OF \s1' \ s2'\ \s1' -1-tl1\ s1''\ \\ \move1 s1' tl1 s1''\] obtain s2'' s2''' tl2 where "s2' -\2\* s2''" and "s2'' -2-tl2\ s2'''" and "\ \move2 s2'' tl2 s2'''" and "s1'' \ s2'''" and "tl1 \ tl2" by blast ultimately have "?P (tl2, s2''')" by(blast intro: rtranclp_trans) hence "?P (Eps ?P)" by(rule someI) hence ?Proceed using \stlsss1 = TCons (tl1, s1'') stlsss1'\ \trsys1.\Runs_table s1'' stlsss1'\ by auto blast thus ?thesis by simp qed qed let ?Tlsim = "\(tl1, s1'') (tl2, s2''). tl1 \ tl2 \ s1'' \ s2''" let ?Bisim = "rel_option bisim" from run1 bisim show "tllist_all2 ?Tlsim ?Bisim stlsss1 (tls1_to_tls2 s2 stlsss1)" proof(coinduction arbitrary: s1 s2 stlsss1) case (tllist_all2 s1 s2 stlsss1) note Runs = \trsys1.\Runs_table s1 stlsss1\ and bisim = \s1 \ s2\ from Runs show ?case proof cases case (Terminate s1') let ?P = "\s2'. s2 -\2\* s2' \ (\tl2 s2''. \ s2' -2-tl2\ s2'') \ s1' \ s2'" from simulation_silents1[OF bisim \s1 -\1\* s1'\] obtain s2' where "s2 -\2\* s2'" and "s1' \ s2'" by blast moreover from no_move1_to_no_move2[OF \s1' \ s2'\ \\tl1 s1''. \ s1' -1-tl1\ s1''\] obtain s2'' where "s2' -\2\* s2''" and "s1' \ s2''" and "\tl2 s2'''. \ s2'' -2-tl2\ s2'''" by blast ultimately have "?P s2''" by(blast intro: rtranclp_trans) hence "?P (Eps ?P)" by(rule someI) thus ?thesis using \stlsss1 = TNil \s1'\\ bisim by(simp) next case (Proceed s1' s1'' stlsss1' tl1) from simulation_silents1[OF bisim \s1 -\1\* s1'\] obtain s2' where "s2 -\2\* s2'" and "s1' \ s2'" by blast moreover from simulation1[OF \s1' \ s2'\ \s1' -1-tl1\ s1''\ \\ \move1 s1' tl1 s1''\] obtain s2'' s2''' tl2 where "s2' -\2\* s2''" and "s2'' -2-tl2\ s2'''" and "\ \move2 s2'' tl2 s2'''" and "s1'' \ s2'''" and "tl1 \ tl2" by blast ultimately have "?P s2 stlsss1 (tl2, s2''')" using \stlsss1 = TCons (tl1, s1'') stlsss1'\ by(auto intro: rtranclp_trans) hence "?P s2 stlsss1 (Eps (?P s2 stlsss1))" by(rule someI) thus ?thesis using \stlsss1 = TCons (tl1, s1'') stlsss1'\ \trsys1.\Runs_table s1'' stlsss1'\ bisim by auto blast qed simp qed qed lemma simulation_\Runs_table2: assumes "s1 \ s2" and "trsys2.\Runs_table s2 stlsss2" shows "\stlsss1. trsys1.\Runs_table s1 stlsss1 \ tllist_all2 (\(tl1, s1'') (tl2, s2''). tl1 \ tl2 \ s1'' \ s2'') (rel_option bisim) stlsss1 stlsss2" using delay_bisimulation_diverge.simulation_\Runs_table1[OF delay_bisimulation_diverge_flip, unfolded flip_simps, OF assms] by(subst tllist_all2_flip[symmetric])(simp only: flip_def split_def) lemma simulation_\Runs1: assumes bisim: "s1 \ s2" and run1: "s1 \1 tls1" shows "\tls2. s2 \2 tls2 \ tllist_all2 tlsim (rel_option bisim) tls1 tls2" proof - from trsys1.\Runs_into_\Runs_table[OF run1] obtain stlsss1 where tls1: "tls1 = tmap fst id stlsss1" and \Runs1: "trsys1.\Runs_table s1 stlsss1" by blast from simulation_\Runs_table1[OF bisim \Runs1] obtain stlsss2 where \Runs2: "trsys2.\Runs_table s2 stlsss2" and tlsim: "tllist_all2 (\(tl1, s1'') (tl2, s2''). tl1 \ tl2 \ s1'' \ s2'') (rel_option bisim) stlsss1 stlsss2" by blast from \Runs2 have "s2 \2 tmap fst id stlsss2" by(rule \Runs_table_into_\Runs) moreover have "tllist_all2 tlsim (rel_option bisim) tls1 (tmap fst id stlsss2)" using tlsim unfolding tls1 by(fastforce simp add: tllist_all2_tmap1 tllist_all2_tmap2 elim: tllist_all2_mono rel_option_mono) ultimately show ?thesis by blast qed lemma simulation_\Runs2: "\ s1 \ s2; s2 \2 tls2 \ \ \tls1. s1 \1 tls1 \ tllist_all2 tlsim (rel_option bisim) tls1 tls2" using delay_bisimulation_diverge.simulation_\Runs1[OF delay_bisimulation_diverge_flip] unfolding flip_simps . end locale delay_bisimulation_final_base = delay_bisimulation_base _ _ _ _ \move1 \move2 + bisimulation_final_base _ _ _ _ final1 final2 for \move1 :: "('s1, 'tl1) trsys" and \move2 :: "('s2, 'tl2) trsys" and final1 :: "'s1 \ bool" and final2 :: "'s2 \ bool" + assumes final1_simulation: "\ s1 \ s2; final1 s1 \ \ \s2'. s2 -\2\* s2' \ s1 \ s2' \ final2 s2'" and final2_simulation: "\ s1 \ s2; final2 s2 \ \ \s1'. s1 -\1\* s1' \ s1' \ s2 \ final1 s1'" begin lemma delay_bisimulation_final_base_flip: "delay_bisimulation_final_base trsys2 trsys1 (flip bisim) \move2 \move1 final2 final1" apply(unfold_locales) apply(unfold flip_simps) by(blast intro: final1_simulation final2_simulation)+ end lemma delay_bisimulation_final_base_flip_simps [flip_simps]: "delay_bisimulation_final_base trsys2 trsys1 (flip bisim) \move2 \move1 final2 final1 = delay_bisimulation_final_base trsys1 trsys2 bisim \move1 \move2 final1 final2" by(auto dest: delay_bisimulation_final_base.delay_bisimulation_final_base_flip simp only: flip_flip) context delay_bisimulation_final_base begin lemma \Runs_terminate_final1: assumes "s1 \1 tls1" and "s2 \2 tls2" and "tllist_all2 tlsim (rel_option bisim) tls1 tls2" and "tfinite tls1" and "terminal tls1 = Some s1'" and "final1 s1'" shows "\s2'. tfinite tls2 \ terminal tls2 = Some s2' \ final2 s2'" using assms(4) assms(1-3,5-) apply(induct arbitrary: tls2 s1 s2 rule: tfinite_induct) -apply(auto 4 4 simp add: tllist_all2_TCons1 tllist_all2_TNil1 rel_option_Some1 trsys1.\Runs_simps trsys2.\Runs_simps dest: final1_simulation elim: converse_rtranclpE) +apply(auto 4 4 simp add: tllist_all2_TCons1 tllist_all2_TNil1 option_rel_Some1 trsys1.\Runs_simps trsys2.\Runs_simps dest: final1_simulation elim: converse_rtranclpE) done lemma \Runs_terminate_final2: "\ s1 \1 tls1; s2 \2 tls2; tllist_all2 tlsim (rel_option bisim) tls1 tls2; tfinite tls2; terminal tls2 = Some s2'; final2 s2' \ \ \s1'. tfinite tls1 \ terminal tls1 = Some s1' \ final1 s1'" using delay_bisimulation_final_base.\Runs_terminate_final1[where tlsim = "flip tlsim", OF delay_bisimulation_final_base_flip] unfolding flip_simps by - end locale delay_bisimulation_diverge_final = delay_bisimulation_diverge + delay_bisimulation_final_base + constrains trsys1 :: "('s1, 'tl1) trsys" and trsys2 :: "('s2, 'tl2) trsys" and bisim :: "('s1, 's2) bisim" and tlsim :: "('tl1, 'tl2) bisim" and \move1 :: "('s1, 'tl1) trsys" and \move2 :: "('s2, 'tl2) trsys" and final1 :: "'s1 \ bool" and final2 :: "'s2 \ bool" begin lemma delay_bisimulation_diverge_final_flip: "delay_bisimulation_diverge_final trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1 final2 final1" apply(rule delay_bisimulation_diverge_final.intro) apply(rule delay_bisimulation_diverge_flip) apply(unfold_locales, unfold flip_simps) apply(blast intro: final1_simulation final2_simulation)+ done end lemma delay_bisimulation_diverge_final_flip_simps [flip_simps]: "delay_bisimulation_diverge_final trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1 final2 final1 = delay_bisimulation_diverge_final trsys1 trsys2 bisim tlsim \move1 \move2 final1 final2" by(auto dest: delay_bisimulation_diverge_final.delay_bisimulation_diverge_final_flip simp only: flip_flip) context delay_bisimulation_diverge_final begin lemma delay_bisimulation_diverge: "delay_bisimulation_diverge trsys1 trsys2 bisim tlsim \move1 \move2" by(unfold_locales) lemma delay_bisimulation_final_base: "delay_bisimulation_final_base trsys1 trsys2 bisim \move1 \move2 final1 final2" by(unfold_locales) lemma final_simulation1: "\ s1 \ s2; s1 -\1-tls1\* s1'; final1 s1' \ \ \s2' tls2. s2 -\2-tls2\* s2' \ s1' \ s2' \ final2 s2' \ tls1 [\] tls2" by(blast dest: simulation1_\rtrancl3p final1_simulation intro: \rtrancl3p_trans[OF _ silent_moves_into_\rtrancl3p, simplified]) lemma final_simulation2: "\ s1 \ s2; s2 -\2-tls2\* s2'; final2 s2' \ \ \s1' tls1. s1 -\1-tls1\* s1' \ s1' \ s2' \ final1 s1' \ tls1 [\] tls2" by(rule delay_bisimulation_diverge_final.final_simulation1[OF delay_bisimulation_diverge_final_flip, unfolded flip_simps]) end locale delay_bisimulation_measure_base = delay_bisimulation_base + constrains trsys1 :: "'s1 \ 'tl1 \ 's1 \ bool" and trsys2 :: "'s2 \ 'tl2 \ 's2 \ bool" and bisim :: "'s1 \ 's2 \ bool" and tlsim :: "'tl1 \ 'tl2 \ bool" and \move1 :: "'s1 \ 'tl1 \ 's1 \ bool" and \move2 :: "'s2 \ 'tl2 \ 's2 \ bool" fixes \1 :: "'s1 \ 's1 \ bool" and \2 :: "'s2 \ 's2 \ bool" locale delay_bisimulation_measure = delay_bisimulation_measure_base _ _ _ _ \move1 \move2 \1 \2 + delay_bisimulation_obs trsys1 trsys2 bisim tlsim \move1 \move2 for \move1 :: "'s1 \ 'tl1 \ 's1 \ bool" and \move2 :: "'s2 \ 'tl2 \ 's2 \ bool" and \1 :: "'s1 \ 's1 \ bool" and \2 :: "'s2 \ 's2 \ bool" + assumes simulation_silent1: "\ s1 \ s2; s1 -\1\ s1' \ \ s1' \ s2 \ \1^++ s1' s1 \ (\s2'. s2 -\2\+ s2' \ s1' \ s2')" and simulation_silent2: "\ s1 \ s2; s2 -\2\ s2' \ \ s1 \ s2' \ \2^++ s2' s2 \ (\s1'. s1 -\1\+ s1' \ s1' \ s2')" and wf_\1: "wfP \1" and wf_\2: "wfP \2" begin lemma delay_bisimulation_measure_flip: "delay_bisimulation_measure trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1 \2 \1" apply(rule delay_bisimulation_measure.intro) apply(rule delay_bisimulation_obs_flip) apply(unfold_locales) apply(unfold flip_simps) apply(rule simulation_silent1 simulation_silent2 wf_\1 wf_\2|assumption)+ done end lemma delay_bisimulation_measure_flip_simps [flip_simps]: "delay_bisimulation_measure trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1 \2 \1 = delay_bisimulation_measure trsys1 trsys2 bisim tlsim \move1 \move2 \1 \2" by(auto dest: delay_bisimulation_measure.delay_bisimulation_measure_flip simp only: flip_simps) context delay_bisimulation_measure begin lemma simulation_silentst1: assumes bisim: "s1 \ s2" and moves: "s1 -\1\+ s1'" shows "s1' \ s2 \ \1^++ s1' s1 \ (\s2'. s2 -\2\+ s2' \ s1' \ s2')" using moves bisim proof induct case (base s1') thus ?case by(auto dest: simulation_silent1) next case (step s1' s1'') hence "s1' \ s2 \ \1\<^sup>+\<^sup>+ s1' s1 \ (\s2'. s2 -\2\+ s2' \ s1' \ s2')" by blast thus ?case proof assume "s1' \ s2 \ \1\<^sup>+\<^sup>+ s1' s1" hence "s1' \ s2" "\1\<^sup>+\<^sup>+ s1' s1" by simp_all with simulation_silent1[OF \s1' \ s2\ \s1' -\1\ s1''\] show ?thesis by(auto) next assume "\s2'. trsys2.silent_move\<^sup>+\<^sup>+ s2 s2' \ s1' \ s2'" then obtain s2' where "s2 -\2\+ s2'" "s1' \ s2'" by blast with simulation_silent1[OF \s1' \ s2'\ \s1' -\1\ s1''\] show ?thesis by(auto intro: tranclp_trans) qed qed lemma simulation_silentst2: "\ s1 \ s2; s2 -\2\+ s2' \ \ s1 \ s2' \ \2^++ s2' s2 \ (\s1'. s1 -\1\+ s1' \ s1' \ s2')" using delay_bisimulation_measure.simulation_silentst1[OF delay_bisimulation_measure_flip] unfolding flip_simps . lemma \diverge_simulation1: assumes diverge1: "s1 -\1\ \" and bisim: "s1 \ s2" shows "s2 -\2\ \" proof - from assms have "s1 -\1\ \ \ s1 \ s2" by blast thus ?thesis using wfP_trancl[OF wf_\1] proof(coinduct rule: trsys2.\diverge_trancl_measure_coinduct) case (\diverge s2 s1) hence "s1 -\1\ \" "s1 \ s2" by simp_all then obtain s1' where "trsys1.silent_move s1 s1'" "s1' -\1\ \" by(fastforce elim: trsys1.\diverge.cases) from simulation_silent1[OF \s1 \ s2\ \trsys1.silent_move s1 s1'\] \s1' -\1\ \\ show ?case by auto qed qed lemma \diverge_simulation2: "\ s2 -\2\ \; s1 \ s2 \ \ s1 -\1\ \" using delay_bisimulation_measure.\diverge_simulation1[OF delay_bisimulation_measure_flip] unfolding flip_simps . lemma \diverge_bisim_inv: "s1 \ s2 \ s1 -\1\ \ \ s2 -\2\ \" by(blast intro: \diverge_simulation1 \diverge_simulation2) end sublocale delay_bisimulation_measure < delay_bisimulation_diverge proof fix s1 s2 s1' assume "s1 \ s2" "s1 -\1\ s1'" from simulation_silent1[OF this] show "\s2'. s2 -\2\* s2' \ s1' \ s2'" by(auto intro: tranclp_into_rtranclp) next fix s1 s2 s2' assume "s1 \ s2" "s2 -\2\ s2'" from simulation_silent2[OF this] show "\s1'. s1 -\1\* s1' \ s1' \ s2'" by(auto intro: tranclp_into_rtranclp) next fix s1 s2 assume "s1 \ s2" thus "s1 -\1\ \ \ s2 -\2\ \" by(rule \diverge_bisim_inv) qed text \ Counter example for @{prop "delay_bisimulation_diverge trsys1 trsys2 bisim tlsim \move1 \move2 \ \\1 \2. delay_bisimulation_measure trsys1 trsys2 bisim tlsim \move1 \move2 \1 \2"} (only \\\moves): \begin{verbatim} --| | v --a ~ x | | | | v v --b ~ y-- | ^ ^ | --| |-- \end{verbatim} \ locale delay_bisimulation_measure_final = delay_bisimulation_measure + delay_bisimulation_final_base + constrains trsys1 :: "('s1, 'tl1) trsys" and trsys2 :: "('s2, 'tl2) trsys" and bisim :: "('s1, 's2) bisim" and tlsim :: "('tl1, 'tl2) bisim" and \move1 :: "('s1, 'tl1) trsys" and \move2 :: "('s2, 'tl2) trsys" and \1 :: "'s1 \ 's1 \ bool" and \2 :: "'s2 \ 's2 \ bool" and final1 :: "'s1 \ bool" and final2 :: "'s2 \ bool" sublocale delay_bisimulation_measure_final < delay_bisimulation_diverge_final by unfold_locales locale \inv = delay_bisimulation_base + constrains trsys1 :: "('s1, 'tl1) trsys" and trsys2 :: "('s2, 'tl2) trsys" and bisim :: "('s1, 's2) bisim" and tlsim :: "('tl1, 'tl2) bisim" and \move1 :: "('s1, 'tl1) trsys" and \move2 :: "('s2, 'tl2) trsys" and \moves1 :: "'s1 \ 's1 \ bool" and \moves2 :: "'s2 \ 's2 \ bool" assumes \inv: "\ s1 \ s2; s1 -1-tl1\ s1'; s2 -2-tl2\ s2'; s1' \ s2'; tl1 \ tl2 \ \ \move1 s1 tl1 s1' \ \move2 s2 tl2 s2'" begin lemma \inv_flip: "\inv trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1" by(unfold_locales)(unfold flip_simps,rule \inv[symmetric]) end lemma \inv_flip_simps [flip_simps]: "\inv trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1 = \inv trsys1 trsys2 bisim tlsim \move1 \move2" by(auto dest: \inv.\inv_flip simp only: flip_simps) locale bisimulation_into_delay = bisimulation + \inv + constrains trsys1 :: "('s1, 'tl1) trsys" and trsys2 :: "('s2, 'tl2) trsys" and bisim :: "('s1, 's2) bisim" and tlsim :: "('tl1, 'tl2) bisim" and \move1 :: "('s1, 'tl1) trsys" and \move2 :: "('s2, 'tl2) trsys" begin lemma bisimulation_into_delay_flip: "bisimulation_into_delay trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1" by(intro_locales)(intro bisimulation_flip \inv_flip)+ end lemma bisimulation_into_delay_flip_simps [flip_simps]: "bisimulation_into_delay trsys2 trsys1 (flip bisim) (flip tlsim) \move2 \move1 = bisimulation_into_delay trsys1 trsys2 bisim tlsim \move1 \move2" by(auto dest: bisimulation_into_delay.bisimulation_into_delay_flip simp only: flip_simps) context bisimulation_into_delay begin lemma simulation_silent1_aux: assumes bisim: "s1 \ s2" and "s1 -\1\ s1'" shows "s1' \ s2 \ \1\<^sup>+\<^sup>+ s1' s1 \ (\s2'. s2 -\2\+ s2' \ s1' \ s2')" proof - from assms obtain tl1 where tr1: "s1 -1-tl1\ s1'" and \1: "\move1 s1 tl1 s1'" by(auto) from simulation1[OF bisim tr1] obtain s2' tl2 where tr2: "s2 -2-tl2\ s2'" and bisim': "s1' \ s2'" and tlsim: "tl1 \ tl2" by blast from \inv[OF bisim tr1 tr2 bisim' tlsim] \1 have \2: "\move2 s2 tl2 s2'" by simp from tr2 \2 have "s2 -\2\+ s2'" by(auto) with bisim' show ?thesis by blast qed lemma simulation_silent2_aux: "\ s1 \ s2; s2 -\2\ s2' \ \ s1 \ s2' \ \2\<^sup>+\<^sup>+ s2' s2 \ (\s1'. s1 -\1\+ s1' \ s1' \ s2')" using bisimulation_into_delay.simulation_silent1_aux[OF bisimulation_into_delay_flip] unfolding flip_simps . lemma simulation1_aux: assumes bisim: "s1 \ s2" and tr1: "s1 -1-tl1\ s1'" and \1: "\ \move1 s1 tl1 s1'" shows "\s2' s2'' tl2. s2 -\2\* s2' \ s2' -2-tl2\ s2'' \ \ \move2 s2' tl2 s2'' \ s1' \ s2'' \ tl1 \ tl2" proof - from simulation1[OF bisim tr1] obtain s2' tl2 where tr2: "s2 -2-tl2\ s2'" and bisim': "s1' \ s2'" and tlsim: "tl1 \ tl2" by blast from \inv[OF bisim tr1 tr2 bisim' tlsim] \1 have \2: "\ \move2 s2 tl2 s2'" by simp with bisim' tr2 tlsim show ?thesis by blast qed lemma simulation2_aux: "\ s1 \ s2; s2 -2-tl2\ s2'; \ \move2 s2 tl2 s2' \ \ \s1' s1'' tl1. s1 -\1\* s1' \ s1' -1-tl1\ s1'' \ \ \move1 s1' tl1 s1'' \ s1'' \ s2' \ tl1 \ tl2" using bisimulation_into_delay.simulation1_aux[OF bisimulation_into_delay_flip] unfolding flip_simps . lemma delay_bisimulation_measure: assumes wf_\1: "wfP \1" and wf_\2: "wfP \2" shows "delay_bisimulation_measure trsys1 trsys2 bisim tlsim \move1 \move2 \1 \2" apply(unfold_locales) apply(rule simulation_silent1_aux simulation_silent2_aux simulation1_aux simulation2_aux wf_\1 wf_\2|assumption)+ done lemma delay_bisimulation: "delay_bisimulation_diverge trsys1 trsys2 bisim tlsim \move1 \move2" proof - interpret delay_bisimulation_measure trsys1 trsys2 bisim tlsim \move1 \move2 "\s s'. False" "\s s'. False" by(blast intro: delay_bisimulation_measure wfP_empty) show ?thesis .. qed end sublocale bisimulation_into_delay < delay_bisimulation_diverge by(rule delay_bisimulation) lemma delay_bisimulation_conv_bisimulation: "delay_bisimulation_diverge trsys1 trsys2 bisim tlsim (\s tl s'. False) (\s tl s'. False) = bisimulation trsys1 trsys2 bisim tlsim" (is "?lhs = ?rhs") proof assume ?lhs then interpret delay_bisimulation_diverge trsys1 trsys2 bisim tlsim "\s tl s'. False" "\s tl s'. False" . show ?rhs by(unfold_locales)(fastforce simp add: \moves_False dest: simulation1 simulation2)+ next assume ?rhs then interpret bisimulation trsys1 trsys2 bisim tlsim . interpret bisimulation_into_delay trsys1 trsys2 bisim tlsim "\s tl s'. False" "\s tl s'. False" by(unfold_locales)(rule refl) show ?lhs by unfold_locales qed context bisimulation_final begin lemma delay_bisimulation_final_base: "delay_bisimulation_final_base trsys1 trsys2 bisim \move1 \move2 final1 final2" by(unfold_locales)(auto simp add: bisim_final) end sublocale bisimulation_final < delay_bisimulation_final_base by(rule delay_bisimulation_final_base) subsection \Transitivity for bisimulations\ definition bisim_compose :: "('s1, 's2) bisim \ ('s2, 's3) bisim \ ('s1, 's3) bisim" (infixr "\\<^sub>B" 60) where "(bisim1 \\<^sub>B bisim2) s1 s3 \ \s2. bisim1 s1 s2 \ bisim2 s2 s3" lemma bisim_composeI [intro]: "\ bisim12 s1 s2; bisim23 s2 s3 \ \ (bisim12 \\<^sub>B bisim23) s1 s3" by(auto simp add: bisim_compose_def) lemma bisim_composeE [elim!]: assumes bisim: "(bisim12 \\<^sub>B bisim23) s1 s3" obtains s2 where "bisim12 s1 s2" "bisim23 s2 s3" by(atomize_elim)(rule bisim[unfolded bisim_compose_def]) lemma bisim_compose_assoc [simp]: "(bisim12 \\<^sub>B bisim23) \\<^sub>B bisim34 = bisim12 \\<^sub>B bisim23 \\<^sub>B bisim34" by(auto simp add: fun_eq_iff) lemma bisim_compose_conv_relcomp: "case_prod (bisim_compose bisim12 bisim23) = (\x. x \ relcomp (Collect (case_prod bisim12)) (Collect (case_prod bisim23)))" by(auto simp add: relcomp_unfold) lemma list_all2_bisim_composeI: "\ list_all2 A xs ys; list_all2 B ys zs \ \ list_all2 (A \\<^sub>B B) xs zs" by(rule list_all2_trans) auto+ lemma delay_bisimulation_diverge_compose: assumes wbisim12: "delay_bisimulation_diverge trsys1 trsys2 bisim12 tlsim12 \move1 \move2" and wbisim23: "delay_bisimulation_diverge trsys2 trsys3 bisim23 tlsim23 \move2 \move3" shows "delay_bisimulation_diverge trsys1 trsys3 (bisim12 \\<^sub>B bisim23) (tlsim12 \\<^sub>B tlsim23) \move1 \move3" proof - interpret trsys1: \trsys trsys1 \move1 . interpret trsys2: \trsys trsys2 \move2 . interpret trsys3: \trsys trsys3 \move3 . interpret wb12: delay_bisimulation_diverge trsys1 trsys2 bisim12 tlsim12 \move1 \move2 by(auto intro: wbisim12) interpret wb23: delay_bisimulation_diverge trsys2 trsys3 bisim23 tlsim23 \move2 \move3 by(auto intro: wbisim23) show ?thesis proof fix s1 s3 s1' assume bisim: "(bisim12 \\<^sub>B bisim23) s1 s3" and tr1: "trsys1.silent_move s1 s1'" from bisim obtain s2 where bisim1: "bisim12 s1 s2" and bisim2: "bisim23 s2 s3" by blast from wb12.simulation_silent1[OF bisim1 tr1] obtain s2' where tr2: "trsys2.silent_moves s2 s2'" and bisim1': "bisim12 s1' s2'" by blast from wb23.simulation_silents1[OF bisim2 tr2] obtain s3' where "trsys3.silent_moves s3 s3'" "bisim23 s2' s3'" by blast with bisim1' show "\s3'. trsys3.silent_moves s3 s3' \ (bisim12 \\<^sub>B bisim23) s1' s3'" by(blast intro: bisim_composeI) next fix s1 s3 s3' assume bisim: "(bisim12 \\<^sub>B bisim23) s1 s3" and tr3: "trsys3.silent_move s3 s3'" from bisim obtain s2 where bisim1: "bisim12 s1 s2" and bisim2: "bisim23 s2 s3" by blast from wb23.simulation_silent2[OF bisim2 tr3] obtain s2' where tr2: "trsys2.silent_moves s2 s2'" and bisim2': "bisim23 s2' s3'" by blast from wb12.simulation_silents2[OF bisim1 tr2] obtain s1' where "trsys1.silent_moves s1 s1'" "bisim12 s1' s2'" by blast with bisim2' show "\s1'. trsys1.silent_moves s1 s1' \ (bisim12 \\<^sub>B bisim23) s1' s3'" by(blast intro: bisim_composeI) next fix s1 s3 tl1 s1' assume bisim: "(bisim12 \\<^sub>B bisim23) s1 s3" and tr1: "trsys1 s1 tl1 s1'" and \1: "\ \move1 s1 tl1 s1'" from bisim obtain s2 where bisim1: "bisim12 s1 s2" and bisim2: "bisim23 s2 s3" by blast from wb12.simulation1[OF bisim1 tr1 \1] obtain s2' s2'' tl2 where tr21: "trsys2.silent_moves s2 s2'" and tr22: "trsys2 s2' tl2 s2''" and \2: "\ \move2 s2' tl2 s2''" and bisim1': "bisim12 s1' s2''" and tlsim1: "tlsim12 tl1 tl2" by blast from wb23.simulation_silents1[OF bisim2 tr21] obtain s3' where tr31: "trsys3.silent_moves s3 s3'" and bisim2': "bisim23 s2' s3'" by blast from wb23.simulation1[OF bisim2' tr22 \2] obtain s3'' s3''' tl3 where "trsys3.silent_moves s3' s3''" "trsys3 s3'' tl3 s3'''" "\ \move3 s3'' tl3 s3'''" "bisim23 s2'' s3'''" "tlsim23 tl2 tl3" by blast with tr31 bisim1' tlsim1 show "\s3' s3'' tl3. trsys3.silent_moves s3 s3' \ trsys3 s3' tl3 s3'' \ \ \move3 s3' tl3 s3'' \ (bisim12 \\<^sub>B bisim23) s1' s3'' \ (tlsim12 \\<^sub>B tlsim23) tl1 tl3" by(blast intro: rtranclp_trans bisim_composeI) next fix s1 s3 tl3 s3' assume bisim: "(bisim12 \\<^sub>B bisim23) s1 s3" and tr3: "trsys3 s3 tl3 s3'" and \3: "\ \move3 s3 tl3 s3'" from bisim obtain s2 where bisim1: "bisim12 s1 s2" and bisim2: "bisim23 s2 s3" by blast from wb23.simulation2[OF bisim2 tr3 \3] obtain s2' s2'' tl2 where tr21: "trsys2.silent_moves s2 s2'" and tr22: "trsys2 s2' tl2 s2''" and \2: "\ \move2 s2' tl2 s2''" and bisim2': "bisim23 s2'' s3'" and tlsim2: "tlsim23 tl2 tl3" by blast from wb12.simulation_silents2[OF bisim1 tr21] obtain s1' where tr11: "trsys1.silent_moves s1 s1'" and bisim1': "bisim12 s1' s2'" by blast from wb12.simulation2[OF bisim1' tr22 \2] obtain s1'' s1''' tl1 where "trsys1.silent_moves s1' s1''" "trsys1 s1'' tl1 s1'''" "\ \move1 s1'' tl1 s1'''" "bisim12 s1''' s2''" "tlsim12 tl1 tl2" by blast with tr11 bisim2' tlsim2 show "\s1' s1'' tl1. trsys1.silent_moves s1 s1' \ trsys1 s1' tl1 s1'' \ \ \move1 s1' tl1 s1'' \ (bisim12 \\<^sub>B bisim23) s1'' s3' \ (tlsim12 \\<^sub>B tlsim23) tl1 tl3" by(blast intro: rtranclp_trans bisim_composeI) next fix s1 s2 assume "(bisim12 \\<^sub>B bisim23) s1 s2" thus "\trsys.\diverge trsys1 \move1 s1 = \trsys.\diverge trsys3 \move3 s2" by(auto simp add: wb12.\diverge_bisim_inv wb23.\diverge_bisim_inv) qed qed lemma bisimulation_bisim_compose: "\ bisimulation trsys1 trsys2 bisim12 tlsim12; bisimulation trsys2 trsys3 bisim23 tlsim23 \ \ bisimulation trsys1 trsys3 (bisim_compose bisim12 bisim23) (bisim_compose tlsim12 tlsim23)" unfolding delay_bisimulation_conv_bisimulation[symmetric] by(rule delay_bisimulation_diverge_compose) lemma delay_bisimulation_diverge_final_compose: fixes \move1 \move2 assumes wbisim12: "delay_bisimulation_diverge_final trsys1 trsys2 bisim12 tlsim12 \move1 \move2 final1 final2" and wbisim23: "delay_bisimulation_diverge_final trsys2 trsys3 bisim23 tlsim23 \move2 \move3 final2 final3" shows "delay_bisimulation_diverge_final trsys1 trsys3 (bisim12 \\<^sub>B bisim23) (tlsim12 \\<^sub>B tlsim23) \move1 \move3 final1 final3" proof - interpret trsys1: \trsys trsys1 \move1 . interpret trsys2: \trsys trsys2 \move2 . interpret trsys3: \trsys trsys3 \move3 . interpret wb12: delay_bisimulation_diverge_final trsys1 trsys2 bisim12 tlsim12 \move1 \move2 final1 final2 by(auto intro: wbisim12) interpret wb23: delay_bisimulation_diverge_final trsys2 trsys3 bisim23 tlsim23 \move2 \move3 final2 final3 by(auto intro: wbisim23) interpret delay_bisimulation_diverge trsys1 trsys3 "bisim12 \\<^sub>B bisim23" "tlsim12 \\<^sub>B tlsim23" \move1 \move3 by(rule delay_bisimulation_diverge_compose)(unfold_locales) show ?thesis proof fix s1 s3 assume "(bisim12 \\<^sub>B bisim23) s1 s3" "final1 s1" from \(bisim12 \\<^sub>B bisim23) s1 s3\ obtain s2 where "bisim12 s1 s2" and "bisim23 s2 s3" .. from wb12.final1_simulation[OF \bisim12 s1 s2\ \final1 s1\] obtain s2' where "trsys2.silent_moves s2 s2'" "bisim12 s1 s2'" "final2 s2'" by blast from wb23.simulation_silents1[OF \bisim23 s2 s3\ \trsys2.silent_moves s2 s2'\] obtain s3' where "trsys3.silent_moves s3 s3'" "bisim23 s2' s3'" by blast from wb23.final1_simulation[OF \bisim23 s2' s3'\ \final2 s2'\] obtain s3'' where "trsys3.silent_moves s3' s3''" "bisim23 s2' s3''" "final3 s3''" by blast from \trsys3.silent_moves s3 s3'\ \trsys3.silent_moves s3' s3''\ have "trsys3.silent_moves s3 s3''" by(rule rtranclp_trans) moreover from \bisim12 s1 s2'\ \bisim23 s2' s3''\ have "(bisim12 \\<^sub>B bisim23) s1 s3''" .. ultimately show "\s3'. trsys3.silent_moves s3 s3' \ (bisim12 \\<^sub>B bisim23) s1 s3' \ final3 s3'" using \final3 s3''\ by iprover next fix s1 s3 assume "(bisim12 \\<^sub>B bisim23) s1 s3" "final3 s3" from \(bisim12 \\<^sub>B bisim23) s1 s3\ obtain s2 where "bisim12 s1 s2" and "bisim23 s2 s3" .. from wb23.final2_simulation[OF \bisim23 s2 s3\ \final3 s3\] obtain s2' where "trsys2.silent_moves s2 s2'" "bisim23 s2' s3" "final2 s2'" by blast from wb12.simulation_silents2[OF \bisim12 s1 s2\ \trsys2.silent_moves s2 s2'\] obtain s1' where "trsys1.silent_moves s1 s1'" "bisim12 s1' s2'" by blast from wb12.final2_simulation[OF \bisim12 s1' s2'\ \final2 s2'\] obtain s1'' where "trsys1.silent_moves s1' s1''" "bisim12 s1'' s2'" "final1 s1''" by blast from \trsys1.silent_moves s1 s1'\ \trsys1.silent_moves s1' s1''\ have "trsys1.silent_moves s1 s1''" by(rule rtranclp_trans) moreover from \bisim12 s1'' s2'\ \bisim23 s2' s3\ have "(bisim12 \\<^sub>B bisim23) s1'' s3" .. ultimately show "\s1'. trsys1.silent_moves s1 s1' \ (bisim12 \\<^sub>B bisim23) s1' s3 \ final1 s1'" using \final1 s1''\ by iprover qed qed end diff --git a/thys/JinjaThreads/Framework/LTS.thy b/thys/JinjaThreads/Framework/LTS.thy --- a/thys/JinjaThreads/Framework/LTS.thy +++ b/thys/JinjaThreads/Framework/LTS.thy @@ -1,1417 +1,1417 @@ (* Title: JinjaThreads/Framework/LTS.thy Author: Andreas Lochbihler *) section \Labelled transition systems\ theory LTS imports "../Basic/Auxiliary" Coinductive.TLList begin no_notation floor ("\_\") lemma rel_option_mono: "\ rel_option R x y; \x y. R x y \ R' x y \ \ rel_option R' x y" by(cases x)(case_tac [!] y, auto) lemma nth_concat_conv: "n < length (concat xss) \ \m n'. concat xss ! n = (xss ! m) ! n' \ n' < length (xss ! m) \ m < length xss \ n = (\i 'b \ 'c) \ 'b \ 'a \ 'c" where "flip f = (\b a. f a b)" text \Create a dynamic list \flip_simps\ of theorems for flip\ ML \ structure FlipSimpRules = Named_Thms ( val name = @{binding flip_simps} val description = "Simplification rules for flip in bisimulations" ) \ setup \FlipSimpRules.setup\ lemma flip_conv [flip_simps]: "flip f b a = f a b" by(simp add: flip_def) lemma flip_flip [flip_simps, simp]: "flip (flip f) = f" by(simp add: flip_def) lemma list_all2_flip [flip_simps]: "list_all2 (flip P) xs ys = list_all2 P ys xs" unfolding flip_def list_all2_conv_all_nth by auto lemma llist_all2_flip [flip_simps]: "llist_all2 (flip P) xs ys = llist_all2 P ys xs" unfolding flip_def llist_all2_conv_all_lnth by auto lemma rtranclp_flipD: assumes "(flip r)^** x y" shows "r^** y x" using assms by(induct rule: rtranclp_induct)(auto intro: rtranclp.rtrancl_into_rtrancl simp add: flip_conv) lemma rtranclp_flip [flip_simps]: "(flip r)^** = flip r^**" by(auto intro!: ext simp add: flip_conv intro: rtranclp_flipD) lemma rel_prod_flip [flip_simps]: "rel_prod (flip R) (flip S) = flip (rel_prod R S)" by(auto intro!: ext simp add: flip_def) lemma rel_option_flip [flip_simps]: "rel_option (flip R) = flip (rel_option R)" by(simp add: fun_eq_iff rel_option_iff flip_def) lemma tllist_all2_flip [flip_simps]: "tllist_all2 (flip P) (flip Q) xs ys \ tllist_all2 P Q ys xs" proof assume "tllist_all2 (flip P) (flip Q) xs ys" thus "tllist_all2 P Q ys xs" by(coinduct rule: tllist_all2_coinduct)(auto dest: tllist_all2_is_TNilD tllist_all2_tfinite2_terminalD tllist_all2_thdD intro: tllist_all2_ttlI simp add: flip_def) next assume "tllist_all2 P Q ys xs" thus "tllist_all2 (flip P) (flip Q) xs ys" by(coinduct rule: tllist_all2_coinduct)(auto dest: tllist_all2_is_TNilD tllist_all2_tfinite2_terminalD tllist_all2_thdD intro: tllist_all2_ttlI simp add: flip_def) qed subsection \Labelled transition systems\ type_synonym ('a, 'b) trsys = "'a \ 'b \ 'a \ bool" locale trsys = fixes trsys :: "('s, 'tl) trsys" ("_/ -_\/ _" [50, 0, 50] 60) begin abbreviation Trsys :: "('s, 'tl list) trsys" ("_/ -_\*/ _" [50,0,50] 60) where "\tl. s -tl\* s' \ rtrancl3p trsys s tl s'" coinductive inf_step :: "'s \ 'tl llist \ bool" ("_ -_\* \" [50, 0] 80) where inf_stepI: "\ trsys a b a'; a' -bs\* \ \ \ a -LCons b bs\* \" coinductive inf_step_table :: "'s \ ('s \ 'tl \ 's) llist \ bool" ("_ -_\*t \" [50, 0] 80) where inf_step_tableI: "\tl. \ trsys s tl s'; s' -stls\*t \ \ \ s -LCons (s, tl, s') stls\*t \" definition inf_step2inf_step_table :: "'s \ 'tl llist \ ('s \ 'tl \ 's) llist" where "inf_step2inf_step_table s tls = unfold_llist (\(s, tls). lnull tls) (\(s, tls). (s, lhd tls, SOME s'. trsys s (lhd tls) s' \ s' -ltl tls\* \)) (\(s, tls). (SOME s'. trsys s (lhd tls) s' \ s' -ltl tls\* \, ltl tls)) (s, tls)" coinductive Rtrancl3p :: "'s \ ('tl, 's) tllist \ bool" where Rtrancl3p_stop: "(\tl s'. \ s -tl\ s') \ Rtrancl3p s (TNil s)" | Rtrancl3p_into_Rtrancl3p: "\tl. \ s -tl\ s'; Rtrancl3p s' tlss \ \ Rtrancl3p s (TCons tl tlss)" inductive_simps Rtrancl3p_simps: "Rtrancl3p s (TNil s')" "Rtrancl3p s (TCons tl' tlss)" inductive_cases Rtrancl3p_cases: "Rtrancl3p s (TNil s')" "Rtrancl3p s (TCons tl' tlss)" coinductive Runs :: "'s \ 'tl llist \ bool" where Stuck: "(\tl s'. \ s -tl\ s') \ Runs s LNil" | Step: "\tl. \ s -tl\ s'; Runs s' tls \ \ Runs s (LCons tl tls)" coinductive Runs_table :: "'s \ ('s \ 'tl \ 's) llist \ bool" where Stuck: "(\tl s'. \ s -tl\ s') \ Runs_table s LNil" | Step: "\tl. \ s -tl\ s'; Runs_table s' stlss \ \ Runs_table s (LCons (s, tl, s') stlss)" inductive_simps Runs_table_simps: "Runs_table s LNil" "Runs_table s (LCons stls stlss)" lemma inf_step_not_finite_llist: assumes r: "s -bs\* \" shows "\ lfinite bs" proof assume "lfinite bs" thus False using r by(induct arbitrary: s rule: lfinite.induct)(auto elim: inf_step.cases) qed lemma inf_step2inf_step_table_LNil [simp]: "inf_step2inf_step_table s LNil = LNil" by(simp add: inf_step2inf_step_table_def) lemma inf_step2inf_step_table_LCons [simp]: fixes tl shows "inf_step2inf_step_table s (LCons tl tls) = LCons (s, tl, SOME s'. trsys s tl s' \ s' -tls\* \) (inf_step2inf_step_table (SOME s'. trsys s tl s' \ s' -tls\* \) tls)" by(simp add: inf_step2inf_step_table_def) lemma lnull_inf_step2inf_step_table [simp]: "lnull (inf_step2inf_step_table s tls) \ lnull tls" by(simp add: inf_step2inf_step_table_def) lemma inf_step2inf_step_table_eq_LNil: "inf_step2inf_step_table s tls = LNil \ tls = LNil" using lnull_inf_step2inf_step_table unfolding lnull_def . lemma lhd_inf_step2inf_step_table [simp]: "\ lnull tls \ lhd (inf_step2inf_step_table s tls) = (s, lhd tls, SOME s'. trsys s (lhd tls) s' \ s' -ltl tls\* \)" by(simp add: inf_step2inf_step_table_def) lemma ltl_inf_step2inf_step_table [simp]: "ltl (inf_step2inf_step_table s tls) = inf_step2inf_step_table (SOME s'. trsys s (lhd tls) s' \ s' -ltl tls\* \) (ltl tls)" by(cases tls) simp_all lemma lmap_inf_step2inf_step_table: "lmap (fst \ snd) (inf_step2inf_step_table s tls) = tls" by(coinduction arbitrary: s tls) auto lemma inf_step_imp_inf_step_table: assumes "s -tls\* \" shows "\stls. s -stls\*t \ \ tls = lmap (fst \ snd) stls" proof - from assms have "s -inf_step2inf_step_table s tls\*t \" proof(coinduction arbitrary: s tls) case (inf_step_table s tls) thus ?case proof cases case (inf_stepI tl s' tls') let ?s' = "SOME s'. trsys s tl s' \ s' -tls'\* \" have "trsys s tl ?s' \ ?s' -tls'\* \" by(rule someI)(blast intro: inf_stepI) thus ?thesis using \tls = LCons tl tls'\ by auto qed qed moreover have "tls = lmap (fst \ snd) (inf_step2inf_step_table s tls)" by(simp only: lmap_inf_step2inf_step_table) ultimately show ?thesis by blast qed lemma inf_step_table_imp_inf_step: "s-stls\*t \ \s -lmap (fst \ snd) stls\* \" proof(coinduction arbitrary: s stls rule: inf_step.coinduct) case (inf_step s tls) thus ?case by cases auto qed lemma Runs_table_into_Runs: "Runs_table s stlss \ Runs s (lmap (\(s, tl, s'). tl) stlss)" proof(coinduction arbitrary: s stlss) case (Runs s tls) thus ?case by (cases)auto qed lemma Runs_into_Runs_table: assumes "Runs s tls" obtains stlss where "tls = lmap (\(s, tl, s'). tl) stlss" and "Runs_table s stlss" proof - define stlss where "stlss s tls = unfold_llist (\(s, tls). lnull tls) (\(s, tls). (s, lhd tls, SOME s'. s -lhd tls\ s' \ Runs s' (ltl tls))) (\(s, tls). (SOME s'. s -lhd tls\ s' \ Runs s' (ltl tls), ltl tls)) (s, tls)" for s tls have [simp]: "\s. stlss s LNil = LNil" "\s tl tls. stlss s (LCons tl tls) = LCons (s, tl, SOME s'. s -tl\ s' \ Runs s' tls) (stlss (SOME s'. s -tl\ s' \ Runs s' tls) tls)" "\s tls. lnull (stlss s tls) \ lnull tls" "\s tls. \ lnull tls \ lhd (stlss s tls) = (s, lhd tls, SOME s'. s -lhd tls\ s' \ Runs s' (ltl tls))" "\s tls. \ lnull tls \ ltl (stlss s tls) = stlss (SOME s'. s -lhd tls\ s' \ Runs s' (ltl tls)) (ltl tls)" by(simp_all add: stlss_def) from assms have "tls = lmap (\(s, tl, s'). tl) (stlss s tls)" proof(coinduction arbitrary: s tls) case Eq_llist thus ?case by cases(auto 4 3 intro: someI2) qed moreover from assms have "Runs_table s (stlss s tls)" proof(coinduction arbitrary: s tls) case (Runs_table s stlss') thus ?case proof(cases) case (Step s' tls' tl) let ?P = "\s'. s -tl\ s' \ Runs s' tls'" from \s -tl\ s'\ \Runs s' tls'\ have "?P s'" .. hence "?P (Eps ?P)" by(rule someI) with Step have ?Step by auto thus ?thesis .. qed simp qed ultimately show ?thesis by(rule that) qed lemma Runs_lappendE: assumes "Runs \ (lappend tls tls')" and "lfinite tls" obtains \' where "\ -list_of tls\* \'" and "Runs \' tls'" proof(atomize_elim) from \lfinite tls\ \Runs \ (lappend tls tls')\ show "\\'. \ -list_of tls\* \' \ Runs \' tls'" proof(induct arbitrary: \) case lfinite_LNil thus ?case by(auto) next case (lfinite_LConsI tls tl) from \Runs \ (lappend (LCons tl tls) tls')\ show ?case unfolding lappend_code proof(cases) case (Step \') from \Runs \' (lappend tls tls') \ \\''. \' -list_of tls\* \'' \ Runs \'' tls'\ \Runs \' (lappend tls tls')\ obtain \'' where "\' -list_of tls\* \''" "Runs \'' tls'" by blast from \\ -tl\ \'\ \\' -list_of tls\* \''\ have "\ -tl # list_of tls\* \''" by(rule rtrancl3p_step_converse) with \lfinite tls\ have "\ -list_of (LCons tl tls)\* \''" by(simp) with \Runs \'' tls'\ show ?thesis by blast qed qed qed lemma Trsys_into_Runs: assumes "s -tls\* s'" and "Runs s' tls'" shows "Runs s (lappend (llist_of tls) tls')" using assms by(induct rule: rtrancl3p_converse_induct)(auto intro: Runs.Step) lemma rtrancl3p_into_Rtrancl3p: "\ rtrancl3p trsys a bs a'; \b a''. \ a' -b\ a'' \ \ Rtrancl3p a (tllist_of_llist a' (llist_of bs))" by(induct rule: rtrancl3p_converse_induct)(auto intro: Rtrancl3p.intros) lemma Rtrancl3p_into_Runs: "Rtrancl3p s tlss \ Runs s (llist_of_tllist tlss)" by(coinduction arbitrary: s tlss rule: Runs.coinduct)(auto elim: Rtrancl3p.cases) lemma Runs_into_Rtrancl3p: assumes "Runs s tls" obtains tlss where "tls = llist_of_tllist tlss" "Rtrancl3p s tlss" proof let ?Q = "\s tls s'. s -lhd tls\ s' \ Runs s' (ltl tls)" define tlss where "tlss = corec_tllist (\(s, tls). lnull tls) (\(s, tls). s) (\(s, tls). lhd tls) (\_. False) undefined (\(s, tls). (SOME s'. ?Q s tls s', ltl tls))" have [simp]: "tlss (s, LNil) = TNil s" "tlss (s, LCons tl tls) = TCons tl (tlss (SOME s'. ?Q s (LCons tl tls) s', tls))" for s tl tls by(auto simp add: tlss_def intro: tllist.expand) show "tls = llist_of_tllist (tlss (s, tls))" using assms by(coinduction arbitrary: s tls)(erule Runs.cases; fastforce intro: someI2) show "Rtrancl3p s (tlss (s, tls))" using assms by(coinduction arbitrary: s tls)(erule Runs.cases; simp; iprover intro: someI2[where Q="trsys _ _"] someI2[where Q="\s'. Runs s' _"]) qed lemma fixes tl assumes "Rtrancl3p s tlss" "tfinite tlss" shows Rtrancl3p_into_Trsys: "Trsys s (list_of (llist_of_tllist tlss)) (terminal tlss)" and terminal_Rtrancl3p_final: "\ terminal tlss -tl\ s'" using assms(2,1) by(induction arbitrary: s rule: tfinite_induct)(auto simp add: Rtrancl3p_simps intro: rtrancl3p_step_converse) end subsection \Labelled transition systems with internal actions\ locale \trsys = trsys + constrains trsys :: "('s, 'tl) trsys" fixes \move :: "('s, 'tl) trsys" begin inductive silent_move :: "'s \ 's \ bool" ("_ -\\ _" [50, 50] 60) where [intro]: "!!tl. \ trsys s tl s'; \move s tl s' \ \ s -\\ s'" declare silent_move.cases [elim] lemma silent_move_iff: "silent_move = (\s s'. (\tl. trsys s tl s' \ \move s tl s'))" by(auto simp add: fun_eq_iff) abbreviation silent_moves :: "'s \ 's \ bool" ("_ -\\* _" [50, 50] 60) where "silent_moves == silent_move^**" abbreviation silent_movet :: "'s \ 's \ bool" ("_ -\\+ _" [50, 50] 60) where "silent_movet == silent_move^++" coinductive \diverge :: "'s \ bool" ("_ -\\ \" [50] 60) where \divergeI: "\ s -\\ s'; s' -\\ \ \ \ s -\\ \" coinductive \inf_step :: "'s \ 'tl llist \ bool" ("_ -\-_\* \" [50, 0] 60) where \inf_step_Cons: "\tl. \ s -\\* s'; s' -tl\ s''; \ \move s' tl s''; s'' -\-tls\* \ \ \ s -\-LCons tl tls\* \" | \inf_step_Nil: "s -\\ \ \ s -\-LNil\* \" coinductive \inf_step_table :: "'s \ ('s \ 's \ 'tl \ 's) llist \ bool" ("_ -\-_\*t \" [50, 0] 80) where \inf_step_table_Cons: "\tl. \ s -\\* s'; s' -tl\ s''; \ \move s' tl s''; s'' -\-tls\*t \ \ \ s -\-LCons (s, s', tl, s'') tls\*t \" | \inf_step_table_Nil: "s -\\ \ \ s -\-LNil\*t \" definition \inf_step2\inf_step_table :: "'s \ 'tl llist \ ('s \ 's \ 'tl \ 's) llist" where "\inf_step2\inf_step_table s tls = unfold_llist (\(s, tls). lnull tls) (\(s, tls). let (s', s'') = SOME (s', s''). s -\\* s' \ s' -lhd tls\ s'' \ \ \move s' (lhd tls) s'' \ s'' -\-ltl tls\* \ in (s, s', lhd tls, s'')) (\(s, tls). let (s', s'') = SOME (s', s''). s -\\* s' \ s' -lhd tls\ s'' \ \ \move s' (lhd tls) s'' \ s'' -\-ltl tls\* \ in (s'', ltl tls)) (s, tls)" definition silent_move_from :: "'s \ 's \ 's \ bool" where "silent_move_from s0 s1 s2 \ silent_moves s0 s1 \ silent_move s1 s2" inductive \rtrancl3p :: "'s \ 'tl list \ 's \ bool" ("_ -\-_\* _" [50, 0, 50] 60) where \rtrancl3p_refl: "\rtrancl3p s [] s" | \rtrancl3p_step: "\tl. \ s -tl\ s'; \ \move s tl s'; \rtrancl3p s' tls s'' \ \ \rtrancl3p s (tl # tls) s''" | \rtrancl3p_\step: "\tl. \ s -tl\ s'; \move s tl s'; \rtrancl3p s' tls s'' \ \ \rtrancl3p s tls s''" coinductive \Runs :: "'s \ ('tl, 's option) tllist \ bool" ("_ \ _" [50, 50] 51) where Terminate: "\ s -\\* s'; \tl s''. \ s' -tl\ s'' \ \ s \ TNil \s'\" | Diverge: "s -\\ \ \ s \ TNil None" | Proceed: "\tl. \ s -\\* s'; s' -tl\ s''; \ \move s' tl s''; s'' \ tls \ \ s \ TCons tl tls" inductive_simps \Runs_simps: "s \ TNil (Some s')" "s \ TNil None" "s \ TCons tl' tls" coinductive \Runs_table :: "'s \ ('tl \ 's, 's option) tllist \ bool" where Terminate: "\ s -\\* s'; \tl s''. \ s' -tl\ s'' \ \ \Runs_table s (TNil \s'\)" | Diverge: "s -\\ \ \ \Runs_table s (TNil None)" | Proceed: "\tl. \ s -\\* s'; s' -tl\ s''; \ \move s' tl s''; \Runs_table s'' tls \ \ \Runs_table s (TCons (tl, s'') tls)" definition silent_move2 :: "'s \ 'tl \ 's \ bool" where "\tl. silent_move2 s tl s' \ s -tl\ s' \ \move s tl s'" abbreviation silent_moves2 :: "'s \ 'tl list \ 's \ bool" where "silent_moves2 \ rtrancl3p silent_move2" coinductive \Runs_table2 :: "'s \ ('tl list \ 's \ 'tl \ 's, ('tl list \ 's) + 'tl llist) tllist \ bool" where Terminate: "\ silent_moves2 s tls s'; \tl s''. \ s' -tl\ s'' \ \ \Runs_table2 s (TNil (Inl (tls, s')))" | Diverge: "trsys.inf_step silent_move2 s tls \ \Runs_table2 s (TNil (Inr tls))" | Proceed: "\tl. \ silent_moves2 s tls s'; s' -tl\ s''; \ \move s' tl s''; \Runs_table2 s'' tlsstlss \ \ \Runs_table2 s (TCons (tls, s', tl, s'') tlsstlss)" inductive_simps \Runs_table2_simps: "\Runs_table2 s (TNil tlss)" "\Runs_table2 s (TCons tlsstls tlsstlss)" lemma inf_step_table_all_\_into_\diverge: "\ s -stls\*t \; \(s, tl, s') \ lset stls. \move s tl s' \ \ s -\\ \" proof(coinduction arbitrary: s stls) case (\diverge s) thus ?case by cases (auto simp add: silent_move_iff, blast) qed lemma inf_step_table_lappend_llist_ofD: "s -lappend (llist_of stls) (LCons (x, tl', x') xs)\*t \ \ (s -map (fst \ snd) stls\* x) \ (x -LCons (x, tl', x') xs\*t \)" proof(induct stls arbitrary: s) case Nil thus ?case by(auto elim: inf_step_table.cases intro: inf_step_table.intros rtrancl3p_refl) next case (Cons st stls) note IH = \\s. s -lappend (llist_of stls) (LCons (x, tl', x') xs)\*t \ \ s -map (fst \ snd) stls\* x \ x -LCons (x, tl', x') xs\*t \\ from \s -lappend (llist_of (st # stls)) (LCons (x, tl', x') xs)\*t \\ show ?case proof cases case (inf_step_tableI s' stls' tl) hence [simp]: "st = (s, tl, s')" "stls' = lappend (llist_of stls) (LCons (x, tl', x') xs)" and "s -tl\ s'" "s' -lappend (llist_of stls) (LCons (x, tl', x') xs)\*t \" by simp_all from IH[OF \s' -lappend (llist_of stls) (LCons (x, tl', x') xs)\*t \\] have "s' -map (fst \ snd) stls\* x" "x -LCons (x, tl', x') xs\*t \" by auto with \s -tl\ s'\ show ?thesis by(auto simp add: o_def intro: rtrancl3p_step_converse) qed qed lemma inf_step_table_lappend_llist_of_\_into_\moves: assumes "lfinite stls" shows "\ s -lappend stls (LCons (x, tl' x') xs)\*t \; \(s, tl, s')\lset stls. \move s tl s' \ \ s -\\* x" using assms proof(induct arbitrary: s rule: lfinite.induct) case lfinite_LNil thus ?case by(auto elim: inf_step_table.cases) next case (lfinite_LConsI stls st) note IH = \\s. \s -lappend stls (LCons (x, tl' x') xs)\*t \; \(s, tl, s')\lset stls. \move s tl s' \ \ s -\\* x\ obtain s1 tl1 s1' where [simp]: "st = (s1, tl1, s1')" by(cases st) from \s -lappend (LCons st stls) (LCons (x, tl' x') xs)\*t \\ show ?case proof cases case (inf_step_tableI X' STLS TL) hence [simp]: "s1 = s" "TL = tl1" "X' = s1'" "STLS = lappend stls (LCons (x, tl' x') xs)" and "s -tl1\ s1'" and "s1' -lappend stls (LCons (x, tl' x') xs)\*t \" by simp_all from \\(s, tl, s')\lset (LCons st stls). \move s tl s'\ have "\move s tl1 s1'" by simp moreover from IH[OF \s1' -lappend stls (LCons (x, tl' x') xs)\*t \\] \\(s, tl, s')\lset (LCons st stls). \move s tl s'\ have "s1' -\\* x" by simp ultimately show ?thesis using \s -tl1\ s1'\ by(auto intro: converse_rtranclp_into_rtranclp) qed qed lemma inf_step_table_into_\inf_step: "s -stls\*t \ \ s -\-lmap (fst \ snd) (lfilter (\(s, tl, s'). \ \move s tl s') stls)\* \" proof(coinduction arbitrary: s stls) case (\inf_step s stls) let ?P = "\(s, tl, s'). \ \move s tl s'" show ?case proof(cases "lfilter ?P stls") case LNil with \inf_step have ?\inf_step_Nil by(auto intro: inf_step_table_all_\_into_\diverge simp add: lfilter_eq_LNil) thus ?thesis .. next case (LCons stls' xs) obtain x tl x' where "stls' = (x, tl, x')" by(cases stls') with LCons have stls: "lfilter ?P stls = LCons (x, tl, x') xs" by simp from lfilter_eq_LConsD[OF this] obtain stls1 stls2 where stls1: "stls = lappend stls1 (LCons (x, tl, x') stls2)" and "lfinite stls1" and \s: "\(s, tl, s')\lset stls1. \move s tl s'" and n\: "\ \move x tl x'" and xs: "xs = lfilter ?P stls2" by blast from \lfinite stls1\ \inf_step \s have "s -\\* x" unfolding stls1 by(rule inf_step_table_lappend_llist_of_\_into_\moves) moreover from \lfinite stls1\ have "llist_of (list_of stls1) = stls1" by(simp add: llist_of_list_of) with \inf_step stls1 have "s -lappend (llist_of (list_of stls1)) (LCons (x, tl, x') stls2)\*t \" by simp from inf_step_table_lappend_llist_ofD[OF this] have "x -LCons (x, tl, x') stls2\*t \" .. hence "x -tl\ x'" "x' -stls2\*t \" by(auto elim: inf_step_table.cases) ultimately have ?\inf_step_Cons using xs n\ by(auto simp add: stls o_def) thus ?thesis .. qed qed lemma inf_step_into_\inf_step: assumes "s -tls\* \" shows "\A. s -\-lnths tls A\* \" proof - from inf_step_imp_inf_step_table[OF assms] obtain stls where "s -stls\*t \" and tls: "tls = lmap (fst \ snd) stls" by blast from \s -stls\*t \\ have "s -\-lmap (fst \ snd) (lfilter (\(s, tl, s'). \ \move s tl s') stls)\* \" by(rule inf_step_table_into_\inf_step) hence "s -\-lnths tls {n. enat n < llength stls \ (\(s, tl, s'). \ \move s tl s') (lnth stls n)}\* \" unfolding lfilter_conv_lnths tls by simp thus ?thesis by blast qed lemma silent_moves_into_\rtrancl3p: "s -\\* s' \ s -\-[]\* s'" by(induct rule: converse_rtranclp_induct)(blast intro: \rtrancl3p.intros)+ lemma \rtrancl3p_into_silent_moves: "s -\-[]\* s' \ s -\\* s'" apply(induct s tls\"[] :: 'tl list" s' rule: \rtrancl3p.induct) apply(auto intro: converse_rtranclp_into_rtranclp) done lemma \rtrancl3p_Nil_eq_\moves: "s -\-[]\* s' \ s -\\* s'" by(blast intro: silent_moves_into_\rtrancl3p \rtrancl3p_into_silent_moves) lemma \rtrancl3p_trans [trans]: "\ s -\-tls\* s'; s' -\-tls'\* s'' \ \ s -\-tls @ tls'\* s''" apply(induct rule: \rtrancl3p.induct) apply(auto intro: \rtrancl3p.intros) done lemma \rtrancl3p_SingletonE: fixes tl assumes red: "s -\-[tl]\* s'''" obtains s' s'' where "s -\\* s'" "s' -tl\ s''" "\ \move s' tl s''" "s'' -\\* s'''" proof(atomize_elim) from red show "\s' s''. s -\\* s' \ s' -tl\ s'' \ \ \move s' tl s'' \ s'' -\\* s'''" proof(induct s tls\"[tl]" s''') case (\rtrancl3p_step s s' s'') from \s -tl\ s'\ \\ \move s tl s'\ \s' -\-[]\* s''\ show ?case by(auto simp add: \rtrancl3p_Nil_eq_\moves) next case (\rtrancl3p_\step s s' s'' tl') then obtain t' t'' where "s' -\\* t'" "t' -tl\ t''" "\ \move t' tl t''" "t'' -\\* s''" by auto moreover from \s -tl'\ s'\ \\move s tl' s'\ have "s -\\* s'" by blast ultimately show ?case by(auto intro: rtranclp_trans) qed qed lemma \rtrancl3p_snocI: "\tl. \ \rtrancl3p s tls s''; s'' -\\* s'''; s''' -tl\ s'; \ \move s''' tl s' \ \ \rtrancl3p s (tls @ [tl]) s'" apply(erule \rtrancl3p_trans) apply(fold \rtrancl3p_Nil_eq_\moves) apply(drule \rtrancl3p_trans) apply(erule (1) \rtrancl3p_step) apply(rule \rtrancl3p_refl) apply simp done lemma \diverge_rtranclp_silent_move: "\ silent_move^** s s'; s' -\\ \ \ \ s -\\ \" by(induct rule: converse_rtranclp_induct)(auto intro: \divergeI) lemma \diverge_trancl_coinduct [consumes 1, case_names \diverge]: assumes X: "X s" and step: "\s. X s \ \s'. silent_move^++ s s' \ (X s' \ s' -\\ \)" shows "s -\\ \" proof - from X have "\s'. silent_move^** s s' \ X s'" by blast thus ?thesis proof(coinduct) case (\diverge s) then obtain s' where "silent_move\<^sup>*\<^sup>* s s'" "X s'" by blast from step[OF \X s'\] obtain s''' where "silent_move^++ s' s'''" "X s''' \ s''' -\\ \" by blast from \silent_move\<^sup>*\<^sup>* s s'\ show ?case proof(cases rule: converse_rtranclpE[consumes 1, case_names refl step]) case refl moreover from tranclpD[OF \silent_move^++ s' s'''\] obtain s'' where "silent_move s' s''" "silent_move^** s'' s'''" by blast ultimately show ?thesis using \silent_move^** s'' s'''\ \X s''' \ s''' -\\ \\ by(auto intro: \diverge_rtranclp_silent_move) next case (step S) moreover from \silent_move\<^sup>*\<^sup>* S s'\ \silent_move^++ s' s'''\ have "silent_move^** S s'''" by(rule rtranclp_trans[OF _ tranclp_into_rtranclp]) ultimately show ?thesis using \X s''' \ s''' -\\ \\ by(auto intro: \diverge_rtranclp_silent_move) qed qed qed lemma \diverge_trancl_measure_coinduct [consumes 2, case_names \diverge]: assumes major: "X s t" "wfP \" and step: "\s t. X s t \ \s' t'. (\ t' t \ s' = s \ silent_move^++ s s') \ (X s' t' \ s' -\\ \)" shows "s -\\ \" proof - { fix s t assume "X s t" with \wfP \\ have "\s' t'. silent_move^++ s s' \ (X s' t' \ s' -\\ \)" proof(induct arbitrary: s rule: wfP_induct[consumes 1]) case (1 t) hence IH: "\s' t'. \ \ t' t; X s' t' \ \ \s'' t''. silent_move^++ s' s'' \ (X s'' t'' \ s'' -\\ \)" by blast from step[OF \X s t\] obtain s' t' where "\ t' t \ s' = s \ silent_move\<^sup>+\<^sup>+ s s'" "X s' t' \ s' -\\ \" by blast from \\ t' t \ s' = s \ silent_move\<^sup>+\<^sup>+ s s'\ show ?case proof assume "\ t' t \ s' = s" hence "\ t' t" and [simp]: "s' = s" by simp_all from \X s' t' \ s' -\\ \\ show ?thesis proof assume "X s' t'" from IH[OF \\ t' t\ this] show ?thesis by simp next assume "s' -\\ \" thus ?thesis by cases(auto simp add: silent_move_iff) qed next assume "silent_move\<^sup>+\<^sup>+ s s'" thus ?thesis using \X s' t' \ s' -\\ \\ by blast qed qed } note X = this from \X s t\ have "\t. X s t" .. thus ?thesis proof(coinduct rule: \diverge_trancl_coinduct) case (\diverge s) then obtain t where "X s t" .. from X[OF this] show ?case by blast qed qed lemma \inf_step2\inf_step_table_LNil [simp]: "\inf_step2\inf_step_table s LNil = LNil" by(simp add: \inf_step2\inf_step_table_def) lemma \inf_step2\inf_step_table_LCons [simp]: fixes s tl ss tls defines "ss \ SOME (s', s''). s -\\* s' \ s' -tl\ s'' \ \ \move s' tl s'' \ s'' -\-tls\* \" shows "\inf_step2\inf_step_table s (LCons tl tls) = LCons (s, fst ss, tl, snd ss) (\inf_step2\inf_step_table (snd ss) tls)" by(simp add: ss_def \inf_step2\inf_step_table_def split_beta) lemma lnull_\inf_step2\inf_step_table [simp]: "lnull (\inf_step2\inf_step_table s tls) \ lnull tls" by(simp add: \inf_step2\inf_step_table_def) lemma lhd_\inf_step2\inf_step_table [simp]: "\ lnull tls \ lhd (\inf_step2\inf_step_table s tls) = (let (s', s'') = SOME (s', s''). s -\\* s' \ s' -lhd tls\ s'' \ \ \move s' (lhd tls) s'' \ s'' -\-ltl tls\* \ in (s, s', lhd tls, s''))" unfolding \inf_step2\inf_step_table_def Let_def by simp lemma ltl_\inf_step2\inf_step_table [simp]: "\ lnull tls \ ltl (\inf_step2\inf_step_table s tls) = (let (s', s'') = SOME (s', s''). s -\\* s' \ s' -lhd tls\ s'' \ \ \move s' (lhd tls) s'' \ s'' -\-ltl tls\* \ in \inf_step2\inf_step_table s'' (ltl tls))" unfolding \inf_step2\inf_step_table_def Let_def by(simp add: split_beta) lemma lmap_\inf_step2\inf_step_table: "lmap (fst \ snd \ snd) (\inf_step2\inf_step_table s tls) = tls" by(coinduction arbitrary: s tls)(auto simp add: split_beta) lemma \inf_step_into_\inf_step_table: "s -\-tls\* \ \ s -\-\inf_step2\inf_step_table s tls\*t \" proof(coinduction arbitrary: s tls) case (\inf_step_table s tls) thus ?case proof(cases) case (\inf_step_Cons s' s'' tls' tl) let ?ss = "SOME (s', s''). s -\\* s' \ s' -tl\ s'' \ \ \move s' tl s'' \ s'' -\-tls'\* \" from \inf_step_Cons have tls: "tls = LCons tl tls'" and "s -\\* s'" "s' -tl\ s''" "\ \move s' tl s''" "s'' -\-tls'\* \" by simp_all hence "(\(s', s''). s -\\* s' \ s' -tl\ s'' \ \ \move s' tl s'' \ s'' -\-tls'\* \) (s', s'')" by simp hence "(\(s', s''). s -\\* s' \ s' -tl\ s'' \ \ \move s' tl s'' \ s'' -\-tls'\* \) ?ss" by(rule someI) with tls have ?\inf_step_table_Cons by auto thus ?thesis .. next case \inf_step_Nil then have ?\inf_step_table_Nil by simp thus ?thesis .. qed qed lemma \inf_step_imp_\inf_step_table: assumes "s -\-tls\* \" shows "\sstls. s -\-sstls\*t \ \ tls = lmap (fst \ snd \ snd) sstls" using \inf_step_into_\inf_step_table[OF assms] by(auto simp only: lmap_\inf_step2\inf_step_table) lemma \inf_step_table_into_\inf_step: "s -\-sstls\*t \ \ s -\-lmap (fst \ snd \ snd) sstls\* \" proof(coinduction arbitrary: s sstls) case (\inf_step s tls) thus ?case by cases(auto simp add: o_def) qed lemma silent_move_fromI [intro]: "\ silent_moves s0 s1; silent_move s1 s2 \ \ silent_move_from s0 s1 s2" by(simp add: silent_move_from_def) lemma silent_move_fromE [elim]: assumes "silent_move_from s0 s1 s2" obtains "silent_moves s0 s1" "silent_move s1 s2" using assms by(auto simp add: silent_move_from_def) lemma rtranclp_silent_move_from_imp_silent_moves: assumes s'x: "silent_move\<^sup>*\<^sup>* s' x" shows "(silent_move_from s')^** x z \ silent_moves s' z" by(induct rule: rtranclp_induct)(auto intro: s'x) lemma \diverge_not_wfP_silent_move_from: assumes "s -\\ \" shows "\ wfP (flip (silent_move_from s))" proof assume "wfP (flip (silent_move_from s))" moreover define Q where "Q = {s'. silent_moves s s' \ s' -\\ \}" hence "s \ Q" using \s -\\ \\ by(auto) ultimately have "\z\Q. \y. silent_move_from s z y \ y \ Q" unfolding wfP_eq_minimal flip_simps by blast then obtain z where "z \ Q" and min: "\y. silent_move_from s z y \ y \ Q" by blast from \z \ Q\ have "silent_moves s z" "z -\\ \" unfolding Q_def by auto from \z -\\ \\ obtain y where "silent_move z y" "y -\\ \" by cases auto from \silent_moves s z\ \silent_move z y\ have "silent_move_from s z y" .. hence "y \ Q" by(rule min) moreover from \silent_moves s z\ \silent_move z y\ \y -\\ \\ have "y \ Q" unfolding Q_def by auto ultimately show False by contradiction qed lemma wfP_silent_move_from_unroll: assumes wfPs': "\s'. s -\\ s' \ wfP (flip (silent_move_from s'))" shows "wfP (flip (silent_move_from s))" unfolding wfP_eq_minimal flip_conv proof(intro allI impI) fix Q and x :: 's assume "x \ Q" show "\z\Q. \y. silent_move_from s z y \ y \ Q" proof(cases "\s'. s -\\ s' \ (\x'. silent_moves s' x' \ x' \ Q)") case False hence "\y. silent_move_from s x y \ \ y \ Q" by(cases "x=s")(auto, blast elim: converse_rtranclpE intro: rtranclp.rtrancl_into_rtrancl) with \x \ Q\ show ?thesis by blast next case True then obtain s' x' where "s -\\ s'" and "silent_moves s' x'" and "x' \ Q" by auto from \s -\\ s'\ have "wfP (flip (silent_move_from s'))" by(rule wfPs') from this \x' \ Q\ obtain z where "z \ Q" and min: "\y. silent_move_from s' z y \ \ y \ Q" and "(silent_move_from s')^** x' z" by (rule wfP_minimalE) (unfold flip_simps, blast) { fix y assume "silent_move_from s z y" with \(silent_move_from s')^** x' z\ \silent_move^** s' x'\ have "silent_move_from s' z y" by(blast intro: rtranclp_silent_move_from_imp_silent_moves) hence "\ y \ Q" by(rule min) } with \z \ Q\ show ?thesis by(auto simp add: intro!: bexI) qed qed lemma not_wfP_silent_move_from_\diverge: assumes "\ wfP (flip (silent_move_from s))" shows "s -\\ \" using assms proof(coinduct) case (\diverge s) { assume wfPs': "\s'. s -\\ s' \ wfP (flip (silent_move_from s'))" hence "wfP (flip (silent_move_from s))" by(rule wfP_silent_move_from_unroll) } with \diverge have "\s'. s -\\ s' \ \ wfP (flip (silent_move_from s'))" by auto thus ?case by blast qed lemma \diverge_neq_wfP_silent_move_from: "s -\\ \ \ wfP (flip (silent_move_from s))" by(auto intro: not_wfP_silent_move_from_\diverge dest: \diverge_not_wfP_silent_move_from) lemma not_\diverge_to_no_\move: assumes "\ s -\\ \" shows "\s'. s -\\* s' \ (\s''. \ s' -\\ s'')" proof - define S where "S = s" from \\ \diverge s\ have "wfP (flip (silent_move_from S))" unfolding S_def using \diverge_neq_wfP_silent_move_from[of s] by simp moreover have "silent_moves S s" unfolding S_def .. ultimately show ?thesis proof(induct rule: wfP_induct') case (wfP s) note IH = \\y. \flip (silent_move_from S) y s; S -\\* y \ \ \s'. y -\\* s' \ (\s''. \ s' -\\ s'')\ show ?case proof(cases "\s'. silent_move s s'") case False thus ?thesis by auto next case True then obtain s' where "s -\\ s'" .. with \S -\\* s\ have "flip (silent_move_from S) s' s" unfolding flip_conv by(rule silent_move_fromI) moreover from \S -\\* s\ \s -\\ s'\ have "S -\\* s'" .. ultimately have "\s''. s' -\\* s'' \ (\s'''. \ s'' -\\ s''')" by(rule IH) then obtain s'' where "s' -\\* s''" "\s'''. \ s'' -\\ s'''" by blast from \s -\\ s'\ \s' -\\* s''\ have "s -\\* s''" by(rule converse_rtranclp_into_rtranclp) with \\s'''. \ s'' -\\ s'''\ show ?thesis by blast qed qed qed lemma \diverge_conv_\Runs: "s -\\ \ \ s \ TNil None" by(auto intro: \Runs.Diverge elim: \Runs.cases) lemma \inf_step_into_\Runs: "s -\-tls\* \ \ s \ tllist_of_llist None tls" proof(coinduction arbitrary: s tls) case (\Runs s tls') thus ?case by cases(auto simp add: \diverge_conv_\Runs) qed lemma \_into_\Runs: "\ s -\\ s'; s' \ tls \ \ s \ tls" by(blast elim: \Runs.cases intro: \Runs.intros \diverge.intros converse_rtranclp_into_rtranclp) lemma \rtrancl3p_into_\Runs: assumes "s -\-tls\* s'" and "s' \ tls'" shows "s \ lappendt (llist_of tls) tls'" using assms by induct(auto intro: \Runs.Proceed \_into_\Runs) lemma \Runs_table_into_\Runs: "\Runs_table s stlsss \ s \ tmap fst id stlsss" proof(coinduction arbitrary: s stlsss) case (\Runs s tls) thus ?case by cases(auto simp add: o_def id_def) qed definition \Runs2\Runs_table :: "'s \ ('tl, 's option) tllist \ ('tl \ 's, 's option) tllist" where "\Runs2\Runs_table s tls = unfold_tllist (\(s, tls). is_TNil tls) (\(s, tls). terminal tls) (\(s, tls). (thd tls, SOME s''. \s'. s -\\* s' \ s' -thd tls\ s'' \ \ \move s' (thd tls) s'' \ s'' \ ttl tls)) (\(s, tls). (SOME s''. \s'. s -\\* s' \ s' -thd tls\ s'' \ \ \move s' (thd tls) s'' \ s'' \ ttl tls, ttl tls)) (s, tls)" lemma is_TNil_\Runs2\Runs_table [simp]: "is_TNil (\Runs2\Runs_table s tls) \ is_TNil tls" thm unfold_tllist.disc by(simp add: \Runs2\Runs_table_def) lemma thd_\Runs2\Runs_table [simp]: "\ is_TNil tls \ thd (\Runs2\Runs_table s tls) = (thd tls, SOME s''. \s'. s -\\* s' \ s' -thd tls\ s'' \ \ \move s' (thd tls) s'' \ s'' \ ttl tls)" by(simp add: \Runs2\Runs_table_def) lemma ttl_\Runs2\Runs_table [simp]: "\ is_TNil tls \ ttl (\Runs2\Runs_table s tls) = \Runs2\Runs_table (SOME s''. \s'. s -\\* s' \ s' -thd tls\ s'' \ \ \move s' (thd tls) s'' \ s'' \ ttl tls) (ttl tls)" by(simp add: \Runs2\Runs_table_def) lemma terminal_\Runs2\Runs_table [simp]: "is_TNil tls \ terminal (\Runs2\Runs_table s tls) = terminal tls" by(simp add: \Runs2\Runs_table_def) lemma \Runs2\Runs_table_simps [simp, nitpick_simp]: "\Runs2\Runs_table s (TNil so) = TNil so" "\tl. \Runs2\Runs_table s (TCons tl tls) = (let s'' = SOME s''. \s'. s -\\* s' \ s' -tl\ s'' \ \ \move s' tl s'' \ s'' \ tls in TCons (tl, s'') (\Runs2\Runs_table s'' tls))" apply(simp add: \Runs2\Runs_table_def) apply(rule tllist.expand) apply(simp_all) done lemma \Runs2\Runs_table_inverse: "tmap fst id (\Runs2\Runs_table s tls) = tls" by(coinduction arbitrary: s tls) auto lemma \Runs_into_\Runs_table: assumes "s \ tls" shows "\stlsss. tls = tmap fst id stlsss \ \Runs_table s stlsss" proof(intro exI conjI) from assms show "\Runs_table s (\Runs2\Runs_table s tls)" proof(coinduction arbitrary: s tls) case (\Runs_table s tls) thus ?case proof cases case (Terminate s') hence ?Terminate by simp thus ?thesis .. next case Diverge hence ?Diverge by simp thus ?thesis by simp next case (Proceed s' s'' tls' tl) let ?P = "\s''. \s'. s -\\* s' \ s' -tl\ s'' \ \ \move s' tl s'' \ s'' \ tls'" from Proceed have "?P s''" by auto hence "?P (Eps ?P)" by(rule someI) hence ?Proceed using \tls = TCons tl tls'\ by(auto simp add: split_beta) thus ?thesis by simp qed qed qed(simp add: \Runs2\Runs_table_inverse) lemma \Runs_lappendtE: assumes "\ \ lappendt tls tls'" and "lfinite tls" obtains \' where "\ -\-list_of tls\* \'" and "\' \ tls'" proof(atomize_elim) from \lfinite tls\ \\ \ lappendt tls tls'\ show "\\'. \ -\-list_of tls\* \' \ \' \ tls'" proof(induct arbitrary: \) case lfinite_LNil thus ?case by(auto intro: \rtrancl3p_refl) next case (lfinite_LConsI tls tl) from \\ \ lappendt (LCons tl tls) tls'\ show ?case unfolding lappendt_LCons proof(cases) case (Proceed \' \'') from \\'' \ lappendt tls tls' \ \\'''. \'' -\-list_of tls\* \''' \ \''' \ tls'\ \\'' \ lappendt tls tls'\ obtain \''' where "\'' -\-list_of tls\* \'''" "\''' \ tls'" by blast from \\' -tl\ \''\ \\ \move \' tl \''\ \\'' -\-list_of tls\* \'''\ have "\' -\-tl # list_of tls\* \'''" by(rule \rtrancl3p_step) with \\ -\\* \'\ have "\ -\-[] @ (tl # list_of tls)\* \'''" unfolding \rtrancl3p_Nil_eq_\moves[symmetric] by(rule \rtrancl3p_trans) with \lfinite tls\ have "\ -\-list_of (LCons tl tls)\* \'''" by(simp add: list_of_LCons) with \\''' \ tls'\ show ?thesis by blast qed qed qed lemma \Runs_total: "\tls. \ \ tls" proof let ?\halt = "\\ \'. \ -\\* \' \ (\tl \''. \ \' -tl\ \'')" let ?\diverge = "\\. \ -\\ \" let ?proceed = "\\ (tl, \''). \\'. \ -\\* \' \ \' -tl\ \'' \ \ \move \' tl \''" define tls where "tls = unfold_tllist (\\. (\\'. ?\halt \ \') \ ?\diverge \) (\\. if \\'. ?\halt \ \' then Some (SOME \'. ?\halt \ \') else None) (\\. fst (SOME tl\'. ?proceed \ tl\')) (\\. snd (SOME tl\'. ?proceed \ tl\')) \" then show "\ \ tls" proof(coinduct \ tls rule: \Runs.coinduct) case (\Runs \ tls) show ?case proof(cases "\\'. ?\halt \ \'") case True hence "?\halt \ (SOME \'. ?\halt \ \')" by(rule someI_ex) hence ?Terminate using True unfolding \Runs by simp thus ?thesis .. next case False note \halt = this show ?thesis proof(cases "?\diverge \") case True hence ?Diverge using False unfolding \Runs by simp thus ?thesis by simp next case False from not_\diverge_to_no_\move[OF this] obtain \' where \_\': "\ -\\* \'" and no_\: "\\''. \ \' -\\ \''" by blast from \_\' \halt obtain tl \'' where "\' -tl\ \''" by auto moreover with no_\[of \''] have "\ \move \' tl \''" by auto ultimately have "?proceed \ (tl, \'')" using \_\' by auto hence "?proceed \ (SOME tl\. ?proceed \ tl\)" by(rule someI) hence ?Proceed using False \halt unfolding \Runs by(subst unfold_tllist.code) fastforce thus ?thesis by simp qed qed qed qed lemma silent_move2_into_silent_move: fixes tl assumes "silent_move2 s tl s'" shows "s -\\ s'" using assms by(auto simp add: silent_move2_def) lemma silent_move_into_silent_move2: assumes "s -\\ s'" shows "\tl. silent_move2 s tl s'" using assms by(auto simp add: silent_move2_def) lemma silent_moves2_into_silent_moves: assumes "silent_moves2 s tls s'" shows "s -\\* s'" using assms by(induct)(blast intro: silent_move2_into_silent_move rtranclp.rtrancl_into_rtrancl)+ lemma silent_moves_into_silent_moves2: assumes "s -\\* s'" shows "\tls. silent_moves2 s tls s'" using assms by(induct)(blast dest: silent_move_into_silent_move2 intro: rtrancl3p_step)+ lemma inf_step_silent_move2_into_\diverge: "trsys.inf_step silent_move2 s tls \ s -\\ \" proof(coinduction arbitrary: s tls) case (\diverge s) thus ?case by(cases rule: trsys.inf_step.cases[consumes 1])(auto intro: silent_move2_into_silent_move) qed lemma \diverge_into_inf_step_silent_move2: assumes "s -\\ \" obtains tls where "trsys.inf_step silent_move2 s tls" proof - define tls where "tls = unfold_llist (\_. False) (\s. fst (SOME (tl, s'). silent_move2 s tl s' \ s' -\\ \)) (\s. snd (SOME (tl, s'). silent_move2 s tl s' \ s' -\\ \)) s" (is "_ = ?tls s") with assms have "s -\\ \ \ tls = ?tls s" by simp hence "trsys.inf_step silent_move2 s tls" proof(coinduct rule: trsys.inf_step.coinduct[consumes 1, case_names inf_step, case_conclusion inf_step step]) case (inf_step s tls) let ?P = "\(tl, s'). silent_move2 s tl s' \ s' -\\ \" from inf_step obtain "s -\\ \" and tls: "tls = ?tls s" .. from \s -\\ \\ obtain s' where "s -\\ s'" "s' -\\ \" by cases from \s -\\ s'\ obtain tl where "silent_move2 s tl s'" by(blast dest: silent_move_into_silent_move2) with \s' -\\ \\ have "?P (tl, s')" by simp hence "?P (Eps ?P)" by(rule someI) thus ?case using tls by(subst (asm) unfold_llist.code)(auto) qed thus thesis by(rule that) qed lemma \Runs_into_\rtrancl3p: assumes runs: "s \ tlss" and fin: "tfinite tlss" and terminal: "terminal tlss = Some s'" shows "\rtrancl3p s (list_of (llist_of_tllist tlss)) s'" using fin runs terminal proof(induct arbitrary: s rule: tfinite_induct) case TNil thus ?case by cases(auto intro: silent_moves_into_\rtrancl3p) next case (TCons tl tlss) from \s \ TCons tl tlss\ obtain s'' s''' where step: "s -\\* s''" and step2: "s'' -tl\ s'''" "\ \move s'' tl s'''" and "s''' \ tlss" by cases from \terminal (TCons tl tlss) = \s'\\ have "terminal tlss = \s'\" by simp with \s''' \ tlss\ have "s''' -\-list_of (llist_of_tllist tlss)\* s'" by(rule TCons) with step2 have "s'' -\-tl # list_of (llist_of_tllist tlss)\* s'" by(rule \rtrancl3p_step) with step have "s -\-[] @ tl # list_of (llist_of_tllist tlss)\* s'" by(rule \rtrancl3p_trans[OF silent_moves_into_\rtrancl3p]) thus ?case using \tfinite tlss\ by simp qed lemma \Runs_terminal_stuck: assumes Runs: "s \ tlss" and fin: "tfinite tlss" and terminal: "terminal tlss = Some s'" and proceed: "s' -tls\ s''" shows False using fin Runs terminal proof(induct arbitrary: s rule: tfinite_induct) case TNil thus ?case using proceed by cases auto next case TCons thus ?case by(fastforce elim: \Runs.cases) qed lemma Runs_table_silent_diverge: "\ Runs_table s stlss; \(s, tl, s') \ lset stlss. \move s tl s'; \ lfinite stlss \ \ s -\\ \" proof(coinduction arbitrary: s stlss) case (\diverge s) thus ?case by cases(auto 5 2) qed lemma Runs_table_silent_rtrancl: assumes "lfinite stlss" and "Runs_table s stlss" and "\(s, tl, s') \ lset stlss. \move s tl s'" shows "s -\\* llast (LCons s (lmap (\(s, tl, s'). s') stlss))" (is ?thesis1) and "llast (LCons s (lmap (\(s, tl, s'). s') stlss)) -tl'\ s'' \ False" (is "PROP ?thesis2") proof - from assms have "?thesis1 \ (llast (LCons s (lmap (\(s, tl, s'). s') stlss)) -tl'\ s'' \ False)" proof(induct arbitrary: s) case lfinite_LNil thus ?case by(auto elim: Runs_table.cases) next case (lfinite_LConsI stlss stls) from \Runs_table s (LCons stls stlss)\ obtain tl s' where [simp]: "stls = (s, tl, s')" and "s -tl\ s'" and Run': "Runs_table s' stlss" by cases from \\(s, tl, s')\lset (LCons stls stlss). \move s tl s'\ have "\move s tl s'" and silent': "\(s, tl, s')\lset stlss. \move s tl s'" by simp_all from \s -tl\ s'\ \\move s tl s'\ have "s -\\ s'" by auto moreover from Run' silent' have "s' -\\* llast (LCons s' (lmap (\(s, tl, s'). s') stlss)) \ (llast (LCons s' (lmap (\(s, tl, s'). s') stlss)) -tl'\ s'' \ False)" by(rule lfinite_LConsI) ultimately show ?case by(auto) qed thus ?thesis1 "PROP ?thesis2" by blast+ qed lemma Runs_table_silent_lappendD: fixes s stlss defines "s' \ llast (LCons s (lmap (\(s, tl, s'). s') stlss))" assumes Runs: "Runs_table s (lappend stlss stlss')" and fin: "lfinite stlss" and silent: "\(s, tl, s') \ lset stlss. \move s tl s'" shows "s -\\* s'" (is ?thesis1) and "Runs_table s' stlss'" (is ?thesis2) and "stlss' \ LNil \ s' = fst (lhd stlss')" (is "PROP ?thesis3") proof - from fin Runs silent have "?thesis1 \ ?thesis2 \ (stlss' \ LNil \ s' = fst (lhd stlss'))" unfolding s'_def proof(induct arbitrary: s) case lfinite_LNil thus ?case by(auto simp add: neq_LNil_conv Runs_table_simps) next case lfinite_LConsI thus ?case by(clarsimp simp add: neq_LNil_conv Runs_table_simps)(blast intro: converse_rtranclp_into_rtranclp) qed thus ?thesis1 ?thesis2 "PROP ?thesis3" by simp_all qed lemma Runs_table_into_\Runs: fixes s stlss defines "tls \ tmap (\(s, tl, s'). tl) id (tfilter None (\(s, tl, s'). \ \move s tl s') (tllist_of_llist (Some (llast (LCons s (lmap (\(s, tl, s'). s') stlss)))) stlss))" (is "_ \ ?conv s stlss") assumes "Runs_table s stlss" shows "\Runs s tls" using assms proof(coinduction arbitrary: s tls stlss) case (\Runs s tls stlss) note tls = \tls = ?conv s stlss\ and Run = \Runs_table s stlss\ show ?case proof(cases tls) case [simp]: (TNil so) from tls have silent: "\(s, tl, s') \ lset stlss. \move s tl s'" by(auto simp add: TNil_eq_tmap_conv tfilter_empty_conv) show ?thesis proof(cases "lfinite stlss") case False with Run silent have "s -\\ \" by(rule Runs_table_silent_diverge) hence ?Diverge using False tls by(simp add: TNil_eq_tmap_conv tfilter_empty_conv) thus ?thesis by simp next case True with Runs_table_silent_rtrancl[OF this Run silent] have ?Terminate using tls by(auto simp add: TNil_eq_tmap_conv tfilter_empty_conv terminal_tllist_of_llist split_def) thus ?thesis by simp qed next case [simp]: (TCons tl tls') from tls obtain s' s'' stlss' where tl': "tfilter None (\(s, tl, s'). \ \move s tl s') (tllist_of_llist \llast (LCons s (lmap (\(s, tl, s'). s') stlss))\ stlss) = TCons (s', tl, s'') stlss'" and tls': "tls' = tmap (\(s, tl, s'). tl) id stlss'" by(simp add: TCons_eq_tmap_conv split_def id_def split_paired_Ex) blast from tfilter_eq_TConsD[OF tl'] obtain stls\ rest where stlss_eq: "tllist_of_llist \llast (LCons s (lmap (\(s, tl, s'). s') stlss))\ stlss = lappendt stls\ (TCons (s', tl, s'') rest)" and fin: "lfinite stls\" and silent: "\(s, tl, s')\lset stls\. \move s tl s'" and "\ \move s' tl s''" and stlss': "stlss' = tfilter None (\(s, tl, s'). \ \move s tl s') rest" by(auto simp add: split_def) from stlss_eq fin obtain rest' where stlss: "stlss = lappend stls\ rest'" and rest': "tllist_of_llist \llast (LCons s (lmap (\(s, tl, s'). s') stlss))\ rest' = TCons (s', tl, s'') rest" unfolding tllist_of_llist_eq_lappendt_conv by auto hence "rest' \ LNil" by clarsimp from Run[unfolded stlss] fin silent have "s -\\* llast (LCons s (lmap (\(s, tl, s'). s') stls\))" and "Runs_table (llast (LCons s (lmap (\(s, tl, s'). s') stls\))) rest'" and "llast (LCons s (lmap (\(s, tl, s'). s') stls\)) = fst (lhd rest')" by(rule Runs_table_silent_lappendD)+(simp add: \rest' \ LNil\) moreover with rest' \rest' \ LNil\ stlss fin obtain rest'' where rest': "rest' = LCons (s', tl, s'') rest''" and rest: "rest = tllist_of_llist \llast (LCons s'' (lmap (\(s, tl, s'). s') rest''))\ rest''" by(clarsimp simp add: neq_LNil_conv llast_LCons lmap_lappend_distrib) ultimately have "s -\\* s'" "s' -tl\ s''" "Runs_table s'' rest''" by(simp_all add: Runs_table_simps) hence ?Proceed using \\ \move s' tl s''\ tls' stlss' rest by(auto simp add: id_def) thus ?thesis by simp qed qed lemma \Runs_table2_into_\Runs: "\Runs_table2 s tlsstlss \ s \ tmap (\(tls, s', tl, s''). tl) (\x. case x of Inl (tls, s') \ Some s' | Inr _ \ None) tlsstlss" proof(coinduction arbitrary: s tlsstlss) case (\Runs s tlsstlss) thus ?case by cases(auto intro: silent_moves2_into_silent_moves inf_step_silent_move2_into_\diverge) qed lemma \Runs_into_\Runs_table2: assumes "s \ tls" obtains tlsstlss where "\Runs_table2 s tlsstlss" and "tls = tmap (\(tls, s', tl, s''). tl) (\x. case x of Inl (tls, s') \ Some s' | Inr _ \ None) tlsstlss" proof - let ?terminal = "\s tls. case terminal tls of None \ Inr (SOME tls'. trsys.inf_step silent_move2 s tls') | Some s' \ let tls' = SOME tls'. silent_moves2 s tls' s' in Inl (tls', s')" let ?P = "\s tls (tls'', s', s''). silent_moves2 s tls'' s' \ s' -thd tls\ s'' \ \ \move s' (thd tls) s'' \ s'' \ ttl tls" define tlsstlss where "tlsstlss s tls = unfold_tllist (\(s, tls). is_TNil tls) (\(s, tls). ?terminal s tls) (\(s, tls). let (tls'', s', s'') = Eps (?P s tls) in (tls'', s', thd tls, s'')) (\(s, tls). let (tls'', s', s'') = Eps (?P s tls) in (s'', ttl tls)) (s, tls)" for s tls have [simp]: "\s tls. is_TNil (tlsstlss s tls) \ is_TNil tls" "\s tls. is_TNil tls \ terminal (tlsstlss s tls) = ?terminal s tls" "\s tls. \ is_TNil tls \ thd (tlsstlss s tls) = (let (tls'', s', s'') = Eps (?P s tls) in (tls'', s', thd tls, s''))" "\s tls. \ is_TNil tls \ ttl (tlsstlss s tls) = (let (tls'', s', s'') = Eps (?P s tls) in tlsstlss s'' (ttl tls))" by(simp_all add: tlsstlss_def split_beta) have [simp]: "\s. tlsstlss s (TNil None) = TNil (Inr (SOME tls'. trsys.inf_step silent_move2 s tls'))" "\s s'. tlsstlss s (TNil (Some s')) = TNil (Inl (SOME tls'. silent_moves2 s tls' s', s'))" unfolding tlsstlss_def by simp_all let ?conv = "tmap (\(tls, s', tl, s''). tl) (\x. case x of Inl (tls, s') \ Some s' | Inr _ \ None)" from assms have "\Runs_table2 s (tlsstlss s tls)" proof(coinduction arbitrary: s tls) case (\Runs_table2 s tls) thus ?case proof(cases) case (Terminate s') let ?P = "\tls'. silent_moves2 s tls' s'" from \s -\\* s'\ obtain tls' where "?P tls'" by(blast dest: silent_moves_into_silent_moves2) hence "?P (Eps ?P)" by(rule someI) with Terminate have ?Terminate by auto thus ?thesis by simp next case Diverge let ?P = "\tls'. trsys.inf_step silent_move2 s tls'" from \s -\\ \\ obtain tls' where "?P tls'" by(rule \diverge_into_inf_step_silent_move2) hence "?P (Eps ?P)" by(rule someI) hence ?Diverge using \tls = TNil None\ by simp thus ?thesis by simp next case (Proceed s' s'' tls' tl) from \s -\\* s'\ obtain tls'' where "silent_moves2 s tls'' s'" by(blast dest: silent_moves_into_silent_moves2) with Proceed have "?P s tls (tls'', s', s'')" by simp hence "?P s tls (Eps (?P s tls))" by(rule someI) hence ?Proceed using Proceed unfolding tlsstlss_def by(subst unfold_tllist.code)(auto simp add: split_def) thus ?thesis by simp qed qed moreover from assms have "tls = ?conv (tlsstlss s tls)" proof(coinduction arbitrary: s tls) case (Eq_tllist s tls) thus ?case proof(cases) case (Proceed s' s'' tls' tl) from \s -\\* s'\ obtain tls'' where "silent_moves2 s tls'' s'" by(blast dest: silent_moves_into_silent_moves2) with Proceed have "?P s tls (tls'', s', s'')" by simp hence "?P s tls (Eps (?P s tls))" by(rule someI) thus ?thesis using \tls = TCons tl tls'\ by auto qed auto qed ultimately show thesis by(rule that) qed lemma \Runs_table2_into_Runs: assumes "\Runs_table2 s tlsstlss" shows "Runs s (lconcat (lappend (lmap (\(tls, s, tl, s'). llist_of (tls @ [tl])) (llist_of_tllist tlsstlss)) (LCons (case terminal tlsstlss of Inl (tls, s') \ llist_of tls | Inr tls \ tls) LNil)))" (is "Runs _ (?conv tlsstlss)") using assms proof(coinduction arbitrary: s tlsstlss) case (Runs s tlsstlss) thus ?case proof(cases) case (Terminate tls' s') from \silent_moves2 s tls' s'\ show ?thesis proof(cases rule: rtrancl3p_converseE) case refl hence ?Stuck using Terminate by simp thus ?thesis .. next case (step tls'' tl s'') from \silent_moves2 s'' tls'' s'\ \\tl s''. \ s' -tl\ s''\ have "\Runs_table2 s'' (TNil (Inl (tls'', s')))" .. with \tls' = tl # tls''\ \silent_move2 s tl s''\ \tlsstlss = TNil (Inl (tls', s'))\ have ?Step by(auto simp add: silent_move2_def intro!: exI) thus ?thesis .. qed next case (Diverge tls') from \trsys.inf_step silent_move2 s tls'\ obtain tl tls'' s' where "silent_move2 s tl s'" and "tls' = LCons tl tls''" "trsys.inf_step silent_move2 s' tls''" by(cases rule: trsys.inf_step.cases[consumes 1]) auto from \trsys.inf_step silent_move2 s' tls''\ have "\Runs_table2 s' (TNil (Inr tls''))" .. hence ?Step using \tlsstlss = TNil (Inr tls')\ \tls' = LCons tl tls''\ \silent_move2 s tl s'\ by(auto simp add: silent_move2_def intro!: exI) thus ?thesis .. next case (Proceed tls' s' s'' tlsstlss' tl) from \silent_moves2 s tls' s'\ have ?Step proof(cases rule: rtrancl3p_converseE) case refl with Proceed show ?thesis by auto next case (step tls'' tl' s''') from \silent_moves2 s''' tls'' s'\ \s' -tl\ s''\ \\ \move s' tl s''\ \\Runs_table2 s'' tlsstlss'\ have "\Runs_table2 s''' (TCons (tls'', s', tl, s'') tlsstlss')" .. with \tls' = tl' # tls''\ \silent_move2 s tl' s'''\ \tlsstlss = TCons (tls', s', tl, s'') tlsstlss'\ show ?thesis by(auto simp add: silent_move2_def intro!: exI) qed thus ?thesis .. qed qed lemma \Runs_table2_silentsD: fixes tl assumes Runs: "\Runs_table2 s tlsstlss" and tset: "(tls, s', tl', s'') \ tset tlsstlss" and set: "tl \ set tls" shows "\s''' s''''. silent_move2 s''' tl s''''" using tset Runs proof(induct arbitrary: s rule: tset_induct) case (find tlsstlss') from \\Runs_table2 s (TCons (tls, s', tl', s'') tlsstlss')\ have "silent_moves2 s tls s'" by cases thus ?case using set by induct auto next case step thus ?case by(auto simp add: \Runs_table2_simps) qed lemma \Runs_table2_terminal_silentsD: assumes Runs: "\Runs_table2 s tlsstlss" and fin: "lfinite (llist_of_tllist tlsstlss)" and terminal: "terminal tlsstlss = Inl (tls, s'')" shows "\s'. silent_moves2 s' tls s''" using fin Runs terminal proof(induct "llist_of_tllist tlsstlss" arbitrary: tlsstlss s) case lfinite_LNil thus ?case by(cases tlsstlss)(auto simp add: \Runs_table2_simps) next case (lfinite_LConsI xs tlsstls) thus ?case by(cases tlsstlss)(auto simp add: \Runs_table2_simps) qed lemma \Runs_table2_terminal_inf_stepD: assumes Runs: "\Runs_table2 s tlsstlss" and fin: "lfinite (llist_of_tllist tlsstlss)" and terminal: "terminal tlsstlss = Inr tls" shows "\s'. trsys.inf_step silent_move2 s' tls" using fin Runs terminal proof(induct "llist_of_tllist tlsstlss" arbitrary: s tlsstlss) case lfinite_LNil thus ?case by(cases tlsstlss)(auto simp add: \Runs_table2_simps) next case (lfinite_LConsI xs tlsstls) thus ?case by(cases tlsstlss)(auto simp add: \Runs_table2_simps) qed lemma \Runs_table2_lappendtD: assumes Runs: "\Runs_table2 s (lappendt tlsstlss tlsstlss')" and fin: "lfinite tlsstlss" shows "\s'. \Runs_table2 s' tlsstlss'" using fin Runs by(induct arbitrary: s)(auto simp add: \Runs_table2_simps) end lemma \moves_False: "\trsys.silent_move r (\s ta s'. False) = (\s s'. False)" by(auto simp add: \trsys.silent_move_iff) lemma \rtrancl3p_False_eq_rtrancl3p: "\trsys.\rtrancl3p r (\s tl s'. False) = rtrancl3p r" proof(intro ext iffI) fix s tls s' assume "\trsys.\rtrancl3p r (\s tl s'. False) s tls s'" thus "rtrancl3p r s tls s'" by(rule \trsys.\rtrancl3p.induct)(blast intro: rtrancl3p_step_converse)+ next fix s tls s' assume "rtrancl3p r s tls s'" thus "\trsys.\rtrancl3p r (\s tl s'. False) s tls s'" by(induct rule: rtrancl3p_converse_induct)(auto intro: \trsys.\rtrancl3p.intros) qed lemma \diverge_empty_\move: "\trsys.\diverge r (\s ta s'. False) = (\s. False)" by(auto intro!: ext elim: \trsys.\diverge.cases \trsys.silent_move.cases) end diff --git a/thys/JinjaThreads/MM/JMM_Common.thy b/thys/JinjaThreads/MM/JMM_Common.thy --- a/thys/JinjaThreads/MM/JMM_Common.thy +++ b/thys/JinjaThreads/MM/JMM_Common.thy @@ -1,1045 +1,1045 @@ (* Title: JinjaThreads/MM/JMM_Common.thy Author: Andreas Lochbihler *) section \JMM Instantiation with Jinja -- common parts\ theory JMM_Common imports JMM_Framework JMM_Typesafe "../Common/BinOp" "../Common/ExternalCallWF" begin context heap begin lemma heap_copy_loc_not_New: assumes "heap_copy_loc a a' al h ob h'" shows "NewHeapElem a'' x \ set ob \ False" using assms by(auto elim: heap_copy_loc.cases) lemma heap_copies_not_New: assumes "heap_copies a a' als h obs h'" and "NewHeapElem a'' x \ set obs" shows "False" using assms by induct(auto dest: heap_copy_loc_not_New) lemma heap_clone_New_same_addr_same: assumes "heap_clone P h a h' \(obs, a')\" and "obs ! i = NewHeapElem a'' x" "i < length obs" and "obs ! j = NewHeapElem a'' x'" "j < length obs" shows "i = j" using assms apply cases apply(fastforce simp add: nth_Cons' gr0_conv_Suc in_set_conv_nth split: if_split_asm dest: heap_copies_not_New)+ done lemma red_external_New_same_addr_same: "\ P,t \ \a\M(vs), h\ -ta\ext \va, h'\; \ta\\<^bsub>o\<^esub> ! i = NewHeapElem a' x; i < length \ta\\<^bsub>o\<^esub>; \ta\\<^bsub>o\<^esub> ! j = NewHeapElem a' x'; j < length \ta\\<^bsub>o\<^esub> \ \ i = j" by(auto elim!: red_external.cases simp add: nth_Cons' split: if_split_asm dest: heap_clone_New_same_addr_same) lemma red_external_aggr_New_same_addr_same: "\ (ta, va, h') \ red_external_aggr P t a M vs h; \ta\\<^bsub>o\<^esub> ! i = NewHeapElem a' x; i < length \ta\\<^bsub>o\<^esub>; \ta\\<^bsub>o\<^esub> ! j = NewHeapElem a' x'; j < length \ta\\<^bsub>o\<^esub> \ \ i = j" by(auto simp add: external_WT_defs.simps red_external_aggr_def nth_Cons' split: if_split_asm if_split_asm dest: heap_clone_New_same_addr_same) lemma heap_copy_loc_read_typeable: assumes "heap_copy_loc a a' al h obs h'" and "ReadMem ad al' v \ set obs" and "P,h \ a@al : T" shows "ad = a \ al'= al" using assms by cases auto lemma heap_copies_read_typeable: assumes "heap_copies a a' als h obs h'" and "ReadMem ad al' v \ set obs" and "list_all2 (\al T. P,h \ a@al : T) als Ts" shows "ad = a \ al' \ set als" using assms proof(induct arbitrary: Ts) case Nil thus ?case by simp next case (Cons al h ob h' als obs h'') from \list_all2 (\al T. P,h \ a@al : T) (al # als) Ts\ obtain T Ts' where Ts [simp]: "Ts = T # Ts'" and "P,h \ a@al : T" and "list_all2 (\al T. P,h \ a@al : T) als Ts'" by(auto simp add: list_all2_Cons1) from \ReadMem ad al' v \ set (ob @ obs)\ show ?case unfolding set_append Un_iff proof assume "ReadMem ad al' v \ set ob" with \heap_copy_loc a a' al h ob h'\ have "ad = a \ al'= al" using \P,h \ a@al : T\ by(rule heap_copy_loc_read_typeable) thus ?thesis by simp next assume "ReadMem ad al' v \ set obs" moreover from \heap_copy_loc a a' al h ob h'\ have "h \ h'" by(rule hext_heap_copy_loc) from \list_all2 (\al T. P,h \ a@al : T) als Ts'\ have "list_all2 (\al T. P,h' \ a@al : T) als Ts'" by(rule List.list_all2_mono)(rule addr_loc_type_hext_mono[OF _ \h \ h'\]) ultimately have "ad = a \ al' \ set als" by(rule Cons) thus ?thesis by simp qed qed lemma heap_clone_read_typeable: assumes clone: "heap_clone P h a h' \(obs, a')\" and read: "ReadMem ad al v \ set obs" shows "ad = a \ (\T'. P,h \ ad@al : T')" using clone proof cases case (ObjClone C H' FDTs obs') let ?als = "map (\((F, D), Tm). CField D F) FDTs" let ?Ts = "map (\(FD, T). fst (the (map_of FDTs FD))) FDTs" note \heap_copies a a' ?als H' obs' h'\ moreover from \obs = NewHeapElem a' (Class_type C) # obs'\ read have "ReadMem ad al v \ set obs'" by simp moreover from \(H', a') \ allocate h (Class_type C)\ have "h \ H'" by(rule hext_allocate) hence "typeof_addr H' a = \Class_type C\" using \typeof_addr h a = \Class_type C\\ by(rule typeof_addr_hext_mono) hence type: "list_all2 (\al T. P,H' \ a@al : T) ?als ?Ts" using \P \ C has_fields FDTs\ - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce intro: addr_loc_type.intros simp add: has_field_def dest: weak_map_of_SomeI) ultimately have "ad = a \ al \ set ?als" by(rule heap_copies_read_typeable) hence [simp]: "ad = a" and "al \ set ?als" by simp_all then obtain F D T where [simp]: "al = CField D F" and "((F, D), T) \ set FDTs" by auto with type \h \ H'\ \typeof_addr h a = \Class_type C\\ show ?thesis - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce elim!: ballE[where x="((F, D), T)"] addr_loc_type.cases dest: typeof_addr_hext_mono intro: addr_loc_type.intros) next case (ArrClone T n H' FDTs obs') let ?als = "map (\((F, D), Tfm). CField D F) FDTs @ map ACell [0..P \ Object has_fields FDTs\ note \heap_copies a a' ?als H' obs' h'\ moreover from \obs = NewHeapElem a' (Array_type T n) # obs'\ read have "ReadMem ad al v \ set obs'" by simp moreover from \(H', a') \ allocate h (Array_type T n)\ have "h \ H'" by(rule hext_allocate) with \typeof_addr h a = \Array_type T n\\ have type': "typeof_addr H' a = \Array_type T n\" by(auto dest: typeof_addr_hext_mono hext_arrD) hence type: "list_all2 (\al T. P,H' \ a@al : T) ?als ?Ts" using FDTs - by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv dest: weak_map_of_SomeI) + by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_same dest: weak_map_of_SomeI) ultimately have "ad = a \ al \ set ?als" by(rule heap_copies_read_typeable) hence [simp]: "ad = a" and "al \ set ?als" by simp_all hence "al \ set (map (\((F, D), Tfm). CField D F) FDTs) \ al \ set (map ACell [0.. set (map (\((F, D), Tfm). CField D F) FDTs)" then obtain F D Tfm where [simp]: "al = CField D F" and "((F, D), Tfm) \ set FDTs" by auto with type type' \h \ H'\ \typeof_addr h a = \Array_type T n\\ show ?thesis - by(fastforce elim!: ballE[where x="((F, D), Tfm)"] addr_loc_type.cases intro: addr_loc_type.intros simp add: list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv) + by(fastforce elim!: ballE[where x="((F, D), Tfm)"] addr_loc_type.cases intro: addr_loc_type.intros simp add: list_all2_append list_all2_map1 list_all2_map2 list_all2_same) next assume "al \ set (map ACell [0..h \ H'\ \typeof_addr h a = \Array_type T n\\ show ?thesis by(fastforce dest: list_all2_nthD[where p=n'] elim: addr_loc_type.cases intro: addr_loc_type.intros) qed qed lemma red_external_read_mem_typeable: assumes red: "P,t \ \a\M(vs), h\ -ta\ext \va, h'\" and read: "ReadMem ad al v \ set \ta\\<^bsub>o\<^esub>" shows "\T'. P,h \ ad@al : T'" using red read by cases(fastforce dest: heap_clone_read_typeable intro: addr_loc_type.intros)+ end context heap_conf begin lemma heap_clone_typeof_addrD: assumes "heap_clone P h a h' \(obs, a')\" and "hconf h" shows "NewHeapElem a'' x \ set obs \ a'' = a' \ typeof_addr h' a' = Some x" using assms by(fastforce elim!: heap_clone.cases dest: allocate_SomeD hext_heap_copies heap_copies_not_New typeof_addr_is_type elim: hext_objD hext_arrD) lemma red_external_New_typeof_addrD: "\ P,t \ \a\M(vs), h\ -ta\ext \va, h'\; NewHeapElem a' x \ set \ta\\<^bsub>o\<^esub>; hconf h \ \ typeof_addr h' a' = Some x" by(erule red_external.cases)(auto dest: heap_clone_typeof_addrD) lemma red_external_aggr_New_typeof_addrD: "\ (ta, va, h') \ red_external_aggr P t a M vs h; NewHeapElem a' x \ set \ta\\<^bsub>o\<^esub>; is_native P (the (typeof_addr h a)) M; hconf h \ \ typeof_addr h' a' = Some x" apply(auto simp add: is_native.simps external_WT_defs.simps red_external_aggr_def split: if_split_asm) apply(blast dest: heap_clone_typeof_addrD)+ done end context heap_conf begin lemma heap_copy_loc_non_speculative_typeable: assumes copy: "heap_copy_loc ad ad' al h obs h'" and sc: "non_speculative P vs (llist_of (map NormalAction obs))" and vs: "vs_conf P h vs" and hconf: "hconf h" and wt: "P,h \ ad@al : T" "P,h \ ad'@al : T" shows "heap_base.heap_copy_loc (heap_read_typed P) heap_write ad ad' al h obs h'" proof - from copy obtain v where obs: "obs = [ReadMem ad al v, WriteMem ad' al v]" and read: "heap_read h ad al v" and "write": "heap_write h ad' al v h'" by cases from obs sc have "v \ vs (ad, al)" by auto with vs wt have v: "P,h \ v :\ T" by(blast dest: vs_confD addr_loc_type_fun)+ with read wt have "heap_read_typed P h ad al v" by(auto intro: heap_read_typedI dest: addr_loc_type_fun) thus ?thesis using "write" unfolding obs by(rule heap_base.heap_copy_loc.intros) qed lemma heap_copy_loc_non_speculative_vs_conf: assumes copy: "heap_copy_loc ad ad' al h obs h'" and sc: "non_speculative P vs (llist_of (take n (map NormalAction obs)))" and vs: "vs_conf P h vs" and hconf: "hconf h" and wt: "P,h \ ad@al : T" "P,h \ ad'@al : T" shows "vs_conf P h' (w_values P vs (take n (map NormalAction obs)))" proof - from copy obtain v where obs: "obs = [ReadMem ad al v, WriteMem ad' al v]" and read: "heap_read h ad al v" and "write": "heap_write h ad' al v h'" by cases from "write" have hext: "h \ h'" by(rule hext_heap_write) with vs have vs': "vs_conf P h' vs" by(rule vs_conf_hext) show ?thesis proof(cases "n > 0") case True with obs sc have "v \ vs (ad, al)" by(auto simp add: take_Cons') with vs wt have v: "P,h \ v :\ T" by(blast dest: vs_confD addr_loc_type_fun)+ with hext wt have "P,h' \ ad'@al : T" "P,h' \ v :\ T" by(blast intro: addr_loc_type_hext_mono conf_hext)+ thus ?thesis using vs' obs by(auto simp add: take_Cons' intro!: vs_confI split: if_split_asm dest: vs_confD) next case False thus ?thesis using vs' by simp qed qed lemma heap_copies_non_speculative_typeable: assumes "heap_copies ad ad' als h obs h'" and "non_speculative P vs (llist_of (map NormalAction obs))" and "vs_conf P h vs" and "hconf h" and "list_all2 (\al T. P,h \ ad@al : T) als Ts" "list_all2 (\al T. P,h \ ad'@al : T) als Ts" shows "heap_base.heap_copies (heap_read_typed P) heap_write ad ad' als h obs h'" using assms proof(induct arbitrary: Ts vs) case Nil show ?case by(auto intro: heap_base.heap_copies.intros) next case (Cons al h ob h' als obs h'') note sc = \non_speculative P vs (llist_of (map NormalAction (ob @ obs)))\ and vs = \vs_conf P h vs\ and hconf = \hconf h\ and wt = \list_all2 (\al T. P,h \ ad@al : T) (al # als) Ts\ \list_all2 (\al T. P,h \ ad'@al : T) (al # als) Ts\ have sc1: "non_speculative P vs (llist_of (map NormalAction ob))" and sc2: "non_speculative P (w_values P vs (map NormalAction ob)) (llist_of (map NormalAction obs))" using sc by(simp_all add: non_speculative_lappend lappend_llist_of_llist_of[symmetric] del: lappend_llist_of_llist_of) from wt obtain T Ts' where Ts: "Ts = T # Ts'" and wt1: "P,h \ ad@al : T" "P,h \ ad'@al : T" and wt2: "list_all2 (\al T. P,h \ ad@al : T) als Ts'" "list_all2 (\al T. P,h \ ad'@al : T) als Ts'" by(auto simp add: list_all2_Cons1) from \heap_copy_loc ad ad' al h ob h'\ sc1 vs hconf wt1 have copy: "heap_base.heap_copy_loc (heap_read_typed P) heap_write ad ad' al h ob h'" by(rule heap_copy_loc_non_speculative_typeable)+ from heap_copy_loc_non_speculative_vs_conf[OF \heap_copy_loc ad ad' al h ob h'\ _ vs hconf wt1, of "length ob"] sc1 have vs': "vs_conf P h' (w_values P vs (map NormalAction ob))" by simp from \heap_copy_loc ad ad' al h ob h'\ have "h \ h'" by(rule hext_heap_copy_loc) with wt2 have wt2': "list_all2 (\al T. P,h' \ ad@al : T) als Ts'" "list_all2 (\al T. P,h' \ ad'@al : T) als Ts'" by -(erule List.list_all2_mono[OF _ addr_loc_type_hext_mono], assumption+)+ from copy hconf wt1 have hconf': "hconf h'" by(rule heap_conf_read.hconf_heap_copy_loc_mono[OF heap_conf_read_heap_read_typed]) from sc2 vs' hconf' wt2' have "heap_base.heap_copies (heap_read_typed P) heap_write ad ad' als h' obs h''" by(rule Cons) with copy show ?case by(rule heap_base.heap_copies.Cons) qed lemma heap_copies_non_speculative_vs_conf: assumes "heap_copies ad ad' als h obs h'" and "non_speculative P vs (llist_of (take n (map NormalAction obs)))" and "vs_conf P h vs" and "hconf h" and "list_all2 (\al T. P,h \ ad@al : T) als Ts" "list_all2 (\al T. P,h \ ad'@al : T) als Ts" shows "vs_conf P h' (w_values P vs (take n (map NormalAction obs)))" using assms proof(induction arbitrary: Ts vs n) case Nil thus ?case by simp next case (Cons al h ob h' als obs h'') note sc = \non_speculative P vs (llist_of (take n (map NormalAction (ob @ obs))))\ and hcl = \heap_copy_loc ad ad' al h ob h'\ and vs = \vs_conf P h vs\ and hconf = \hconf h\ and wt = \list_all2 (\al T. P,h \ ad@al : T) (al # als) Ts\ \list_all2 (\al T. P,h \ ad'@al : T) (al # als) Ts\ let ?vs' = "w_values P vs (take n (map NormalAction ob))" from sc have sc1: "non_speculative P vs (llist_of (take n (map NormalAction ob)))" and sc2: "non_speculative P ?vs' (llist_of (take (n - length ob) (map NormalAction obs)))" by(simp_all add: lappend_llist_of_llist_of[symmetric] non_speculative_lappend del: lappend_llist_of_llist_of) from wt obtain T Ts' where Ts: "Ts = T # Ts'" and wt1: "P,h \ ad@al : T" "P,h \ ad'@al : T" and wt2: "list_all2 (\al T. P,h \ ad@al : T) als Ts'" "list_all2 (\al T. P,h \ ad'@al : T) als Ts'" by(auto simp add: list_all2_Cons1) from hcl sc1 vs hconf wt1 have vs': "vs_conf P h' ?vs'" by(rule heap_copy_loc_non_speculative_vs_conf) show ?case proof(cases "n < length ob") case True from \heap_copies ad ad' als h' obs h''\ have "h' \ h''" by(rule hext_heap_copies) with vs' have "vs_conf P h'' ?vs'" by(rule vs_conf_hext) thus ?thesis using True by simp next case False note sc2 vs' moreover from False sc1 have sc1': "non_speculative P vs (llist_of (map NormalAction ob))" by simp with hcl have "heap_base.heap_copy_loc (heap_read_typed P) heap_write ad ad' al h ob h'" using vs hconf wt1 by(rule heap_copy_loc_non_speculative_typeable) hence "hconf h'" using hconf wt1 by(rule heap_conf_read.hconf_heap_copy_loc_mono[OF heap_conf_read_heap_read_typed]) moreover from hcl have "h \ h'" by(rule hext_heap_copy_loc) with wt2 have wt2': "list_all2 (\al T. P,h' \ ad@al : T) als Ts'" "list_all2 (\al T. P,h' \ ad'@al : T) als Ts'" by -(erule List.list_all2_mono[OF _ addr_loc_type_hext_mono], assumption+)+ ultimately have "vs_conf P h'' (w_values P ?vs' (take (n - length ob) (map NormalAction obs)))" by(rule Cons.IH) with False show ?thesis by simp qed qed lemma heap_clone_non_speculative_typeable_Some: assumes clone: "heap_clone P h ad h' \(obs, ad')\" and sc: "non_speculative P vs (llist_of (map NormalAction obs))" and vs: "vs_conf P h vs" and hconf: "hconf h" shows "heap_base.heap_clone allocate typeof_addr (heap_read_typed P) heap_write P h ad h' \(obs, ad')\" using clone proof(cases) case (ObjClone C h'' FDTs obs') note FDTs = \P \ C has_fields FDTs\ and obs = \obs = NewHeapElem ad' (Class_type C) # obs'\ let ?als = "map (\((F, D), Tfm). CField D F) FDTs" let ?Ts = "map (\(FD, T). fst (the (map_of FDTs FD))) FDTs" let ?vs = "w_value P vs (NormalAction (NewHeapElem ad' (Class_type C) :: ('addr, 'thread_id) obs_event))" from \(h'', ad') \ allocate h (Class_type C)\ have hext: "h \ h''" by(rule hext_heap_ops) hence type: "typeof_addr h'' ad = \Class_type C\" using \typeof_addr h ad = \Class_type C\\ by(rule typeof_addr_hext_mono) note \heap_copies ad ad' ?als h'' obs' h'\ moreover from sc have "non_speculative P ?vs (llist_of (map NormalAction obs'))" by(simp add: obs) moreover from \P \ C has_fields FDTs\ have "is_class P C" by(rule has_fields_is_class) hence "is_htype P (Class_type C)" by simp with vs \(h'', ad') \ allocate h (Class_type C)\ have "vs_conf P h'' ?vs" by(rule vs_conf_allocate) moreover from \(h'', ad') \ allocate h (Class_type C)\ hconf \is_htype P (Class_type C)\ have "hconf h''" by(rule hconf_allocate_mono) moreover from type FDTs have "list_all2 (\al T. P,h'' \ ad@al : T) ?als ?Ts" - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce intro: addr_loc_type.intros simp add: has_field_def dest: weak_map_of_SomeI) moreover from \(h'', ad') \ allocate h (Class_type C)\ \is_htype P (Class_type C)\ have "typeof_addr h'' ad' = \Class_type C\" by(auto dest: allocate_SomeD) with FDTs have "list_all2 (\al T. P,h'' \ ad'@al : T) ?als ?Ts" - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce intro: addr_loc_type.intros simp add: has_field_def dest: weak_map_of_SomeI) ultimately have copy: "heap_base.heap_copies (heap_read_typed P) heap_write ad ad' (map (\((F, D), Tfm). CField D F) FDTs) h'' obs' h'" by(rule heap_copies_non_speculative_typeable)+ from \typeof_addr h ad = \Class_type C\\ \(h'', ad') \ allocate h (Class_type C)\ FDTs copy show ?thesis unfolding obs by(rule heap_base.heap_clone.intros) next case (ArrClone T n h'' FDTs obs') note obs = \obs = NewHeapElem ad' (Array_type T n) # obs'\ and new = \(h'', ad') \ allocate h (Array_type T n)\ and FDTs = \P \ Object has_fields FDTs\ let ?als = "map (\((F, D), Tfm). CField D F) FDTs @ map ACell [0.. h''" by(rule hext_heap_ops) hence type: "typeof_addr h'' ad = \Array_type T n\" using \typeof_addr h ad = \Array_type T n\\ by(rule typeof_addr_hext_mono) note \heap_copies ad ad' ?als h'' obs' h'\ moreover from sc have "non_speculative P ?vs (llist_of (map NormalAction obs'))" by(simp add: obs) moreover from \typeof_addr h ad = \Array_type T n\\ \hconf h\ have "is_htype P (Array_type T n)" by(auto dest: typeof_addr_is_type) with vs new have "vs_conf P h'' ?vs" by(rule vs_conf_allocate) moreover from new hconf \is_htype P (Array_type T n)\ have "hconf h''" by(rule hconf_allocate_mono) moreover from type FDTs have "list_all2 (\al T. P,h'' \ ad@al : T) ?als ?Ts" - by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv dest: weak_map_of_SomeI) + by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_same dest: weak_map_of_SomeI) moreover from new \is_htype P (Array_type T n)\ have "typeof_addr h'' ad' = \Array_type T n\" by(auto dest: allocate_SomeD) hence "list_all2 (\al T. P,h'' \ ad'@al : T) ?als ?Ts" using FDTs - by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv dest: weak_map_of_SomeI) + by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_same dest: weak_map_of_SomeI) ultimately have copy: "heap_base.heap_copies (heap_read_typed P) heap_write ad ad' (map (\((F, D), Tfm). CField D F) FDTs @ map ACell [0..typeof_addr h ad = \Array_type T n\\ new FDTs copy show ?thesis unfolding obs by(rule heap_base.heap_clone.ArrClone) qed lemma heap_clone_non_speculative_vs_conf_Some: assumes clone: "heap_clone P h ad h' \(obs, ad')\" and sc: "non_speculative P vs (llist_of (take n (map NormalAction obs)))" and vs: "vs_conf P h vs" and hconf: "hconf h" shows "vs_conf P h' (w_values P vs (take n (map NormalAction obs)))" using clone proof(cases) case (ObjClone C h'' FDTs obs') note FDTs = \P \ C has_fields FDTs\ and obs = \obs = NewHeapElem ad' (Class_type C) # obs'\ let ?als = "map (\((F, D), Tfm). CField D F) FDTs" let ?Ts = "map (\(FD, T). fst (the (map_of FDTs FD))) FDTs" let ?vs = "w_value P vs (NormalAction (NewHeapElem ad' (Class_type C) :: ('addr, 'thread_id) obs_event))" from \(h'', ad') \ allocate h (Class_type C)\ have hext: "h \ h''" by(rule hext_heap_ops) hence type: "typeof_addr h'' ad = \Class_type C\" using \typeof_addr h ad = \Class_type C\\ by(rule typeof_addr_hext_mono) note \heap_copies ad ad' ?als h'' obs' h'\ moreover from sc have "non_speculative P ?vs (llist_of (take (n - 1) (map NormalAction obs')))" by(simp add: obs take_Cons' split: if_split_asm) moreover from \P \ C has_fields FDTs\ have "is_class P C" by(rule has_fields_is_class) hence "is_htype P (Class_type C)" by simp with vs \(h'', ad') \ allocate h (Class_type C)\ have "vs_conf P h'' ?vs" by(rule vs_conf_allocate) moreover from \(h'', ad') \ allocate h (Class_type C)\ hconf \is_htype P (Class_type C)\ have "hconf h''" by(rule hconf_allocate_mono) moreover from type FDTs have "list_all2 (\al T. P,h'' \ ad@al : T) ?als ?Ts" - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce intro: addr_loc_type.intros simp add: has_field_def dest: weak_map_of_SomeI) moreover from \(h'', ad') \ allocate h (Class_type C)\ \is_htype P (Class_type C)\ have "typeof_addr h'' ad' = \Class_type C\" by(auto dest: allocate_SomeD) with FDTs have "list_all2 (\al T. P,h'' \ ad'@al : T) ?als ?Ts" - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce intro: addr_loc_type.intros simp add: has_field_def dest: weak_map_of_SomeI) ultimately have vs': "vs_conf P h' (w_values P ?vs (take (n - 1) (map NormalAction obs')))" by(rule heap_copies_non_speculative_vs_conf) show ?thesis proof(cases "n > 0") case True with obs vs' show ?thesis by(simp add: take_Cons') next case False from \heap_copies ad ad' ?als h'' obs' h'\ have "h'' \ h'" by(rule hext_heap_copies) with \h \ h''\ have "h \ h'" by(rule hext_trans) with vs have "vs_conf P h' vs" by(rule vs_conf_hext) thus ?thesis using False by simp qed next case (ArrClone T N h'' FDTs obs') note obs = \obs = NewHeapElem ad' (Array_type T N) # obs'\ and new = \(h'', ad') \ allocate h (Array_type T N)\ and FDTs = \P \ Object has_fields FDTs\ let ?als = "map (\((F, D), Tfm). CField D F) FDTs @ map ACell [0.. h''" by(rule hext_heap_ops) hence type: "typeof_addr h'' ad = \Array_type T N\" using \typeof_addr h ad = \Array_type T N\\ by(rule typeof_addr_hext_mono) note \heap_copies ad ad' ?als h'' obs' h'\ moreover from sc have "non_speculative P ?vs (llist_of (take (n - 1) (map NormalAction obs')))" by(simp add: obs take_Cons' split: if_split_asm) moreover from \typeof_addr h ad = \Array_type T N\\ \hconf h\ have "is_htype P (Array_type T N)" by(auto dest: typeof_addr_is_type) with vs new have "vs_conf P h'' ?vs" by(rule vs_conf_allocate) moreover from new hconf \is_htype P (Array_type T N)\ have "hconf h''" by(rule hconf_allocate_mono) moreover from type FDTs have "list_all2 (\al T. P,h'' \ ad@al : T) ?als ?Ts" - by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv dest: weak_map_of_SomeI) + by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_same dest: weak_map_of_SomeI) moreover from new \is_htype P (Array_type T N)\ have "typeof_addr h'' ad' = \Array_type T N\" by(auto dest: allocate_SomeD) hence "list_all2 (\al T. P,h'' \ ad'@al : T) ?als ?Ts" using FDTs - by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv dest: weak_map_of_SomeI) + by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_same dest: weak_map_of_SomeI) ultimately have vs': "vs_conf P h' (w_values P ?vs (take (n - 1) (map NormalAction obs')))" by(rule heap_copies_non_speculative_vs_conf) show ?thesis proof(cases "n > 0") case True with obs vs' show ?thesis by(simp add: take_Cons') next case False from \heap_copies ad ad' ?als h'' obs' h'\ have "h'' \ h'" by(rule hext_heap_copies) with \h \ h''\ have "h \ h'" by(rule hext_trans) with vs have "vs_conf P h' vs" by(rule vs_conf_hext) thus ?thesis using False by simp qed qed lemma heap_clone_non_speculative_typeable_None: assumes "heap_clone P h ad h' None" shows "heap_base.heap_clone allocate typeof_addr (heap_read_typed P) heap_write P h ad h' None" using assms by(cases)(blast intro: heap_base.heap_clone.intros)+ lemma red_external_non_speculative_typeable: assumes red: "P,t \ \a\M(vs), h\ -ta\ext \va, h'\" and sc: "non_speculative P Vs (llist_of (map NormalAction \ta\\<^bsub>o\<^esub>))" and vs: "vs_conf P h Vs" and hconf: "hconf h" shows "heap_base.red_external addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr (heap_read_typed P) heap_write P t h a M vs ta va h'" using assms by(cases)(auto intro: heap_base.red_external.intros heap_clone_non_speculative_typeable_None heap_clone_non_speculative_typeable_Some dest: hext_heap_clone elim: vs_conf_hext) lemma red_external_non_speculative_vs_conf: assumes red: "P,t \ \a\M(vs), h\ -ta\ext \va, h'\" and sc: "non_speculative P Vs (llist_of (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" and vs: "vs_conf P h Vs" and hconf: "hconf h" shows "vs_conf P h' (w_values P Vs (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" using assms by(cases)(auto intro: heap_base.red_external.intros heap_clone_non_speculative_vs_conf_Some dest: hext_heap_clone elim: vs_conf_hext simp add: take_Cons') lemma red_external_aggr_non_speculative_typeable: assumes red: "(ta, va, h') \ red_external_aggr P t a M vs h" and sc: "non_speculative P Vs (llist_of (map NormalAction \ta\\<^bsub>o\<^esub>))" and vs: "vs_conf P h Vs" and hconf: "hconf h" and native: "is_native P (the (typeof_addr h a)) M" shows "(ta, va, h') \ heap_base.red_external_aggr addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr (heap_read_typed P) heap_write P t a M vs h" using assms by(cases "the (typeof_addr h a)")(auto 4 3 simp add: is_native.simps external_WT_defs.simps red_external_aggr_def heap_base.red_external_aggr_def split: if_split_asm split del: if_split del: disjCI intro: heap_clone_non_speculative_typeable_None heap_clone_non_speculative_typeable_Some dest: sees_method_decl_above) lemma red_external_aggr_non_speculative_vs_conf: assumes red: "(ta, va, h') \ red_external_aggr P t a M vs h" and sc: "non_speculative P Vs (llist_of (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" and vs: "vs_conf P h Vs" and hconf: "hconf h" and native: "is_native P (the (typeof_addr h a)) M" shows "vs_conf P h' (w_values P Vs (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" using assms by(cases "the (typeof_addr h a)")(auto 4 3 simp add: is_native.simps external_WT_defs.simps red_external_aggr_def heap_base.red_external_aggr_def take_Cons' split: if_split_asm split del: if_split del: disjCI intro: heap_clone_non_speculative_vs_conf_Some dest: hext_heap_clone elim: vs_conf_hext dest: sees_method_decl_above) end declare split_paired_Ex [simp del] declare eq_upto_seq_inconsist_simps [simp] context heap_progress begin lemma heap_copy_loc_non_speculative_read: assumes hrt: "heap_read_typeable hconf P" and vs: "vs_conf P h vs" and type: "P,h \ a@al : T" "P,h \ a'@al : T" and hconf: "hconf h" and copy: "heap_copy_loc a a' al h obs h'" and i: "i < length obs" and read: "obs ! i = ReadMem a'' al'' v" and v: "v' \ w_values P vs (map NormalAction (take i obs)) (a'', al'')" shows "\obs' h''. heap_copy_loc a a' al h obs' h'' \ i < length obs' \ take i obs' = take i obs \ obs' ! i = ReadMem a'' al'' v' \ length obs' \ length obs \ non_speculative P vs (llist_of (map NormalAction obs'))" using copy proof cases case (1 v'') with read i have [simp]: "i = 0" "v'' = v" "a'' = a" "al'' = al" by(simp_all add: nth_Cons split: nat.split_asm) from v have "v' \ vs (a, al)" by simp with vs type have conf: "P,h \ v' :\ T" by(auto dest: addr_loc_type_fun vs_confD) let ?obs'' = "[ReadMem a al v', WriteMem a' al v']" from hrt type(1) conf hconf have "heap_read h a al v'" by(rule heap_read_typeableD) moreover from heap_write_total[OF hconf type(2) conf] obtain h'' where "heap_write h a' al v' h''" .. ultimately have "heap_copy_loc a a' al h ?obs'' h''" .. thus ?thesis using 1 \v' \ vs (a, al)\ by(auto) qed lemma heap_copies_non_speculative_read: assumes hrt: "heap_read_typeable hconf P" and copies: "heap_copies a a' als h obs h'" and vs: "vs_conf P h vs" and type1: "list_all2 (\al T. P,h \ a@al : T) als Ts" and type2: "list_all2 (\al T. P,h \ a'@al : T) als Ts" and hconf: "hconf h" and i: "i < length obs" and read: "obs ! i = ReadMem a'' al'' v" and v: "v' \ w_values P vs (map NormalAction (take i obs)) (a'', al'')" and ns: "non_speculative P vs (llist_of (map NormalAction (take i obs)))" shows "\obs' h''. heap_copies a a' als h obs' h'' \ i < length obs' \ take i obs' = take i obs \ obs' ! i = ReadMem a'' al'' v' \ length obs' \ length obs" (is "?concl als h obs vs i") using copies vs type1 type2 hconf i read v ns proof(induction arbitrary: Ts vs i) case Nil thus ?case by simp next case (Cons al h ob h' als obs h'' Ts vs) note copy = \heap_copy_loc a a' al h ob h'\ note vs = \vs_conf P h vs\ note type1 = \list_all2 (\al T. P,h \ a@al : T) (al # als) Ts\ and type2 = \list_all2 (\al T. P,h \ a'@al : T) (al # als) Ts\ note hconf = \hconf h\ note i = \i < length (ob @ obs)\ note read = \(ob @ obs) ! i = ReadMem a'' al'' v\ note v = \v' \ w_values P vs (map NormalAction (take i (ob @ obs))) (a'', al'')\ note ns = \non_speculative P vs (llist_of (map NormalAction (take i (ob @ obs))))\ from type1 obtain T Ts' where Ts: "Ts = T # Ts'" and type1': "P,h \ a@al : T" and type1'': "list_all2 (\al T. P,h \ a@al : T) als Ts'" by(auto simp add: list_all2_Cons1) from type2 Ts have type2': "P,h \ a'@al : T" and type2'': "list_all2 (\al T. P,h \ a'@al : T) als Ts'" by simp_all show ?case proof(cases "i < length ob") case True with read v have "ob ! i = ReadMem a'' al'' v" and "v' \ w_values P vs (map NormalAction (take i ob)) (a'', al'')" by(simp_all add: nth_append) from heap_copy_loc_non_speculative_read[OF hrt vs type1' type2' hconf copy True this] obtain ob' H'' where copy': "heap_copy_loc a a' al h ob' H''" and i': "i < length ob'" and "take i ob' = take i ob" and "ob' ! i = ReadMem a'' al'' v'" and "length ob' \ length ob" and ns: "non_speculative P vs (llist_of (map NormalAction ob'))" by blast moreover { from copy' have hext: "h \ H''" by(rule hext_heap_copy_loc) have "hconf H''" by(rule heap_conf_read.hconf_heap_copy_loc_mono[OF heap_conf_read_heap_read_typed])(rule heap_copy_loc_non_speculative_typeable[OF copy' ns vs hconf type1' type2'], fact+) moreover from type1'' have "list_all2 (\al T. P,H'' \ a@al : T) als Ts'" by(rule List.list_all2_mono)(rule addr_loc_type_hext_mono[OF _ hext]) moreover from type2'' have "list_all2 (\al T. P,H'' \ a'@al : T) als Ts'" by(rule List.list_all2_mono)(rule addr_loc_type_hext_mono[OF _ hext]) moreover note calculation } from heap_copies_progress[OF this] obtain obs' h''' where *: "heap_copies a a' als H'' obs' h'''" by blast moreover note heap_copies_length[OF *] moreover note heap_copy_loc_length[OF copy'] moreover note heap_copies_length[OF \heap_copies a a' als h' obs h''\] ultimately show ?thesis using True by(auto intro!: heap_copies.Cons exI simp add: nth_append) next case False let ?vs' = "w_values P vs (map NormalAction ob)" let ?i' = "i - length ob" from ns False obtain ns': "non_speculative P vs (llist_of (map NormalAction ob))" and ns'': "non_speculative P ?vs' (llist_of (map NormalAction (take ?i' obs)))" by(simp add: lappend_llist_of_llist_of[symmetric] non_speculative_lappend del: lappend_llist_of_llist_of) from heap_copy_loc_non_speculative_vs_conf[OF copy _ vs hconf type1' type2', where n="length ob"] ns' have "vs_conf P h' ?vs'" by simp moreover from copy have hext: "h \ h'" by(rule hext_heap_copy_loc) from type1'' have "list_all2 (\al T. P,h' \ a@al : T) als Ts'" by(rule List.list_all2_mono)(rule addr_loc_type_hext_mono[OF _ hext]) moreover from type2'' have "list_all2 (\al T. P,h' \ a'@al : T) als Ts'" by(rule List.list_all2_mono)(rule addr_loc_type_hext_mono[OF _ hext]) moreover have "hconf h'" by(rule heap_conf_read.hconf_heap_copy_loc_mono[OF heap_conf_read_heap_read_typed])(rule heap_copy_loc_non_speculative_typeable[OF copy ns' vs hconf type1' type2'], fact+) moreover from i False have "?i' < length obs" by simp moreover from read False have "obs ! ?i' = ReadMem a'' al'' v" by(simp add: nth_append) moreover from v False have "v' \ w_values P ?vs' (map NormalAction (take ?i' obs)) (a'', al'')" by(simp) ultimately have "?concl als h' obs ?vs' ?i'" using ns'' by(rule Cons.IH) thus ?thesis using False copy by safe(auto intro!: heap_copies.Cons exI simp add: nth_append) qed qed lemma heap_clone_non_speculative_read: assumes hrt: "heap_read_typeable hconf P" and clone: "heap_clone P h a h' \(obs, a')\" and vs: "vs_conf P h vs" and hconf: "hconf h" and i: "i < length obs" and read: "obs ! i = ReadMem a'' al'' v" and v: "v' \ w_values P vs (map NormalAction (take i obs)) (a'', al'')" and ns: "non_speculative P vs (llist_of (map NormalAction (take i obs)))" shows "\obs' h''. heap_clone P h a h'' \(obs', a')\ \ i < length obs' \ take i obs' = take i obs \ obs' ! i = ReadMem a'' al'' v' \ length obs' \ length obs" using clone proof cases case (ObjClone C h'' FDTs obs') note obs = \obs = NewHeapElem a' (Class_type C) # obs'\ note FDTs = \P \ C has_fields FDTs\ let ?als = "map (\((F, D), Tm). CField D F) FDTs" let ?Ts = "map (\(FD, T). fst (the (map_of FDTs FD))) FDTs" let ?vs = "w_value P vs (NormalAction (NewHeapElem a' (Class_type C)) :: ('addr, 'thread_id) obs_event action)" let ?i = "i - 1" from i read obs have i_0: "i > 0" by(simp add: nth_Cons' split: if_split_asm) from \P \ C has_fields FDTs\ have "is_class P C" by(rule has_fields_is_class) with \(h'', a') \ allocate h (Class_type C)\ have type_a': "typeof_addr h'' a' = \Class_type C\" and hext: "h \ h''" by(auto dest: allocate_SomeD hext_allocate) note \heap_copies a a' ?als h'' obs' h'\ moreover from \typeof_addr h a = \Class_type C\\ hconf have "is_htype P (Class_type C)" by(rule typeof_addr_is_type) with vs \(h'', a') \ allocate h (Class_type C)\ have "vs_conf P h'' ?vs" by(rule vs_conf_allocate) moreover from hext \typeof_addr h a = \Class_type C\\ have "typeof_addr h'' a = \Class_type C\" by(rule typeof_addr_hext_mono) hence "list_all2 (\al T. P,h'' \ a@al : T) ?als ?Ts" using FDTs - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce intro: addr_loc_type.intros simp add: has_field_def dest: weak_map_of_SomeI) moreover from FDTs type_a' have "list_all2 (\al T. P,h'' \ a'@al : T) ?als ?Ts" - unfolding list_all2_map1 list_all2_map2 list_all2_refl_conv + unfolding list_all2_map1 list_all2_map2 list_all2_same by(fastforce intro: addr_loc_type.intros simp add: has_field_def dest: weak_map_of_SomeI) moreover from \(h'', a') \ allocate h (Class_type C)\ hconf \is_htype P (Class_type C)\ have "hconf h''" by(rule hconf_allocate_mono) moreover from i read i_0 obs have "?i < length obs'" "obs' ! ?i = ReadMem a'' al'' v" by simp_all moreover from v i_0 obs have "v' \ w_values P ?vs (map NormalAction (take ?i obs')) (a'', al'')" by(simp add: take_Cons') moreover from ns i_0 obs have "non_speculative P ?vs (llist_of (map NormalAction (take ?i obs')))" by(simp add: take_Cons') ultimately have "\obs'' h'''. heap_copies a a' ?als h'' obs'' h''' \ ?i < length obs'' \ take ?i obs'' = take ?i obs' \ obs'' ! ?i = ReadMem a'' al'' v' \ length obs'' \ length obs'" by(rule heap_copies_non_speculative_read[OF hrt]) thus ?thesis using \typeof_addr h a = \Class_type C\\ \(h'', a') \ allocate h (Class_type C)\ FDTs obs i_0 by(auto 4 4 intro: heap_clone.ObjClone simp add: take_Cons') next case (ArrClone T n h'' FDTs obs') note obs = \obs = NewHeapElem a' (Array_type T n) # obs'\ note FDTs = \P \ Object has_fields FDTs\ let ?als = "map (\((F, D), Tfm). CField D F) FDTs @ map ACell [0.. 0" by(simp add: nth_Cons' split: if_split_asm) from \typeof_addr h a = \Array_type T n\\ hconf have "is_htype P (Array_type T n)" by(rule typeof_addr_is_type) with \(h'', a') \ allocate h (Array_type T n)\ have type_a': "typeof_addr h'' a' = \Array_type T n\" and hext: "h \ h''" by(auto dest: allocate_SomeD hext_allocate) note \heap_copies a a' ?als h'' obs' h'\ moreover from vs \(h'', a') \ allocate h (Array_type T n)\ \is_htype P (Array_type T n)\ have "vs_conf P h'' ?vs" by(rule vs_conf_allocate) moreover from hext \typeof_addr h a = \Array_type T n\\ have type'a: "typeof_addr h'' a = \Array_type T n\" by(auto intro: hext_arrD) from type'a FDTs have "list_all2 (\al T. P,h'' \ a@al : T) ?als ?Ts" - by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv dest: weak_map_of_SomeI) + by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_same dest: weak_map_of_SomeI) moreover from type_a' FDTs have "list_all2 (\al T. P,h'' \ a'@al : T) ?als ?Ts" - by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_refl_conv dest: weak_map_of_SomeI) + by(fastforce intro: list_all2_all_nthI addr_loc_type.intros simp add: has_field_def list_all2_append list_all2_map1 list_all2_map2 list_all2_same dest: weak_map_of_SomeI) moreover from \(h'', a') \ allocate h (Array_type T n)\ hconf \is_htype P (Array_type T n)\ have "hconf h''" by(rule hconf_allocate_mono) moreover from i read i_0 obs have "?i < length obs'" "obs' ! ?i = ReadMem a'' al'' v" by simp_all moreover from v i_0 obs have "v' \ w_values P ?vs (map NormalAction (take ?i obs')) (a'', al'')" by(simp add: take_Cons') moreover from ns i_0 obs have "non_speculative P ?vs (llist_of (map NormalAction (take ?i obs')))" by(simp add: take_Cons') ultimately have "\obs'' h'''. heap_copies a a' ?als h'' obs'' h''' \ ?i < length obs'' \ take ?i obs'' = take ?i obs' \ obs'' ! ?i = ReadMem a'' al'' v' \ length obs'' \ length obs'" by(rule heap_copies_non_speculative_read[OF hrt]) thus ?thesis using \typeof_addr h a = \Array_type T n\\ \(h'', a') \ allocate h (Array_type T n)\ FDTs obs i_0 by(auto 4 4 intro: heap_clone.ArrClone simp add: take_Cons') qed lemma red_external_non_speculative_read: assumes hrt: "heap_read_typeable hconf P" and vs: "vs_conf P (shr s) vs" and red: "P,t \ \a\M(vs'), shr s\ -ta\ext \va,h'\" and aok: "final_thread.actions_ok final s t ta" and hconf: "hconf (shr s)" and i: "i < length \ta\\<^bsub>o\<^esub>" and read: "\ta\\<^bsub>o\<^esub> ! i = ReadMem a'' al'' v" and v: "v' \ w_values P vs (map NormalAction (take i \ta\\<^bsub>o\<^esub>)) (a'', al'')" and ns: "non_speculative P vs (llist_of (map NormalAction (take i \ta\\<^bsub>o\<^esub>)))" shows "\ta'' va'' h''. P,t \ \a\M(vs'), shr s\ -ta''\ext \va'', h''\ \ final_thread.actions_ok final s t ta'' \ i < length \ta''\\<^bsub>o\<^esub> \ take i \ta''\\<^bsub>o\<^esub> = take i \ta\\<^bsub>o\<^esub> \ \ta''\\<^bsub>o\<^esub> ! i = ReadMem a'' al'' v' \ length \ta''\\<^bsub>o\<^esub> \ length \ta\\<^bsub>o\<^esub>" using red i read proof cases case [simp]: (RedClone obs a') from heap_clone_non_speculative_read[OF hrt \heap_clone P (shr s) a h' \(obs, a')\\ vs hconf, of i a'' al'' v v'] i read v ns show ?thesis using aok by(fastforce intro: red_external.RedClone simp add: final_thread.actions_ok_iff) qed(auto simp add: nth_Cons) lemma red_external_aggr_non_speculative_read: assumes hrt: "heap_read_typeable hconf P" and vs: "vs_conf P (shr s) vs" and red: "(ta, va, h') \ red_external_aggr P t a M vs' (shr s)" and native: "is_native P (the (typeof_addr (shr s) a)) M" and aok: "final_thread.actions_ok final s t ta" and hconf: "hconf (shr s)" and i: "i < length \ta\\<^bsub>o\<^esub>" and read: "\ta\\<^bsub>o\<^esub> ! i = ReadMem a'' al'' v" and v: "v' \ w_values P vs (map NormalAction (take i \ta\\<^bsub>o\<^esub>)) (a'', al'')" and ns: "non_speculative P vs (llist_of (map NormalAction (take i \ta\\<^bsub>o\<^esub>)))" shows "\ta'' va'' h''. (ta'', va'', h'') \ red_external_aggr P t a M vs' (shr s) \ final_thread.actions_ok final s t ta'' \ i < length \ta''\\<^bsub>o\<^esub> \ take i \ta''\\<^bsub>o\<^esub> = take i \ta\\<^bsub>o\<^esub> \ \ta''\\<^bsub>o\<^esub> ! i = ReadMem a'' al'' v' \ length \ta''\\<^bsub>o\<^esub> \ length \ta\\<^bsub>o\<^esub>" using red native aok hconf i read v ns apply(simp add: red_external_aggr_def final_thread.actions_ok_iff ex_disj_distrib conj_disj_distribR split nth_Cons' del: if_split split: if_split_asm disj_split_asm) apply(drule heap_clone_non_speculative_read[OF hrt _ vs hconf, of _ _ _ _ i a'' al'' v v']) apply simp_all apply(fastforce) done end declare split_paired_Ex [simp] declare eq_upto_seq_inconsist_simps [simp del] context allocated_heap begin lemma heap_copy_loc_allocated_same: assumes "heap_copy_loc a a' al h obs h'" shows "allocated h' = allocated h" using assms by cases(auto del: subsetI simp: heap_write_allocated_same) lemma heap_copy_loc_allocated_mono: "heap_copy_loc a a' al h obs h' \ allocated h \ allocated h'" by(simp add: heap_copy_loc_allocated_same) lemma heap_copies_allocated_same: assumes "heap_copies a a' al h obs h'" shows "allocated h' = allocated h" using assms by(induct)(auto simp add: heap_copy_loc_allocated_same) lemma heap_copies_allocated_mono: "heap_copies a a' al h obs h' \ allocated h \ allocated h'" by(simp add: heap_copies_allocated_same) lemma heap_clone_allocated_mono: assumes "heap_clone P h a h' aobs" shows "allocated h \ allocated h'" using assms by cases(blast del: subsetI intro: heap_copies_allocated_mono allocate_allocated_mono intro: subset_trans)+ lemma red_external_allocated_mono: assumes "P,t \ \a\M(vs), h\ -ta\ext \va, h'\" shows "allocated h \ allocated h'" using assms by(cases)(blast del: subsetI intro: heap_clone_allocated_mono heap_write_allocated_same)+ lemma red_external_aggr_allocated_mono: "\ (ta, va, h') \ red_external_aggr P t a M vs h; is_native P (the (typeof_addr h a)) M \ \ allocated h \ allocated h'" by(cases "the (typeof_addr h a)")(auto simp add: is_native.simps external_WT_defs.simps red_external_aggr_def split: if_split_asm dest: heap_clone_allocated_mono sees_method_decl_above) lemma heap_clone_allocatedD: assumes "heap_clone P h a h' \(obs, a')\" and "NewHeapElem a'' x \ set obs" shows "a'' \ allocated h' \ a'' \ allocated h" using assms by cases(auto dest: allocate_allocatedD heap_copies_allocated_mono heap_copies_not_New) lemma red_external_allocatedD: "\ P,t \ \a\M(vs), h\ -ta\ext \va, h'\; NewHeapElem a' x \ set \ta\\<^bsub>o\<^esub> \ \ a' \ allocated h' \ a' \ allocated h" by(erule red_external.cases)(auto dest: heap_clone_allocatedD) lemma red_external_aggr_allocatedD: "\ (ta, va, h') \ red_external_aggr P t a M vs h; NewHeapElem a' x \ set \ta\\<^bsub>o\<^esub>; is_native P (the (typeof_addr h a)) M \ \ a' \ allocated h' \ a' \ allocated h" by(auto simp add: is_native.simps external_WT_defs.simps red_external_aggr_def split: if_split_asm dest: heap_clone_allocatedD sees_method_decl_above) lemma heap_clone_NewHeapElemD: assumes "heap_clone P h a h' \(obs, a')\" and "ad \ allocated h'" and "ad \ allocated h" shows "\CTn. NewHeapElem ad CTn \ set obs" using assms by cases(auto dest!: allocate_allocatedD heap_copies_allocated_same) lemma heap_clone_fail_allocated_same: assumes "heap_clone P h a h' None" shows "allocated h' = allocated h" using assms by(cases)(auto) lemma red_external_NewHeapElemD: "\ P,t \ \a\M(vs), h\ -ta\ext \va, h'\; a' \ allocated h'; a' \ allocated h \ \ \CTn. NewHeapElem a' CTn \ set \ta\\<^bsub>o\<^esub>" by(erule red_external.cases)(auto dest: heap_clone_NewHeapElemD heap_clone_fail_allocated_same) lemma red_external_aggr_NewHeapElemD: "\ (ta, va, h') \ red_external_aggr P t a M vs h; a' \ allocated h'; a' \ allocated h; is_native P (the (typeof_addr h a)) M \ \ \CTn. NewHeapElem a' CTn \ set \ta\\<^bsub>o\<^esub>" by(cases "the (typeof_addr h a)")(auto simp add: is_native.simps external_WT_defs.simps red_external_aggr_def split: if_split_asm dest: heap_clone_fail_allocated_same heap_clone_NewHeapElemD sees_method_decl_above) end context heap_base begin lemma binop_known_addrs: assumes ok: "start_heap_ok" shows "binop bop v1 v2 = \Inl v\ \ ka_Val v \ ka_Val v1 \ ka_Val v2 \ set start_addrs" and "binop bop v1 v2 = \Inr a\ \ a \ ka_Val v1 \ ka_Val v2 \ set start_addrs" apply(cases bop, auto split: if_split_asm)[1] apply(cases bop, auto split: if_split_asm simp add: addr_of_sys_xcpt_start_addr[OF ok]) done lemma heap_copy_loc_known_addrs_ReadMem: assumes "heap_copy_loc a a' al h ob h'" and "ReadMem ad al' v \ set ob" shows "ad = a" using assms by cases simp lemma heap_copies_known_addrs_ReadMem: assumes "heap_copies a a' als h obs h'" and "ReadMem ad al v \ set obs" shows "ad = a" using assms by(induct)(auto dest: heap_copy_loc_known_addrs_ReadMem) lemma heap_clone_known_addrs_ReadMem: assumes "heap_clone P h a h' \(obs, a')\" and "ReadMem ad al v \ set obs" shows "ad = a" using assms by cases(auto dest: heap_copies_known_addrs_ReadMem) lemma red_external_known_addrs_ReadMem: "\ P,t \ \a\M(vs), h\ -ta\ext \va,h'\; ReadMem ad al v \ set \ta\\<^bsub>o\<^esub> \ \ ad \ {thread_id2addr t, a} \ (\(ka_Val ` set vs)) \ set start_addrs" by(erule red_external.cases)(simp_all add: heap_clone_known_addrs_ReadMem) lemma red_external_aggr_known_addrs_ReadMem: "\ (ta, va, h') \ red_external_aggr P t a M vs h; ReadMem ad al v \ set \ta\\<^bsub>o\<^esub> \ \ ad \ {thread_id2addr t, a} \ (\(ka_Val ` set vs)) \ set start_addrs" apply(auto simp add: red_external_aggr_def split: if_split_asm dest: heap_clone_known_addrs_ReadMem) done lemma heap_copy_loc_known_addrs_WriteMem: assumes "heap_copy_loc a a' al h ob h'" and "ob ! n = WriteMem ad al' (Addr a'')" "n < length ob" shows "a'' \ new_obs_addrs (take n ob)" using assms by cases(auto simp add: nth_Cons new_obs_addrs_def split: nat.split_asm) lemma heap_copies_known_addrs_WriteMem: assumes "heap_copies a a' als h obs h'" and "obs ! n = WriteMem ad al (Addr a'')" "n < length obs" shows "a'' \ new_obs_addrs (take n obs)" using assms by(induct arbitrary: n)(auto simp add: nth_append new_obs_addrs_def dest: heap_copy_loc_known_addrs_WriteMem split: if_split_asm) lemma heap_clone_known_addrs_WriteMem: assumes "heap_clone P h a h' \(obs, a')\" and "obs ! n = WriteMem ad al (Addr a'')" "n < length obs" shows "a'' \ new_obs_addrs (take n obs)" using assms by cases(auto simp add: nth_Cons new_obs_addrs_def split: nat.split_asm dest: heap_copies_known_addrs_WriteMem) lemma red_external_known_addrs_WriteMem: "\ P,t \ \a\M(vs), h\ -ta\ext \va,h'\; \ta\\<^bsub>o\<^esub> ! n = WriteMem ad al (Addr a'); n < length \ta\\<^bsub>o\<^esub> \ \ a' \ {thread_id2addr t, a} \ (\(ka_Val ` set vs)) \ set start_addrs \ new_obs_addrs (take n \ta\\<^bsub>o\<^esub>)" by(erule red_external.cases)(auto dest: heap_clone_known_addrs_WriteMem) lemma red_external_aggr_known_addrs_WriteMem: "\ (ta, va, h') \ red_external_aggr P t a M vs h; \ta\\<^bsub>o\<^esub> ! n = WriteMem ad al (Addr a'); n < length \ta\\<^bsub>o\<^esub> \ \ a' \ {thread_id2addr t, a} \ (\(ka_Val ` set vs)) \ set start_addrs \ new_obs_addrs (take n \ta\\<^bsub>o\<^esub>)" apply(auto simp add: red_external_aggr_def split: if_split_asm dest: heap_clone_known_addrs_WriteMem) done lemma red_external_known_addrs_mono: assumes ok: "start_heap_ok" and red: "P,t \ \a\M(vs), h\ -ta\ext \va, h'\" shows "(case va of RetVal v \ ka_Val v | RetExc a \ {a} | RetStaySame \ {}) \ {thread_id2addr t, a} \ (\(ka_Val ` set vs)) \ set start_addrs \ new_obs_addrs \ta\\<^bsub>o\<^esub>" using red by cases(auto simp add: addr_of_sys_xcpt_start_addr[OF ok] new_obs_addrs_def heap_clone.simps) lemma red_external_aggr_known_addrs_mono: assumes ok: "start_heap_ok" and red: "(ta, va, h') \ red_external_aggr P t a M vs h" "is_native P (the (typeof_addr h a)) M" shows "(case va of RetVal v \ ka_Val v | RetExc a \ {a} | RetStaySame \ {}) \ {thread_id2addr t, a} \ (\(ka_Val ` set vs)) \ set start_addrs \ new_obs_addrs \ta\\<^bsub>o\<^esub>" using red apply(cases "the (typeof_addr h a)") apply(auto simp add: red_external_aggr_def addr_of_sys_xcpt_start_addr[OF ok] new_obs_addrs_def heap_clone.simps split: if_split_asm) apply(auto simp add: is_native.simps elim!: external_WT_defs.cases dest: sees_method_decl_above) done lemma red_external_NewThread_idD: "\ P,t \ \a\M(vs), h\ -ta\ext \va, h'\; NewThread t' (C, M', a') h'' \ set \ta\\<^bsub>t\<^esub> \ \ t' = addr2thread_id a \ a' = a" by(erule red_external.cases) simp_all lemma red_external_aggr_NewThread_idD: "\ (ta, va, h') \ red_external_aggr P t a M vs h; NewThread t' (C, M', a') h'' \ set \ta\\<^bsub>t\<^esub> \ \ t' = addr2thread_id a \ a' = a" apply(auto simp add: red_external_aggr_def split: if_split_asm) done end locale heap'' = heap' addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write P for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and P :: "'m prog" + assumes allocate_typeof_addr_SomeD: "\ (h', a) \ allocate h hT; typeof_addr a \ None \ \ typeof_addr a = \hT\" begin lemma heap_copy_loc_New_type_match: "\ h.heap_copy_loc a a' al h obs h'; NewHeapElem ad CTn \ set obs; typeof_addr ad \ None \ \ typeof_addr ad = \CTn\" by(erule h.heap_copy_loc.cases) simp lemma heap_copies_New_type_match: "\ h.heap_copies a a' als h obs h'; NewHeapElem ad CTn \ set obs; typeof_addr ad \ None \ \ typeof_addr ad = \CTn\" by(induct rule: h.heap_copies.induct)(auto dest: heap_copy_loc_New_type_match) lemma heap_clone_New_type_match: "\ h.heap_clone P h a h' \(obs, a')\; NewHeapElem ad CTn \ set obs; typeof_addr ad \ None \ \ typeof_addr ad = \CTn\" by(erule h.heap_clone.cases)(auto dest: allocate_typeof_addr_SomeD heap_copies_New_type_match) lemma red_external_New_type_match: "\ h.red_external P t a M vs h ta va h'; NewHeapElem ad CTn \ set \ta\\<^bsub>o\<^esub>; typeof_addr ad \ None \ \ typeof_addr ad = \CTn\" by(erule h.red_external.cases)(auto dest: heap_clone_New_type_match) lemma red_external_aggr_New_type_match: "\ (ta, va, h') \ h.red_external_aggr P t a M vs h; NewHeapElem ad CTn \ set \ta\\<^bsub>o\<^esub>; typeof_addr ad \ None \ \ typeof_addr ad = \CTn\" by(auto simp add: h.red_external_aggr_def split: if_split_asm dest: heap_clone_New_type_match) end end diff --git a/thys/JinjaThreads/MM/JMM_DRF.thy b/thys/JinjaThreads/MM/JMM_DRF.thy --- a/thys/JinjaThreads/MM/JMM_DRF.thy +++ b/thys/JinjaThreads/MM/JMM_DRF.thy @@ -1,401 +1,401 @@ (* Title: JinjaThreads/MM/JMM_DRF.thy Author: Andreas Lochbihler *) section \The data race free guarantee of the JMM\ theory JMM_DRF imports JMM_Spec begin context drf begin lemma drf_lemma: assumes wf: "P \ (E, ws) \" and E: "E \ \" and sync: "correctly_synchronized P \" and read_before: "\r. r \ read_actions E \ P,E \ ws r \hb r" shows "sequentially_consistent P (E, ws)" proof(rule ccontr) let ?Q = "{r. r \ read_actions E \ \ P,E \ r \mrw ws r}" assume "\ ?thesis" then obtain r where "r \ read_actions E" "\ P,E \ r \mrw ws r" by(auto simp add: sequentially_consistent_def) hence "r \ ?Q" by simp with wf_action_order[of E] obtain r' where "r' \ ?Q" and "(action_order E)\<^sup>\\<^sup>\^** r' r" and r'_min: "\a. (action_order E)\<^sup>\\<^sup>\ a r' \ a \ ?Q" by(rule wfP_minimalE) blast from \r' \ ?Q\ have r': "r' \ read_actions E" and not_mrw: "\ P,E \ r' \mrw ws r'" by blast+ from r' obtain ad al v where obs_r': "action_obs E r' = NormalAction (ReadMem ad al v)" by(cases) auto from wf have ws: "is_write_seen P E ws" and tsa_ok: "thread_start_actions_ok E" by(rule wf_exec_is_write_seenD wf_exec_thread_start_actions_okD)+ from is_write_seenD[OF ws r' obs_r'] have ws_r: "ws r' \ write_actions E" and adal: "(ad, al) \ action_loc P E (ws r')" and v: "v = value_written P E (ws r') (ad, al)" and not_hb: "\ P,E \ r' \hb ws r'" by auto from r' have "P,E \ ws r' \hb r'" by(rule read_before) hence "E \ ws r' \a r'" by(rule happens_before_into_action_order) from not_mrw have "\w'. w' \ write_actions E \ (ad, al) \ action_loc P E w' \ \ P,E \ w' \hb ws r' \ \ P,E \ w' \so ws r' \ \ P,E \ r' \hb w' \ \ P,E \ r' \so w' \ E \ w' \a r'" proof(rule contrapos_np) assume inbetween: "\ ?thesis" note r' moreover from obs_r' have "(ad, al) \ action_loc P E r'" by simp moreover note \E \ ws r' \a r'\ ws_r adal moreover { fix w' assume "w' \ write_actions E" "(ad, al) \ action_loc P E w'" with inbetween have "P,E \ w' \hb ws r' \ P,E \ w' \so ws r' \ P,E \ r' \hb w' \ P,E \ r' \so w' \ \ E \ w' \a r'" by simp moreover from total_onPD[OF total_action_order, of w' E r'] \w' \ write_actions E\ r' have "E \ w' \a r' \ E \ r' \a w'" by(auto dest: read_actions_not_write_actions) ultimately have "E \ w' \a ws r' \ E \ r' \a w'" unfolding sync_order_def by(blast intro: happens_before_into_action_order) } ultimately show "P,E \ r' \mrw ws r'" by(rule most_recent_write_for.intros) qed then obtain w' where w': "w' \ write_actions E" and adal_w': "(ad, al) \ action_loc P E w'" and "\ P,E \ w' \hb ws r'" "\ P,E \ r' \hb w'" "E \ w' \a r'" and so: "\ P,E \ w' \so ws r'" "\ P,E \ r' \so w'" by blast have "ws r' \ w'" using \\ P,E \ w' \hb ws r'\ ws_r by(auto intro: happens_before_refl) have vol: "\ is_volatile P al" proof assume vol_al: "is_volatile P al" with r' obs_r' have "r' \ sactions P E" by cases(rule sactionsI, simp_all) moreover from w' vol_al adal_w' have "w' \ sactions P E" by(cases)(auto intro: sactionsI elim!: is_write_action.cases) ultimately have "P,E \ w' \so r' \ w' = r' \ P,E \ r' \so w'" using total_sync_order[of P E] by(blast dest: total_onPD) moreover have "w' \ r'" using w' r' by(auto dest: read_actions_not_write_actions) ultimately have "P,E \ w' \so r'" using \\ P,E \ r' \so w'\ by simp moreover from ws_r vol_al adal have "ws r' \ sactions P E" by(cases)(auto intro: sactionsI elim!: is_write_action.cases) with total_sync_order[of P E] \w' \ sactions P E\ \\ P,E \ w' \so ws r'\ \ws r' \ w'\ have "P,E \ ws r' \so w'" by(blast dest: total_onPD) ultimately show False using is_write_seenD[OF ws r' obs_r'] w' adal_w' vol_al \ws r' \ w'\ by auto qed { fix a assume "a < r'" and "a \ read_actions E" hence "(action_order E)\<^sup>\\<^sup>\ a r'" using r' obs_r' by(auto intro: action_orderI) from r'_min[OF this] \a \ read_actions E\ have "P,E \ a \mrw ws a" by simp } from \_sequential_completion[OF E wf this, of r'] r' obtain E' ws' where "E' \ \" "P \ (E', ws') \" and eq: "ltake (enat r') E = ltake (enat r') E'" and sc': "sequentially_consistent P (E', ws')" and r'': "action_tid E r' = action_tid E' r'" "action_obs E r' \ action_obs E' r'" and "r' \ actions E'" by auto from \P \ (E', ws') \\ have tsa_ok': "thread_start_actions_ok E'" by(rule wf_exec_thread_start_actions_okD) from \r' \ read_actions E\ have "enat r' < llength E" by(auto elim: read_actions.cases actionsE) moreover from \r' \ actions E'\ have "enat r' < llength E'" by(auto elim: actionsE) ultimately have eq': "ltake (enat (Suc r')) E [\] ltake (enat (Suc r')) E'" using eq[THEN eq_into_sim_actions] r'' by(auto simp add: ltake_Suc_conv_snoc_lnth sim_actions_def split_beta action_tid_def action_obs_def intro!: llist_all2_lappendI) from r' have r'': "r' \ read_actions E'" by(rule read_actions_change_prefix[OF _eq']) simp from obs_r' have "(ad, al) \ action_loc P E r'" by simp hence adal_r'': "(ad, al) \ action_loc P E' r'" by(subst (asm) action_loc_change_prefix[OF eq']) simp from \\ P,E \ w' \hb ws r'\ have "\ is_new_action (action_obs E w')" proof(rule contrapos_nn) assume new_w': "is_new_action (action_obs E w')" show "P,E \ w' \hb ws r'" proof(cases "is_new_action (action_obs E (ws r'))") case True with adal new_w' adal_w' w' ws_r have "ws r' \ new_actions_for P E (ad, al)" "w' \ new_actions_for P E (ad, al)" by(auto simp add: new_actions_for_def) with \E \ \\ have "ws r' = w'" by(rule \_new_actions_for_fun) thus ?thesis using w' by(auto intro: happens_before_refl) next case False with tsa_ok w' ws_r new_w' show ?thesis by(auto intro: happens_before_new_not_new) qed qed with \E \ w' \a r'\ have "w' \ r'" by(auto elim!: action_orderE) moreover from w' r' have "w' \ r'" by(auto intro: read_actions_not_write_actions) ultimately have "w' < r'" by simp with w' have "w' \ write_actions E'" by(auto intro: write_actions_change_prefix[OF _ eq']) hence "w' \ actions E'" by simp from adal_w' \w' < r'\ have "(ad, al) \ action_loc P E' w'" by(subst action_loc_change_prefix[symmetric, OF eq']) simp_all from vol \r' \ read_actions E'\ \w' \ write_actions E'\ \(ad, al) \ action_loc P E' w'\ adal_r'' have "P,E' \ r' \ w'" unfolding non_volatile_conflict_def by auto with sync \E' \ \\ \P \ (E', ws') \\ sc' \r' \ actions E'\ \w' \ actions E'\ have hb'_r'_w': "P,E' \ r' \hb w' \ P,E' \ w' \hb r'" by(rule correctly_synchronizedD[rule_format]) hence "P,E \ r' \hb w' \ P,E \ w' \hb r'" using \w' < r'\ by(auto intro: happens_before_change_prefix[OF _ tsa_ok eq'[symmetric]]) with \\ P,E \ r' \hb w'\ have "P,E \ w' \hb r'" by simp have "P,E \ ws r' \hb w'" proof(cases "is_new_action (action_obs E (ws r'))") case False with \E \ ws r' \a r'\ have "ws r' \ r'" by(auto elim!: action_orderE) moreover from ws_r r' have "ws r' \ r'" by(auto dest: read_actions_not_write_actions) ultimately have "ws r' < r'" by simp with ws_r have "ws r' \ write_actions E'" by(auto intro: write_actions_change_prefix[OF _ eq']) hence "ws r' \ actions E'" by simp from adal \ws r' < r'\ have "(ad, al) \ action_loc P E' (ws r')" by(subst action_loc_change_prefix[symmetric, OF eq']) simp_all hence "P,E' \ ws r' \ w'" using \ws r' \ write_actions E'\ \w' \ write_actions E'\ \(ad, al) \ action_loc P E' w'\ vol unfolding non_volatile_conflict_def by auto with sync \E' \ \\ \P \ (E', ws') \\ sc' \ws r' \ actions E'\ \w' \ actions E'\ have "P,E' \ ws r' \hb w' \ P,E' \ w' \hb ws r'" by(rule correctly_synchronizedD[rule_format]) thus "P,E \ ws r' \hb w'" using \w' < r'\ \ws r' < r'\ \\ P,E \ w' \hb ws r'\ by(auto dest: happens_before_change_prefix[OF _ tsa_ok eq'[symmetric]]) next case True with tsa_ok ws_r w' \\ is_new_action (action_obs E w')\ show "P,E \ ws r' \hb w'" by(auto intro: happens_before_new_not_new) qed moreover from wf have "is_write_seen P E ws" by(rule wf_exec_is_write_seenD) ultimately have "w' = ws r'" using is_write_seenD[OF \is_write_seen P E ws\ \r' \ read_actions E\ obs_r'] \w' \ write_actions E\ \(ad, al) \ action_loc P E w'\ \P,E \ w' \hb r'\ by auto with porder_happens_before[of E P] \\ P,E \ w' \hb ws r'\ ws_r show False by(auto dest: refl_onPD[where a="ws r'"] elim!: porder_onE) qed lemma justified_action_committedD: assumes justified: "P \ (E, ws) weakly_justified_by J" and a: "a \ actions E" obtains n a' where "a = action_translation (J n) a'" "a' \ committed (J n)" proof(atomize_elim) from justified have "actions E = (\n. action_translation (J n) ` committed (J n))" by(simp add: is_commit_sequence_def) with a show "\n a'. a = action_translation (J n) a' \ a' \ committed (J n)" by auto qed theorem drf_weak: assumes sync: "correctly_synchronized P \" and legal: "weakly_legal_execution P \ (E, ws)" shows "sequentially_consistent P (E, ws)" using legal_wf_execD[OF legal] legal_\D[OF legal] sync proof(rule drf_lemma) fix r assume "r \ read_actions E" from legal obtain J where E: "E \ \" and wf_exec: "P \ (E, ws) \" and J: "P \ (E, ws) weakly_justified_by J" and range_J: "range (justifying_exec \ J) \ \" by(rule legal_executionE) let ?E = "\n. justifying_exec (J n)" and ?ws = "\n. justifying_ws (J n)" and ?C = "\n. committed (J n)" and ?\ = "\n. action_translation (J n)" from \r \ read_actions E\ have "r \ actions E" by simp with J obtain n r' where r: "r = action_translation (J n) r'" and r': "r' \ ?C n" by(rule justified_action_committedD) note \r \ read_actions E\ moreover from J have wfan: "wf_action_translation_on (?E n) E (?C n) (?\ n)" by(simp add: wf_action_translations_def) hence "action_obs (?E n) r' \ action_obs E r" using r' unfolding r by(blast dest: wf_action_translation_on_actionD) moreover from J r' have "r' \ actions (?E n)" by(auto simp add: committed_subset_actions_def) ultimately have "r' \ read_actions (?E n)" unfolding r by cases(auto intro: read_actions.intros) hence "P,E \ ws (?\ n r') \hb ?\ n r'" using \r' \ ?C n\ proof(induct n arbitrary: r') case 0 from J have "?C 0 = {}" by(simp add: is_commit_sequence_def) with 0 have False by simp thus ?case .. next case (Suc n r) note r = \r \ read_actions (?E (Suc n))\ from J have wfan: "wf_action_translation_on (?E n) E (?C n) (?\ n)" and wfaSn: "wf_action_translation_on (?E (Suc n)) E (?C (Suc n)) (?\ (Suc n))" by(simp_all add: wf_action_translations_def) from wfaSn have injSn: "inj_on (?\ (Suc n)) (actions (?E (Suc n)))" by(rule wf_action_translation_on_inj_onD) from J have C_sub_A: "?C (Suc n) \ actions (?E (Suc n))" by(simp add: committed_subset_actions_def) from J have wf: "P \ (?E (Suc n), ?ws (Suc n)) \" by(simp add: justification_well_formed_def) moreover from range_J have "?E (Suc n) \ \" by auto ultimately have sc: "sequentially_consistent P (?E (Suc n), ?ws (Suc n))" using sync proof(rule drf_lemma) fix r' assume r': "r' \ read_actions (?E (Suc n))" hence "r' \ actions (?E (Suc n))" by simp show "P,?E (Suc n) \ ?ws (Suc n) r' \hb r'" proof(cases "?\ (Suc n) r' \ ?\ n ` ?C n") case True then obtain r'' where r'': "r'' \ ?C n" and r'_r'': "?\ (Suc n) r' = ?\ n r''" by(auto) from r'' wfan have "action_tid (?E n) r'' = action_tid E (?\ n r'')" and "action_obs (?E n) r'' \ action_obs E (?\ n r'')" by(blast dest: wf_action_translation_on_actionD)+ moreover from J have "?\ n ` ?C n \ ?\ (Suc n) ` ?C (Suc n)" by(simp add: is_commit_sequence_def) with r'' have "?\ (Suc n) r' \ ?\ (Suc n) ` ?C (Suc n)" unfolding r'_r'' by auto hence "r' \ ?C (Suc n)" - unfolding inj_on_image_mem_iff[OF injSn C_sub_A \r' \ actions (?E (Suc n))\] . + unfolding inj_on_image_mem_iff[OF injSn \r' \ actions (?E (Suc n))\ C_sub_A] . with wfaSn have "action_tid (?E (Suc n)) r' = action_tid E (?\ (Suc n) r')" and "action_obs (?E (Suc n)) r' \ action_obs E (?\ (Suc n) r')" by(blast dest: wf_action_translation_on_actionD)+ ultimately have tid: "action_tid (?E n) r'' = action_tid (?E (Suc n)) r'" and obs: "action_obs (?E n) r'' \ action_obs (?E (Suc n)) r'" unfolding r'_r'' by(auto intro: sim_action_trans sim_action_sym) from J have "?C n \ actions (?E n)" by(simp add: committed_subset_actions_def) with r'' have "r'' \ actions (?E n)" by blast with r' obs have "r'' \ read_actions (?E n)" by cases(auto intro: read_actions.intros) hence hb'': "P,E \ ws (?\ n r'') \hb ?\ n r''" using \r'' \ ?C n\ by(rule Suc) have r_conv_inv: "r' = inv_into (actions (?E (Suc n))) (?\ (Suc n)) (?\ n r'')" using \r' \ actions (?E (Suc n))\ unfolding r'_r''[symmetric] by(simp add: inv_into_f_f[OF injSn]) with \r'' \ ?C n\ r' J \r'' \ read_actions (?E n)\ have ws_eq[symmetric]: "?\ (Suc n) (?ws (Suc n) r') = ws (?\ n r'')" - by(simp add: write_seen_committed_def Let_def) + by(simp add: write_seen_committed_def) with r'_r''[symmetric] hb'' have "P,E \ ?\ (Suc n) (?ws (Suc n) r') \hb ?\ (Suc n) r'" by simp moreover from J r' \r' \ committed (J (Suc n))\ have "ws (?\ (Suc n) r') \ ?\ (Suc n) ` ?C (Suc n)" by(rule weakly_justified_write_seen_hb_read_committed) then obtain w' where w': "ws (?\ (Suc n) r') = ?\ (Suc n) w'" and committed_w': "w' \ ?C (Suc n)" by blast with C_sub_A have w'_action: "w' \ actions (?E (Suc n))" by auto hence w'_def: "w' = inv_into (actions (?E (Suc n))) (?\ (Suc n)) (ws (?\ (Suc n) r'))" using injSn unfolding w' by simp from J r' \r' \ committed (J (Suc n))\ have hb_eq: "P,E \ ws (?\ (Suc n) r') \hb ?\ (Suc n) r' \ P,?E (Suc n) \ w' \hb r'" unfolding w'_def by(simp add: happens_before_committed_weak_def) from r' obtain ad al v where "action_obs (?E (Suc n)) r' = NormalAction (ReadMem ad al v)" by(cases) from is_write_seenD[OF wf_exec_is_write_seenD[OF wf] r' this] have "?ws (Suc n) r' \ actions (?E (Suc n))" by(auto) with injSn have "w' = ?ws (Suc n) r'" unfolding w'_def ws_eq[folded r'_r''] by(rule inv_into_f_f) thus ?thesis using hb'' hb_eq w'_action r'_r''[symmetric] w' injSn by simp next case False with J r' show ?thesis by(auto simp add: uncommitted_reads_see_hb_def) qed qed from r have "r \ actions (?E (Suc n))" by simp let ?w = "inv_into (actions (?E (Suc n))) (?\ (Suc n)) (ws (?\ (Suc n) r))" from J r \r \ ?C (Suc n)\ have ws_rE_comm: "ws (?\ (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" by(rule weakly_justified_write_seen_hb_read_committed) hence "?w \ ?C (Suc n)" using C_sub_A by(auto simp add: inv_into_f_f[OF injSn]) with C_sub_A have w: "?w \ actions (?E (Suc n))" by blast from ws_rE_comm C_sub_A have w_eq: "?\ (Suc n) ?w = ws (?\ (Suc n) r)" by(auto simp: f_inv_into_f[where f="?\ (Suc n)"]) from r obtain ad al v where obsr: "action_obs (?E (Suc n)) r = NormalAction (ReadMem ad al v)" by cases hence adal_r: "(ad, al) \ action_loc P (?E (Suc n)) r" by simp from J wfaSn \r \ ?C (Suc n)\ have obs_sim: "action_obs (?E (Suc n)) r \ action_obs E (?\ (Suc n) r)" "?\ (Suc n) r \ actions E" by(auto dest: wf_action_translation_on_actionD simp add: committed_subset_actions_def is_commit_sequence_def) with obsr have rE: "?\ (Suc n) r \ read_actions E" by(fastforce intro: read_actions.intros) from obs_sim obsr obtain v' where obsrE: "action_obs E (?\ (Suc n) r) = NormalAction (ReadMem ad al v')" by auto from wf_exec have "is_write_seen P E ws" by(rule wf_exec_is_write_seenD) from is_write_seenD[OF this rE obsrE] have "ws (?\ (Suc n) r) \ write_actions E" and "(ad, al) \ action_loc P E (ws (?\ (Suc n) r))" and nhb: "\ P,E \ ?\ (Suc n) r \hb ws (?\ (Suc n) r)" and vol: "is_volatile P al \ \ P,E \ ?\ (Suc n) r \so ws (?\ (Suc n) r)" by simp_all show ?case proof(cases "is_volatile P al") case False from wf_action_translation_on_actionD[OF wfaSn \?w \ ?C (Suc n)\] have "action_obs (?E (Suc n)) ?w \ action_obs E (?\ (Suc n) ?w)" by simp with w_eq have obs_sim_w: "action_obs (?E (Suc n)) ?w \ action_obs E (ws (?\ (Suc n) r))" by simp with \ws (?\ (Suc n) r) \ write_actions E\ \?w \ actions (?E (Suc n))\ have "?w \ write_actions (?E (Suc n))" by cases(fastforce intro: write_actions.intros is_write_action.intros elim!: is_write_action.cases) from \(ad, al) \ action_loc P E (ws (?\ (Suc n) r))\ obs_sim_w have "(ad, al) \ action_loc P (?E (Suc n)) ?w" by cases(auto intro: action_loc_aux_intros) with r adal_r \?w \ write_actions (?E (Suc n))\ False have "P,?E (Suc n) \ r \ ?w" by(auto simp add: non_volatile_conflict_def) with sc \r \ actions (?E (Suc n))\ w have "P,?E (Suc n) \ r \hb ?w \ P,?E (Suc n) \ ?w \hb r" by(rule correctly_synchronizedD[rule_format, OF sync \?E (Suc n) \ \\ wf]) moreover from J r \r \ ?C (Suc n)\ have "P,?E (Suc n) \ ?w \hb r \ P,E \ ws (?\ (Suc n) r) \hb ?\ (Suc n) r" and "\ P,?E (Suc n) \ r \hb ?w" by(simp_all add: happens_before_committed_weak_def) ultimately show ?thesis by auto next case True with rE obsrE have "?\ (Suc n) r \ sactions P E" by cases (auto intro: sactionsI) moreover from \ws (?\ (Suc n) r) \ write_actions E\ \(ad, al) \ action_loc P E (ws (?\ (Suc n) r))\ True have "ws (?\ (Suc n) r) \ sactions P E" by cases(auto intro!: sactionsI elim: is_write_action.cases) moreover have "?\ (Suc n) r \ ws (?\ (Suc n) r)" using \ws (?\ (Suc n) r) \ write_actions E\ rE by(auto dest: read_actions_not_write_actions) ultimately have "P,E \ ws (?\ (Suc n) r) \so ?\ (Suc n) r" using total_sync_order[of P E] vol[OF True] by(auto dest: total_onPD) moreover from \ws (?\ (Suc n) r) \ write_actions E\ \(ad, al) \ action_loc P E (ws (?\ (Suc n) r))\ True have "P \ (action_tid E (ws (?\ (Suc n) r)), action_obs E (ws (?\ (Suc n) r))) \sw (action_tid E (?\ (Suc n) r), action_obs E (?\ (Suc n) r))" by cases(fastforce elim!: is_write_action.cases intro: synchronizes_with.intros addr_locsI simp add: obsrE) ultimately have "P,E \ ws (?\ (Suc n) r) \sw ?\ (Suc n) r" by(rule sync_withI) thus ?thesis unfolding po_sw_def by blast qed qed thus "P,E \ ws r \hb r" unfolding r . qed corollary drf: "\ correctly_synchronized P \; legal_execution P \ (E, ws) \ \ sequentially_consistent P (E, ws)" by(erule drf_weak)(rule legal_imp_weakly_legal_execution) end end diff --git a/thys/JinjaThreads/MM/JMM_Framework.thy b/thys/JinjaThreads/MM/JMM_Framework.thy --- a/thys/JinjaThreads/MM/JMM_Framework.thy +++ b/thys/JinjaThreads/MM/JMM_Framework.thy @@ -1,3395 +1,3395 @@ (* Title: JinjaThreads/MM/JMM_Framework.thy Author: Andreas Lochbihler *) section \Combination of locales for heap operations and interleaving\ theory JMM_Framework imports JMM_Heap "../Framework/FWInitFinLift" "../Common/WellForm" begin lemma enat_plus_eq_enat_conv: \ \Move to Extended\_Nat\ "enat m + n = enat k \ k \ m \ n = enat (k - m)" by(cases n) auto declare convert_new_thread_action_id [simp] context heap begin lemma init_fin_lift_state_start_state: "init_fin_lift_state s (start_state f P C M vs) = start_state (\C M Ts T meth vs. (s, f C M Ts T meth vs)) P C M vs" by(simp add: start_state_def init_fin_lift_state_def split_beta fun_eq_iff) lemma non_speculative_start_heap_obs: "non_speculative P vs (llist_of (map snd (lift_start_obs start_tid start_heap_obs)))" apply(rule non_speculative_nthI) using start_heap_obs_not_Read by(clarsimp simp add: lift_start_obs_def lnth_LCons o_def eSuc_enat[symmetric] in_set_conv_nth split: nat.split_asm) lemma ta_seq_consist_start_heap_obs: "ta_seq_consist P Map.empty (llist_of (map snd (lift_start_obs start_tid start_heap_obs)))" using start_heap_obs_not_Read by(auto intro: ta_seq_consist_nthI simp add: lift_start_obs_def o_def lnth_LCons in_set_conv_nth split: nat.split_asm) end context allocated_heap begin lemma w_addrs_lift_start_heap_obs: "w_addrs (w_values P vs (map snd (lift_start_obs start_tid start_heap_obs))) \ w_addrs vs" by(simp add: lift_start_obs_def o_def w_addrs_start_heap_obs) end context heap begin lemma w_values_start_heap_obs_typeable: assumes wf: "wf_syscls P" and mrws: "v \ w_values P (\_. {}) (map snd (lift_start_obs start_tid start_heap_obs)) (ad, al)" shows "\T. P,start_heap \ ad@al : T \ P,start_heap \ v :\ T" proof - from in_w_valuesD[OF mrws] obtain obs' wa obs'' where eq: "map snd (lift_start_obs start_tid start_heap_obs) = obs' @ wa # obs''" and "is_write_action wa" and adal: "(ad, al) \ action_loc_aux P wa" and vwa: "value_written_aux P wa al = v" by blast from \is_write_action wa\ show ?thesis proof cases case (WriteMem ad' al' v') with vwa adal eq have "WriteMem ad al v \ set start_heap_obs" by(auto simp add: map_eq_append_conv Cons_eq_append_conv lift_start_obs_def) thus ?thesis by(rule start_heap_write_typeable) next case (NewHeapElem ad' hT) with vwa adal eq have "NewHeapElem ad hT \ set start_heap_obs" by(auto simp add: map_eq_append_conv Cons_eq_append_conv lift_start_obs_def) hence "typeof_addr start_heap ad = \hT\" by(rule NewHeapElem_start_heap_obsD[OF wf]) thus ?thesis using adal vwa NewHeapElem apply(cases hT) apply(auto intro!: addr_loc_type.intros dest: has_field_decl_above) apply(frule has_field_decl_above) apply(auto intro!: addr_loc_type.intros dest: has_field_decl_above) done qed qed lemma start_state_vs_conf: "wf_syscls P \ vs_conf P start_heap (w_values P (\_. {}) (map snd (lift_start_obs start_tid start_heap_obs)))" by(rule vs_confI)(rule w_values_start_heap_obs_typeable) end subsection \JMM traces for Jinja semantics\ context multithreaded_base begin inductive_set \ :: "('l,'t,'x,'m,'w) state \ ('t \ 'o) llist set" for \ :: "('l,'t,'x,'m,'w) state" where "mthr.Runs \ E' \ lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E') \ \ \" lemma actions_\E_aux: fixes \ E' defines "E == lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E')" assumes mthr: "mthr.Runs \ E'" and a: "enat a < llength E" obtains m n t ta where "lnth E a = (t, \ta\\<^bsub>o\<^esub> ! n)" and "n < length \ta\\<^bsub>o\<^esub>" and "enat m < llength E'" and "a = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + n" and "lnth E' m = (t, ta)" proof - from lnth_lconcat_conv[OF a[unfolded E_def], folded E_def] obtain m n where "lnth E a = lnth (lnth (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E') m) n" and "enat n < llength (lnth (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E') m)" and "enat m < llength (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E')" and "enat a = (\i(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E') i)) + enat n" by blast moreover obtain t ta where "lnth E' m = (t, ta)" by(cases "lnth E' m") ultimately have E_a: "lnth E a = (t, \ta\\<^bsub>o\<^esub> ! n)" and n: "n < length \ta\\<^bsub>o\<^esub>" and m: "enat m < llength E'" and a: "enat a = (\i(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E') i)) + enat n" - by(simp_all add: lnth_llist_of) + by(simp_all) note a also have "(\i(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E') i)) = sum (enat \ (\i. length \snd (lnth E' i)\\<^bsub>o\<^esub>)) {.. = enat (\isnd (lnth E' i)\\<^bsub>o\<^esub>)" - by(subst sum_hom)(simp_all add: zero_enat_def) + by(subst sum_comp_morphism)(simp_all add: zero_enat_def) finally have a: "a = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + n" by simp with E_a n m show thesis using \lnth E' m = (t, ta)\ by(rule that) qed lemma actions_\E: assumes E: "E \ \ \" and a: "enat a < llength E" obtains E' m n t ta where "E = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E')" and "mthr.Runs \ E'" and "lnth E a = (t, \ta\\<^bsub>o\<^esub> ! n)" and "n < length \ta\\<^bsub>o\<^esub>" and "enat m < llength E'" and "a = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + n" and "lnth E' m = (t, ta)" proof - from E obtain E' ws where E: "E = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E')" and "mthr.Runs \ E'" by(rule \.cases) blast from \mthr.Runs \ E'\ a[unfolded E] show ?thesis by(rule actions_\E_aux)(fold E, rule that[OF E \mthr.Runs \ E'\]) qed end context \multithreaded_wf begin text \Alternative characterisation for @{term "\"}\ lemma \_conv_Runs: "\ \ = lconcat ` lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ` llist_of_tllist ` {E. mthr.\Runs \ E}" (is "?lhs = ?rhs") proof(intro equalityI subsetI) fix E assume "E \ ?rhs" then obtain E' where E: "E = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (llist_of_tllist E'))" and \Runs: "mthr.\Runs \ E'" by(blast) obtain E'' where E': "E' = tmap (\(tls, s', tl, s''). tl) (case_sum (\(tls, s'). \s'\) Map.empty) E''" and \Runs': "mthr.\Runs_table2 \ E''" using \Runs by(rule mthr.\Runs_into_\Runs_table2) have "mthr.Runs \ (lconcat (lappend (lmap (\(tls, s, tl, s'). llist_of (tls @ [tl])) (llist_of_tllist E'')) (LCons (case terminal E'' of Inl (tls, s') \ llist_of tls | Inr tls \ tls) LNil)))" (is "mthr.Runs _ ?E'''") using \Runs' by(rule mthr.\Runs_table2_into_Runs) moreover let ?tail = "\E''. case terminal E'' of Inl (tls, s') \ llist_of tls | Inr tls \ tls" { have "E = lconcat (lfilter (\xs. \ lnull xs) (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (llist_of_tllist E')))" unfolding E by(simp add: lconcat_lfilter_neq_LNil) also have "\ = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (lmap (\(tls, s', tta, s''). tta) (lfilter (\(tls, s', (t, ta), s''). \ta\\<^bsub>o\<^esub> \ []) (llist_of_tllist E''))))" by(simp add: E' lfilter_lmap llist.map_comp o_def split_def) also from \mthr.\Runs_table2 \ E''\ have "lmap (\(tls, s', tta, s''). tta) (lfilter (\(tls, s', (t, ta), s''). \ta\\<^bsub>o\<^esub> \ []) (llist_of_tllist E'')) = lfilter (\(t, ta). \ta\\<^bsub>o\<^esub> \ []) (lconcat (lappend (lmap (\(tls, s, tl, s'). llist_of (tls @ [tl])) (llist_of_tllist E'')) (LCons (?tail E'') LNil)))" (is "?lhs \ E'' = ?rhs \ E''") proof(coinduction arbitrary: \ E'' rule: llist.coinduct_strong) case (Eq_llist \ E'') have ?lnull by(cases "lfinite (llist_of_tllist E'')")(fastforce split: sum.split_asm simp add: split_beta lset_lconcat_lfinite lappend_inf mthr.silent_move2_def dest: mthr.\Runs_table2_silentsD[OF Eq_llist] mthr.\Runs_table2_terminal_silentsD[OF Eq_llist] mthr.\Runs_table2_terminal_inf_stepD[OF Eq_llist] m\move_silentD inf_step_silentD silent_moves2_silentD split: sum.split_asm)+ moreover have ?LCons proof(intro impI conjI) assume lhs': "\ lnull (lmap (\(tls, s', tta, s''). tta) (lfilter (\(tls, s', (t, ta), s''). \ta\\<^bsub>o\<^esub> \ []) (llist_of_tllist E'')))" (is "\ lnull ?lhs'") and "\ lnull (lfilter (\(t, ta). \ta\\<^bsub>o\<^esub> \ []) (lconcat (lappend (lmap (\(tls, s, tl, s'). llist_of (tls @ [tl])) (llist_of_tllist E'')) (LCons (case terminal E'' of Inl (tls, s') \ llist_of tls | Inr tls \ tls) LNil))))" (is "\ lnull ?rhs'") note \Runs' = \mthr.\Runs_table2 \ E''\ from lhs' obtain tl tls' where "?lhs \ E'' = LCons tl tls'" by(auto simp only: not_lnull_conv) then obtain tls s' s'' tlsstlss' where tls': "tls' = lmap (\(tls, s', tta, s''). tta) tlsstlss'" and filter: "lfilter (\(tls, s', (t, ta), s''). obs_a ta \ []) (llist_of_tllist E'') = LCons (tls, s', tl, s'') tlsstlss'" using lhs' by(fastforce simp add: lmap_eq_LCons_conv) from lfilter_eq_LConsD[OF filter] obtain us vs where eq: "llist_of_tllist E'' = lappend us (LCons (tls, s', tl, s'') vs)" and fin: "lfinite us" and empty: "\(tls, s', (t, ta), s'')\lset us. obs_a ta = []" and neq_empty: "obs_a (snd tl) \ []" and tlsstlss': "tlsstlss' = lfilter (\(tls, s', (t, ta), s''). obs_a ta \ []) vs" by(auto simp add: split_beta) from eq obtain E''' where E'': "E'' = lappendt us E'''" and eq': "llist_of_tllist E''' = LCons (tls, s', tl, s'') vs" and terminal: "terminal E''' = terminal E''" unfolding llist_of_tllist_eq_lappend_conv by auto from \Runs' fin E'' obtain \' where \Runs'': "mthr.\Runs_table2 \' E'''" by(auto dest: mthr.\Runs_table2_lappendtD) then obtain \'' E'''' where "mthr.\Runs_table2 \'' E''''" "E''' = TCons (tls, s', tl, s'') E''''" using eq' by cases auto moreover from \Runs' E'' fin have "\(tls, s, tl, s')\lset us. \(t, ta)\set tls. ta = \" by(fastforce dest: mthr.\Runs_table2_silentsD m\move_silentD simp add: mthr.silent_move2_def) hence "lfilter (\(t, ta). obs_a ta \ []) (lconcat (lmap (\(tls, s, tl, s'). llist_of (tls @ [tl])) us)) = LNil" using empty by(auto simp add: lfilter_empty_conv lset_lconcat_lfinite split_beta) moreover from \Runs'' eq' have "snd ` set tls \ {\}" by(cases)(fastforce dest: silent_moves2_silentD)+ hence "[(t, ta)\tls . obs_a ta \ []] = []" by(auto simp add: filter_empty_conv split_beta) ultimately show "lhd ?lhs' = lhd ?rhs'" and "(\\ E''. ltl ?lhs' = lmap (\(tls, s', tta, s''). tta) (lfilter (\(tls, s', (t, ta), s''). \ta\\<^bsub>o\<^esub> \ []) (llist_of_tllist E'')) \ ltl ?rhs' = lfilter (\(t, ta). \ta\\<^bsub>o\<^esub> \ []) (lconcat (lappend (lmap (\(tls, s, tl, s'). llist_of (tls @ [tl])) (llist_of_tllist E'')) (LCons (case terminal E'' of Inl (tls, s') \ llist_of tls | Inr tls \ tls) LNil))) \ \trsys.\Runs_table2 redT m\move \ E'') \ ltl ?lhs' = ltl ?rhs'" using lhs' E'' fin tls' tlsstlss' filter eq' neq_empty by(auto simp add: lmap_lappend_distrib lappend_assoc split_beta filter_empty_conv simp del: split_paired_Ex) qed ultimately show ?case .. qed also have "lmap (\(t, ta). llist_of (map (Pair t) (obs_a ta))) \ = lfilter (\obs. \ lnull obs) (lmap (\(t, ta). llist_of (map (Pair t) (obs_a ta))) (lconcat (lappend (lmap (\(tls, s, tl, s'). llist_of (tls @ [tl])) (llist_of_tllist E'')) (LCons (?tail E'') LNil))))" unfolding lfilter_lmap by(simp add: o_def split_def llist_of_eq_LNil_conv) finally have "E = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?E''')" by(simp add: lconcat_lfilter_neq_LNil) } ultimately show "E \ ?lhs" by(blast intro: \.intros) next fix E assume "E \ ?lhs" then obtain E' where E: "E = lconcat (lmap (\(t, ta). llist_of (map (Pair t) (obs_a ta))) E')" and Runs: "mthr.Runs \ E'" by(blast elim: \.cases) from Runs obtain E'' where E': "E' = lmap (\(s, tl, s'). tl) E''" and Runs': "mthr.Runs_table \ E''" by(rule mthr.Runs_into_Runs_table) have "mthr.\Runs \ (tmap (\(s, tl, s'). tl) id (tfilter None (\(s, tl, s'). \ m\move s tl s') (tllist_of_llist (Some (llast (LCons \ (lmap (\(s, tl, s'). s') E'')))) E'')))" (is "mthr.\Runs _ ?E'''") using Runs' by(rule mthr.Runs_table_into_\Runs) moreover have "(\(s, (t, ta), s'). obs_a ta \ []) = (\(s, (t, ta), s'). obs_a ta \ [] \ \ m\move s (t, ta) s')" by(rule ext)(auto dest: m\move_silentD) hence "E = lconcat (lmap (\(t, ta). llist_of (map (Pair t) (obs_a ta))) (llist_of_tllist ?E'''))" unfolding E E' by(subst (1 2) lconcat_lfilter_neq_LNil[symmetric])(simp add: lfilter_lmap lfilter_lfilter o_def split_def) ultimately show "E \ ?rhs" by(blast) qed end text \Running threads have been started before\ definition Status_no_wait_locks :: "('l,'t,status \ 'x) thread_info \ bool" where "Status_no_wait_locks ts \ (\t status x ln. ts t = \((status, x), ln)\ \ status \ Running \ ln = no_wait_locks)" lemma Status_no_wait_locks_PreStartD: "\ln. \ Status_no_wait_locks ts; ts t = \((PreStart, x), ln)\ \ \ ln = no_wait_locks" unfolding Status_no_wait_locks_def by blast lemma Status_no_wait_locks_FinishedD: "\ln. \ Status_no_wait_locks ts; ts t = \((Finished, x), ln)\ \ \ ln = no_wait_locks" unfolding Status_no_wait_locks_def by blast lemma Status_no_wait_locksI: "(\t status x ln. \ ts t = \((status, x), ln)\; status = PreStart \ status = Finished \ \ ln = no_wait_locks) \ Status_no_wait_locks ts" unfolding Status_no_wait_locks_def apply clarify apply(case_tac status) apply auto done context heap_base begin lemma Status_no_wait_locks_start_state: "Status_no_wait_locks (thr (init_fin_lift_state status (start_state f P C M vs)))" by(clarsimp simp add: Status_no_wait_locks_def init_fin_lift_state_def start_state_def split_beta) end context multithreaded_base begin lemma init_fin_preserve_Status_no_wait_locks: assumes ok: "Status_no_wait_locks (thr s)" and redT: "multithreaded_base.redT init_fin_final init_fin (map NormalAction \ convert_RA) s tta s'" shows "Status_no_wait_locks (thr s')" using redT proof(cases rule: multithreaded_base.redT.cases[consumes 1, case_names redT_normal redT_acquire]) case redT_acquire with ok show ?thesis by(auto intro!: Status_no_wait_locksI dest: Status_no_wait_locks_PreStartD Status_no_wait_locks_FinishedD split: if_split_asm) next case redT_normal show ?thesis proof(rule Status_no_wait_locksI) fix t' status' x' ln' assume tst': "thr s' t' = \((status', x'), ln')\" and status: "status' = PreStart \ status' = Finished" show "ln' = no_wait_locks" proof(cases "thr s t'") case None with redT_normal tst' show ?thesis by(fastforce elim!: init_fin.cases dest: redT_updTs_new_thread simp add: final_thread.actions_ok_iff split: if_split_asm) next case (Some sxln) obtain status'' x'' ln'' where [simp]: "sxln = ((status'', x''), ln'')" by(cases sxln) auto show ?thesis proof(cases "fst tta = t'") case True with redT_normal tst' status show ?thesis by(auto simp add: expand_finfun_eq fun_eq_iff) next case False with tst' redT_normal Some status have "status'' = status'" "ln'' = ln'" by(force dest: redT_updTs_Some simp add: final_thread.actions_ok_iff)+ with ok Some status show ?thesis by(auto dest: Status_no_wait_locks_PreStartD Status_no_wait_locks_FinishedD) qed qed qed qed lemma init_fin_Running_InitialThreadAction: assumes redT: "multithreaded_base.redT init_fin_final init_fin (map NormalAction \ convert_RA) s tta s'" and not_running: "\x ln. thr s t \ \((Running, x), ln)\" and running: "thr s' t = \((Running, x'), ln')\" shows "tta = (t, \InitialThreadAction\)" using redT proof(cases rule: multithreaded_base.redT.cases[consumes 1, case_names redT_normal redT_acquire]) case redT_acquire with running not_running show ?thesis by(auto split: if_split_asm) next case redT_normal show ?thesis proof(cases "thr s t") case None with redT_normal running not_running show ?thesis by(fastforce simp add: final_thread.actions_ok_iff elim: init_fin.cases dest: redT_updTs_new_thread split: if_split_asm) next case (Some a) with redT_normal running not_running show ?thesis apply(cases a) apply(auto simp add: final_thread.actions_ok_iff split: if_split_asm elim: init_fin.cases) apply((drule (1) redT_updTs_Some)?, fastforce)+ done qed qed end context if_multithreaded begin lemma init_fin_Trsys_preserve_Status_no_wait_locks: assumes ok: "Status_no_wait_locks (thr s)" and Trsys: "if.mthr.Trsys s ttas s'" shows "Status_no_wait_locks (thr s')" using Trsys ok by(induct)(blast dest: init_fin_preserve_Status_no_wait_locks)+ lemma init_fin_Trsys_Running_InitialThreadAction: assumes redT: "if.mthr.Trsys s ttas s'" and not_running: "\x ln. thr s t \ \((Running, x), ln)\" and running: "thr s' t = \((Running, x'), ln')\" shows "(t, \InitialThreadAction\) \ set ttas" using redT not_running running proof(induct arbitrary: x' ln') case rtrancl3p_refl thus ?case by(fastforce) next case (rtrancl3p_step s ttas s' tta s'') thus ?case by(cases "\x ln. thr s' t = \((Running, x), ln)\")(fastforce dest: init_fin_Running_InitialThreadAction)+ qed end locale heap_multithreaded_base = heap_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write + mthr: multithreaded_base final r convert_RA for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and final :: "'x \ bool" and r :: "('addr, 'thread_id, 'x, 'heap, 'addr, ('addr, 'thread_id) obs_event) semantics" ("_ \ _ -_\ _" [50,0,0,50] 80) and convert_RA :: "'addr released_locks \ ('addr, 'thread_id) obs_event list" sublocale heap_multithreaded_base < mthr: if_multithreaded_base final r convert_RA . context heap_multithreaded_base begin abbreviation \_start :: "(cname \ mname \ ty list \ ty \ 'md \ 'addr val list \ 'x) \ 'md prog \ cname \ mname \ 'addr val list \ status \ ('thread_id \ ('addr, 'thread_id) obs_event action) llist set" where "\_start f P C M vs status \ lappend (llist_of (lift_start_obs start_tid start_heap_obs)) ` mthr.if.\ (init_fin_lift_state status (start_state f P C M vs))" end locale heap_multithreaded = heap_multithreaded_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write final r convert_RA + heap addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write P + mthr: multithreaded final r convert_RA for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and final :: "'x \ bool" and r :: "('addr, 'thread_id, 'x, 'heap, 'addr, ('addr, 'thread_id) obs_event) semantics" ("_ \ _ -_\ _" [50,0,0,50] 80) and convert_RA :: "'addr released_locks \ ('addr, 'thread_id) obs_event list" and P :: "'md prog" sublocale heap_multithreaded < mthr: if_multithreaded final r convert_RA by(unfold_locales) sublocale heap_multithreaded < "if": jmm_multithreaded mthr.init_fin_final mthr.init_fin "map NormalAction \ convert_RA" P . context heap_multithreaded begin lemma thread_start_actions_ok_init_fin_RedT: assumes Red: "mthr.if.RedT (init_fin_lift_state status (start_state f P C M vs)) ttas s'" (is "mthr.if.RedT ?start_state _ _") shows "thread_start_actions_ok (llist_of (lift_start_obs start_tid start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) ttas)))" (is "thread_start_actions_ok (llist_of (?obs_prefix @ ?E'))") proof(rule thread_start_actions_okI) let ?E = "llist_of (?obs_prefix @ ?E')" fix a assume a: "a \ actions ?E" and new: "\ is_new_action (action_obs ?E a)" show "\i \ a. action_obs ?E i = InitialThreadAction \ action_tid ?E i = action_tid ?E a" proof(cases "action_tid ?E a = start_tid") case True thus ?thesis by(auto simp add: lift_start_obs_def action_tid_def action_obs_def) next case False let ?a = "a - length ?obs_prefix" from False have a_len: "a \ length ?obs_prefix" by(rule contrapos_np)(auto simp add: lift_start_obs_def action_tid_def lnth_LCons nth_append split: nat.split) hence [simp]: "action_tid ?E a = action_tid (llist_of ?E') ?a" "action_obs ?E a = action_obs (llist_of ?E') ?a" by(simp_all add: action_tid_def nth_append action_obs_def) from False have not_running: "\x ln. thr ?start_state (action_tid (llist_of ?E') ?a) \ \((Running, x), ln)\" by(auto simp add: start_state_def split_beta init_fin_lift_state_def split: if_split_asm) from a a_len have "?a < length ?E'" by(simp add: actions_def) from nth_concat_conv[OF this] obtain m n where E'_a: "?E' ! ?a = (\(t, ta). (t, \ta\\<^bsub>o\<^esub> ! n)) (ttas ! m)" and n: "n < length \snd (ttas ! m)\\<^bsub>o\<^esub>" and m: "m < length ttas" and a_conv: "?a = (\i(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) ttas ! i)) + n" by(clarsimp simp add: split_def) from Red obtain s'' s''' where Red1: "mthr.if.RedT ?start_state (take m ttas) s''" and red: "mthr.if.redT s'' (ttas ! m) s'''" and Red2: "mthr.if.RedT s''' (drop (Suc m) ttas) s'" unfolding mthr.if.RedT_def by(subst (asm) (4) id_take_nth_drop[OF m])(blast elim: rtrancl3p_appendE rtrancl3p_converseE) from E'_a m n have [simp]: "action_tid (llist_of ?E') ?a = fst (ttas ! m)" by(simp add: action_tid_def split_def) from red obtain status x ln where tst: "thr s'' (fst (ttas ! m)) = \((status, x), ln)\" by cases auto show ?thesis proof(cases "status = PreStart \ status = Finished") case True from Red1 have "Status_no_wait_locks (thr s'')" unfolding mthr.if.RedT_def by(rule mthr.init_fin_Trsys_preserve_Status_no_wait_locks[OF Status_no_wait_locks_start_state]) with True tst have "ln = no_wait_locks" by(auto dest: Status_no_wait_locks_PreStartD Status_no_wait_locks_FinishedD) with red tst True have "\snd (ttas ! m)\\<^bsub>o\<^esub> = [InitialThreadAction]" by(cases) auto hence "action_obs ?E a = InitialThreadAction" using a_conv n a_len E'_a by(simp add: action_obs_def nth_append split_beta) thus ?thesis by(auto) next case False hence "status = Running" by(cases status) auto with tst mthr.init_fin_Trsys_Running_InitialThreadAction[OF Red1[unfolded mthr.if.RedT_def] not_running] have "(fst (ttas ! m), \InitialThreadAction\) \ set (take m ttas)" using E'_a by(auto simp add: action_tid_def split_beta) then obtain i where i: "i < m" and nth_i: "ttas ! i = (fst (ttas ! m), \InitialThreadAction\)" unfolding in_set_conv_nth by auto let ?i' = "length (concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (take i ttas)))" let ?i = "length ?obs_prefix + ?i'" from i m nth_i have "?i' < length (concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (take m ttas)))" apply(simp add: length_concat o_def split_beta) apply(subst (6) id_take_nth_drop[where i=i]) apply(simp_all add: take_map[symmetric] min_def) done also from m have "\ \ ?a" unfolding a_conv by(simp add: length_concat sum_list_sum_nth min_def split_def atLeast0LessThan) finally have "?i < a" using a_len by simp moreover from i m nth_i have "?i' < length ?E'" apply(simp add: length_concat o_def split_def) apply(subst (7) id_take_nth_drop[where i=i]) apply(simp_all add: take_map[symmetric]) done from nth_i i E'_a a_conv m have "lnth ?E ?i = (fst (ttas ! m), InitialThreadAction)" by(simp add: lift_start_obs_def nth_append length_concat o_def split_def)(rule nth_concat_eqI[where k=0 and i=i], simp_all add: take_map o_def split_def) ultimately show ?thesis using E'_a by(cases "ttas ! m")(auto simp add: action_obs_def action_tid_def nth_append intro!: exI[where x="?i"]) qed qed qed (* TODO: use previous lemma for proof *) lemma thread_start_actions_ok_init_fin: assumes E: "E \ mthr.if.\ (init_fin_lift_state status (start_state f P C M vs))" shows "thread_start_actions_ok (lappend (llist_of (lift_start_obs start_tid start_heap_obs)) E)" (is "thread_start_actions_ok ?E") proof(rule thread_start_actions_okI) let ?start_heap_obs = "lift_start_obs start_tid start_heap_obs" let ?start_state = "init_fin_lift_state status (start_state f P C M vs)" fix a assume a: "a \ actions ?E" and a_new: "\ is_new_action (action_obs ?E a)" show "\i. i \ a \ action_obs ?E i = InitialThreadAction \ action_tid ?E i = action_tid ?E a" proof(cases "action_tid ?E a = start_tid") case True thus ?thesis by(auto simp add: lift_start_obs_def action_tid_def action_obs_def) next case False let ?a = "a - length ?start_heap_obs" from False have "a \ length ?start_heap_obs" by(rule contrapos_np)(auto simp add: lift_start_obs_def action_tid_def lnth_LCons lnth_lappend1 split: nat.split) hence [simp]: "action_tid ?E a = action_tid E ?a" "action_obs ?E a = action_obs E ?a" by(simp_all add: action_tid_def lnth_lappend2 action_obs_def) from False have not_running: "\x ln. thr ?start_state (action_tid E ?a) \ \((Running, x), ln)\" by(auto simp add: start_state_def split_beta init_fin_lift_state_def split: if_split_asm) from E obtain E' where E': "E = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E')" and \Runs: "mthr.if.mthr.Runs ?start_state E'" by(rule mthr.if.\.cases) from a E' \a \ length ?start_heap_obs\ have enat_a: "enat ?a < llength (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E'))" by(cases "llength (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E'))")(auto simp add: actions_def) with \Runs obtain m n t ta where a_obs: "lnth (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E')) (a - length ?start_heap_obs) = (t, \ta\\<^bsub>o\<^esub> ! n)" and n: "n < length \ta\\<^bsub>o\<^esub>" and m: "enat m < llength E'" and a_conv: "?a = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + n" and E'_m: "lnth E' m = (t, ta)" by(rule mthr.if.actions_\E_aux) from a_obs have [simp]: "action_tid E ?a = t" "action_obs E ?a = \ta\\<^bsub>o\<^esub> ! n" by(simp_all add: E' action_tid_def action_obs_def) let ?E' = "ldropn (Suc m) E'" let ?m_E' = "ltake (enat m) E'" have E'_unfold: "E' = lappend (ltake (enat m) E') (LCons (lnth E' m) ?E')" unfolding ldropn_Suc_conv_ldropn[OF m] by simp hence "mthr.if.mthr.Runs ?start_state (lappend ?m_E' (LCons (lnth E' m) ?E'))" using \Runs by simp then obtain \' where \_\': "mthr.if.mthr.Trsys ?start_state (list_of ?m_E') \'" and \Runs': "mthr.if.mthr.Runs \' (LCons (lnth E' m) ?E')" by(rule mthr.if.mthr.Runs_lappendE) simp from \Runs' obtain \''' where red_a: "mthr.if.redT \' (t, ta) \'''" and \Runs'': "mthr.if.mthr.Runs \''' ?E'" unfolding E'_m by cases from red_a obtain status x ln where tst: "thr \' t = \((status, x), ln)\" by cases auto show ?thesis proof(cases "status = PreStart \ status = Finished") case True have "Status_no_wait_locks (thr \')" by(rule mthr.init_fin_Trsys_preserve_Status_no_wait_locks[OF _ \_\'])(rule Status_no_wait_locks_start_state) with True tst have "ln = no_wait_locks" by(auto dest: Status_no_wait_locks_PreStartD Status_no_wait_locks_FinishedD) with red_a tst True have "\ta\\<^bsub>o\<^esub> = [InitialThreadAction]" by(cases) auto hence "action_obs E ?a = InitialThreadAction" using a_obs n unfolding E' by(simp add: action_obs_def) thus ?thesis by(auto) next case False hence "status = Running" by(cases status) auto with tst mthr.init_fin_Trsys_Running_InitialThreadAction[OF \_\' not_running] have "(action_tid E ?a, \InitialThreadAction\) \ set (list_of (ltake (enat m) E'))" using a_obs E' by(auto simp add: action_tid_def) then obtain i where "i < m" "enat i < llength E'" and nth_i: "lnth E' i = (action_tid E ?a, \InitialThreadAction\)" unfolding in_set_conv_nth by(cases "llength E'")(auto simp add: length_list_of_conv_the_enat lnth_ltake) let ?i' = "\isnd (lnth E' i)\\<^bsub>o\<^esub>" let ?i = "length ?start_heap_obs + ?i'" from \i < m\ have "(\isnd (lnth E' i)\\<^bsub>o\<^esub>) = ?i' + (\i=i..snd (lnth E' i)\\<^bsub>o\<^esub>)" unfolding atLeast0LessThan[symmetric] by(subst sum.atLeastLessThan_concat) simp_all hence "?i' \ ?a" unfolding a_conv by simp hence "?i \ a" using \a \ length ?start_heap_obs\ by arith from \?i' \ ?a\ have "enat ?i' < llength E" using enat_a E' by(simp add: le_less_trans[where y="enat ?a"]) from lnth_lconcat_conv[OF this[unfolded E'], folded E'] obtain k l where nth_i': "lnth E ?i' = lnth (lnth (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E') k) l" and l: "l < length \snd (lnth E' k)\\<^bsub>o\<^esub>" and k: "enat k < llength E'" and i_conv: "enat ?i' = (\i(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E') i)) + enat l" by(fastforce simp add: split_beta) have "(\i(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E') i)) = (\i (\i. length \snd (lnth E' i)\\<^bsub>o\<^esub>)) i)" by(rule sum.cong)(simp_all add: less_trans[where y="enat k"] split_beta k) also have "\ = enat (\isnd (lnth E' i)\\<^bsub>o\<^esub>)" - by(rule sum_hom)(simp_all add: zero_enat_def) + by(rule sum_comp_morphism)(simp_all add: zero_enat_def) finally have i_conv: "?i' = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + l" using i_conv by simp have [simp]: "i = k" proof(rule ccontr) assume "i \ k" thus False unfolding neq_iff proof assume "i < k" hence "(\isnd (lnth E' i)\\<^bsub>o\<^esub>) = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + (\i=i..snd (lnth E' i)\\<^bsub>o\<^esub>)" unfolding atLeast0LessThan[symmetric] by(subst sum.atLeastLessThan_concat) simp_all with i_conv have "(\i=i..snd (lnth E' i)\\<^bsub>o\<^esub>) = l" "l = 0" by simp_all moreover have "(\i=i..snd (lnth E' i)\\<^bsub>o\<^esub>) \ length \snd (lnth E' i)\\<^bsub>o\<^esub>" by(subst sum.atLeast_Suc_lessThan[OF \i < k\]) simp ultimately show False using nth_i by simp next assume "k < i" hence "?i' = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + (\i=k..snd (lnth E' i)\\<^bsub>o\<^esub>)" unfolding atLeast0LessThan[symmetric] by(subst sum.atLeastLessThan_concat) simp_all with i_conv have "(\i=k..snd (lnth E' i)\\<^bsub>o\<^esub>) = l" by simp moreover have "(\i=k..snd (lnth E' i)\\<^bsub>o\<^esub>) \ length \snd (lnth E' k)\\<^bsub>o\<^esub>" by(subst sum.atLeast_Suc_lessThan[OF \k < i\]) simp ultimately show False using l by simp qed qed with l nth_i have [simp]: "l = 0" by simp hence "lnth E ?i' = (action_tid E ?a, InitialThreadAction)" using nth_i nth_i' k by simp with \?i \ a\ show ?thesis by(auto simp add: action_tid_def action_obs_def lnth_lappend2) qed qed qed end text \In the subsequent locales, \convert_RA\ refers to @{term "convert_RA"} and is no longer a parameter!\ lemma convert_RA_not_write: "\ln. ob \ set (convert_RA ln) \ \ is_write_action (NormalAction ob)" by(auto simp add: convert_RA_def) lemma ta_seq_consist_convert_RA: fixes ln shows "ta_seq_consist P vs (llist_of ((map NormalAction \ convert_RA) ln))" proof(rule ta_seq_consist_nthI) fix i ad al v assume "enat i < llength (llist_of ((map NormalAction \ convert_RA) ln :: ('b, 'c) obs_event action list))" and "lnth (llist_of ((map NormalAction \ convert_RA) ln :: ('b, 'c) obs_event action list)) i = NormalAction (ReadMem ad al v)" hence "ReadMem ad al v \ set (convert_RA ln :: ('b, 'c) obs_event list)" by(auto simp add: in_set_conv_nth) hence False by(auto simp add: convert_RA_def) thus "\b. mrw_values P vs (list_of (ltake (enat i) (llist_of ((map NormalAction \ convert_RA) ln)))) (ad, al) = \(v, b)\" .. qed lemma ta_hb_consistent_convert_RA: "\ln. ta_hb_consistent P E (llist_of (map (Pair t) ((map NormalAction \ convert_RA) ln)))" by(rule ta_hb_consistent_not_ReadI)(auto simp add: convert_RA_def) locale allocated_multithreaded = allocated_heap addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write allocated P + mthr: multithreaded final r convert_RA for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and allocated :: "'heap \ 'addr set" and final :: "'x \ bool" and r :: "('addr, 'thread_id, 'x, 'heap, 'addr, ('addr, 'thread_id) obs_event) semantics" ("_ \ _ -_\ _" [50,0,0,50] 80) and P :: "'md prog" + assumes red_allocated_mono: "t \ (x, m) -ta\ (x', m') \ allocated m \ allocated m'" and red_New_allocatedD: "\ t \ (x, m) -ta\ (x', m'); NewHeapElem ad CTn \ set \ta\\<^bsub>o\<^esub> \ \ ad \ allocated m' \ ad \ allocated m" and red_allocated_NewD: "\ t \ (x, m) -ta\ (x', m'); ad \ allocated m'; ad \ allocated m \ \ \CTn. NewHeapElem ad CTn \ set \ta\\<^bsub>o\<^esub>" and red_New_same_addr_same: "\ t \ (x, m) -ta\ (x', m'); \ta\\<^bsub>o\<^esub> ! i = NewHeapElem a CTn; i < length \ta\\<^bsub>o\<^esub>; \ta\\<^bsub>o\<^esub> ! j = NewHeapElem a CTn'; j < length \ta\\<^bsub>o\<^esub> \ \ i = j" sublocale allocated_multithreaded < heap_multithreaded addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write final r convert_RA P by(unfold_locales) context allocated_multithreaded begin lemma redT_allocated_mono: assumes "mthr.redT \ (t, ta) \'" shows "allocated (shr \) \ allocated (shr \')" using assms by cases(auto dest: red_allocated_mono del: subsetI) lemma RedT_allocated_mono: assumes "mthr.RedT \ ttas \'" shows "allocated (shr \) \ allocated (shr \')" using assms unfolding mthr.RedT_def by induct(auto dest!: redT_allocated_mono intro: subset_trans del: subsetI) lemma init_fin_allocated_mono: "t \ (x, m) -ta\i (x', m') \ allocated m \ allocated m'" by(cases rule: mthr.init_fin.cases)(auto dest: red_allocated_mono) lemma init_fin_redT_allocated_mono: assumes "mthr.if.redT \ (t, ta) \'" shows "allocated (shr \) \ allocated (shr \')" using assms by cases(auto dest: init_fin_allocated_mono del: subsetI) lemma init_fin_RedT_allocated_mono: assumes "mthr.if.RedT \ ttas \'" shows "allocated (shr \) \ allocated (shr \')" using assms unfolding mthr.if.RedT_def by induct(auto dest!: init_fin_redT_allocated_mono intro: subset_trans del: subsetI) lemma init_fin_red_New_allocatedD: assumes "t \ (x, m) -ta\i (x', m')" "NormalAction (NewHeapElem ad CTn) \ set \ta\\<^bsub>o\<^esub>" shows "ad \ allocated m' \ ad \ allocated m" using assms by cases(auto dest: red_New_allocatedD) lemma init_fin_red_allocated_NewD: assumes "t \ (x, m) -ta\i (x', m')" "ad \ allocated m'" "ad \ allocated m" shows "\CTn. NormalAction (NewHeapElem ad CTn) \ set \ta\\<^bsub>o\<^esub>" using assms by(cases)(auto dest!: red_allocated_NewD) lemma init_fin_red_New_same_addr_same: assumes "t \ (x, m) -ta\i (x', m')" and "\ta\\<^bsub>o\<^esub> ! i = NormalAction (NewHeapElem a CTn)" "i < length \ta\\<^bsub>o\<^esub>" and "\ta\\<^bsub>o\<^esub> ! j = NormalAction (NewHeapElem a CTn')" "j < length \ta\\<^bsub>o\<^esub>" shows "i = j" using assms by cases(auto dest: red_New_same_addr_same) lemma init_fin_redT_allocated_NewHeapElemD: assumes "mthr.if.redT s (t, ta) s'" and "ad \ allocated (shr s')" and "ad \ allocated (shr s)" shows "\CTn. NormalAction (NewHeapElem ad CTn) \ set \ta\\<^bsub>o\<^esub>" using assms by(cases)(auto dest: init_fin_red_allocated_NewD) lemma init_fin_RedT_allocated_NewHeapElemD: assumes "mthr.if.RedT s ttas s'" and "ad \ allocated (shr s')" and "ad \ allocated (shr s)" shows "\t ta CTn. (t, ta) \ set ttas \ NormalAction (NewHeapElem ad CTn) \ set \ta\\<^bsub>o\<^esub>" using assms proof(induct rule: mthr.if.RedT_induct') case refl thus ?case by simp next case (step ttas s' t ta s'') thus ?case by(cases "ad \ allocated (shr s')")(fastforce simp del: split_paired_Ex dest: init_fin_redT_allocated_NewHeapElemD)+ qed lemma \_new_actions_for_unique: assumes E: "E \ \_start f P C M vs status" and a: "a \ new_actions_for P E adal" and a': "a' \ new_actions_for P E adal" shows "a = a'" using a a' proof(induct a a' rule: wlog_linorder_le) case symmetry thus ?case by simp next case (le a a') note a = \a \ new_actions_for P E adal\ and a' = \a' \ new_actions_for P E adal\ and a_a' = \a \ a'\ obtain ad al where adal: "adal = (ad, al)" by(cases adal) let ?init_obs = "lift_start_obs start_tid start_heap_obs" let ?start_state = "init_fin_lift_state status (start_state f P C M vs)" have distinct: "distinct (filter (\obs. \a CTn. obs = NormalAction (NewHeapElem a CTn)) (map snd ?init_obs))" unfolding start_heap_obs_def by(fastforce intro: inj_onI intro!: distinct_filter simp add: distinct_map distinct_zipI1 distinct_initialization_list) from start_addrs_allocated have dom_start_state: "{a. \CTn. NormalAction (NewHeapElem a CTn) \ snd ` set ?init_obs} \ allocated (shr ?start_state)" by(fastforce simp add: init_fin_lift_state_conv_simps shr_start_state dest: NewHeapElem_start_heap_obs_start_addrsD subsetD) show ?case proof(cases "a' < length ?init_obs") case True with a' adal E obtain t_a' CTn_a' where CTn_a': "?init_obs ! a' = (t_a', NormalAction (NewHeapElem ad CTn_a'))" by(cases "?init_obs ! a'")(fastforce elim!: is_new_action.cases action_loc_aux_cases simp add: action_obs_def lnth_lappend1 new_actions_for_def )+ from True a_a' have len_a: "a < length ?init_obs" by simp with a adal E obtain t_a CTn_a where CTn_a: "?init_obs ! a = (t_a, NormalAction (NewHeapElem ad CTn_a))" by(cases "?init_obs ! a")(fastforce elim!: is_new_action.cases action_loc_aux_cases simp add: action_obs_def lnth_lappend1 new_actions_for_def )+ from CTn_a CTn_a' True len_a have "NormalAction (NewHeapElem ad CTn_a') \ snd ` set ?init_obs" and "NormalAction (NewHeapElem ad CTn_a) \ snd ` set ?init_obs" unfolding set_conv_nth by(fastforce intro: rev_image_eqI)+ hence [simp]: "CTn_a' = CTn_a" using distinct_start_addrs' by(auto simp add: in_set_conv_nth distinct_conv_nth start_heap_obs_def start_addrs_def) blast from distinct_filterD[OF distinct, of a' a "NormalAction (NewHeapElem ad CTn_a)"] len_a True CTn_a CTn_a' show "a = a'" by simp next case False obtain n where n: "length ?init_obs = n" by blast with False have "n \ a'" by simp from E obtain E'' where E: "E = lappend (llist_of ?init_obs) E''" and E'': "E'' \ mthr.if.\ ?start_state" by auto from E'' obtain E' where E': "E'' = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E')" and \Runs: "mthr.if.mthr.Runs ?start_state E'" by(rule mthr.if.\.cases) from E E'' a' n \n \ a'\ adal have a': "a' - n \ new_actions_for P E'' adal" by(auto simp add: new_actions_for_def lnth_lappend2 action_obs_def actions_lappend elim: actionsE) from a' have "a' - n \ actions E''" by(auto elim: new_actionsE) hence "enat (a' - n) < llength E''" by(rule actionsE) with \Runs obtain a'_m a'_n t_a' ta_a' where E_a': "lnth E'' (a' - n) = (t_a', \ta_a'\\<^bsub>o\<^esub> ! a'_n)" and a'_n: "a'_n < length \ta_a'\\<^bsub>o\<^esub>" and a'_m: "enat a'_m < llength E'" and a'_conv: "a' - n = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + a'_n" and E'_a'_m: "lnth E' a'_m = (t_a', ta_a')" unfolding E' by(rule mthr.if.actions_\E_aux) from a' have "is_new_action (action_obs E'' (a' - n))" and "(ad, al) \ action_loc P E'' (a' - n)" unfolding adal by(auto elim: new_actionsE) then obtain CTn' where "action_obs E'' (a' - n) = NormalAction (NewHeapElem ad CTn')" by cases(fastforce)+ hence New_ta_a': "\ta_a'\\<^bsub>o\<^esub> ! a'_n = NormalAction (NewHeapElem ad CTn')" using E_a' a'_n unfolding action_obs_def by simp show ?thesis proof(cases "a < n") case True with a adal E n obtain t_a CTn_a where "?init_obs ! a = (t_a, NormalAction (NewHeapElem ad CTn_a))" by(cases "?init_obs ! a")(fastforce elim!: is_new_action.cases simp add: action_obs_def lnth_lappend1 new_actions_for_def)+ with subsetD[OF dom_start_state, of ad] n True have a_shr_\: "ad \ allocated (shr ?start_state)" by(fastforce simp add: set_conv_nth intro: rev_image_eqI) have E'_unfold': "E' = lappend (ltake (enat a'_m) E') (LCons (lnth E' a'_m) (ldropn (Suc a'_m) E'))" unfolding ldropn_Suc_conv_ldropn[OF a'_m] by simp hence "mthr.if.mthr.Runs ?start_state (lappend (ltake (enat a'_m) E') (LCons (lnth E' a'_m) (ldropn (Suc a'_m) E')))" using \Runs by simp then obtain \' where \_\': "mthr.if.mthr.Trsys ?start_state (list_of (ltake (enat a'_m) E')) \'" and \Runs': "mthr.if.mthr.Runs \' (LCons (lnth E' a'_m) (ldropn (Suc a'_m) E'))" by(rule mthr.if.mthr.Runs_lappendE) simp from \Runs' obtain \'' where red_a': "mthr.if.redT \' (t_a', ta_a') \''" and \Runs'': "mthr.if.mthr.Runs \'' (ldropn (Suc a'_m) E')" unfolding E'_a'_m by cases from New_ta_a' a'_n have "NormalAction (NewHeapElem ad CTn') \ set \ta_a'\\<^bsub>o\<^esub>" unfolding in_set_conv_nth by blast with red_a' obtain x_a' x'_a' m'_a' where red'_a': "mthr.init_fin t_a' (x_a', shr \') ta_a' (x'_a', m'_a')" and \''': "redT_upd \' t_a' ta_a' x'_a' m'_a' \''" and ts_t_a': "thr \' t_a' = \(x_a', no_wait_locks)\" by cases auto from red'_a' \NormalAction (NewHeapElem ad CTn') \ set \ta_a'\\<^bsub>o\<^esub>\ obtain ta'_a' X_a' X'_a' where x_a': "x_a' = (Running, X_a')" and x'_a': "x'_a' = (Running, X'_a')" and ta_a': "ta_a' = convert_TA_initial (convert_obs_initial ta'_a')" and red''_a': "t_a' \ \X_a', shr \'\ -ta'_a'\ \X'_a', m'_a'\" by cases fastforce+ from ta_a' New_ta_a' a'_n have New_ta'_a': "\ta'_a'\\<^bsub>o\<^esub> ! a'_n = NewHeapElem ad CTn'" and a'_n': "a'_n < length \ta'_a'\\<^bsub>o\<^esub>" by auto hence "NewHeapElem ad CTn' \ set \ta'_a'\\<^bsub>o\<^esub>" unfolding in_set_conv_nth by blast with red''_a' have allocated_ad': "ad \ allocated (shr \')" by(auto dest: red_New_allocatedD) have "allocated (shr ?start_state) \ allocated (shr \')" using \_\' unfolding mthr.if.RedT_def[symmetric] by(rule init_fin_RedT_allocated_mono) hence False using allocated_ad' a_shr_\ by blast thus ?thesis .. next case False hence "n \ a" by simp from E E'' a n \n \ a\ adal have a: "a - n \ new_actions_for P E'' adal" by(auto simp add: new_actions_for_def lnth_lappend2 action_obs_def actions_lappend elim: actionsE) from a have "a - n \ actions E''" by(auto elim: new_actionsE) hence "enat (a - n) < llength E''" by(rule actionsE) with \Runs obtain a_m a_n t_a ta_a where E_a: "lnth E'' (a - n) = (t_a, \ta_a\\<^bsub>o\<^esub> ! a_n)" and a_n: "a_n < length \ta_a\\<^bsub>o\<^esub>" and a_m: "enat a_m < llength E'" and a_conv: "a - n = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + a_n" and E'_a_m: "lnth E' a_m = (t_a, ta_a)" unfolding E' by(rule mthr.if.actions_\E_aux) from a have "is_new_action (action_obs E'' (a - n))" and "(ad, al) \ action_loc P E'' (a - n)" unfolding adal by(auto elim: new_actionsE) then obtain CTn where "action_obs E'' (a - n) = NormalAction (NewHeapElem ad CTn)" by cases(fastforce)+ hence New_ta_a: " \ta_a\\<^bsub>o\<^esub> ! a_n = NormalAction (NewHeapElem ad CTn)" using E_a a_n unfolding action_obs_def by simp let ?E' = "ldropn (Suc a_m) E'" have E'_unfold: "E' = lappend (ltake (enat a_m) E') (LCons (lnth E' a_m) ?E')" unfolding ldropn_Suc_conv_ldropn[OF a_m] by simp hence "mthr.if.mthr.Runs ?start_state (lappend (ltake (enat a_m) E') (LCons (lnth E' a_m) ?E'))" using \Runs by simp then obtain \' where \_\': "mthr.if.mthr.Trsys ?start_state (list_of (ltake (enat a_m) E')) \'" and \Runs': "mthr.if.mthr.Runs \' (LCons (lnth E' a_m) ?E')" by(rule mthr.if.mthr.Runs_lappendE) simp from \Runs' obtain \'' where red_a: "mthr.if.redT \' (t_a, ta_a) \''" and \Runs'': "mthr.if.mthr.Runs \'' ?E'" unfolding E'_a_m by cases from New_ta_a a_n have "NormalAction (NewHeapElem ad CTn) \ set \ta_a\\<^bsub>o\<^esub>" unfolding in_set_conv_nth by blast with red_a obtain x_a x'_a m'_a where red'_a: "mthr.init_fin t_a (x_a, shr \') ta_a (x'_a, m'_a)" and \''': "redT_upd \' t_a ta_a x'_a m'_a \''" and ts_t_a: "thr \' t_a = \(x_a, no_wait_locks)\" by cases auto from red'_a \NormalAction (NewHeapElem ad CTn) \ set \ta_a\\<^bsub>o\<^esub>\ obtain ta'_a X_a X'_a where x_a: "x_a = (Running, X_a)" and x'_a: "x'_a = (Running, X'_a)" and ta_a: "ta_a = convert_TA_initial (convert_obs_initial ta'_a)" and red''_a: "t_a \ (X_a, shr \') -ta'_a\ (X'_a, m'_a)" by cases fastforce+ from ta_a New_ta_a a_n have New_ta'_a: "\ta'_a\\<^bsub>o\<^esub> ! a_n = NewHeapElem ad CTn" and a_n': "a_n < length \ta'_a\\<^bsub>o\<^esub>" by auto hence "NewHeapElem ad CTn \ set \ta'_a\\<^bsub>o\<^esub>" unfolding in_set_conv_nth by blast with red''_a have allocated_m'_a_ad: "ad \ allocated m'_a" by(auto dest: red_New_allocatedD) have "a_m \ a'_m" proof(rule ccontr) assume "\ ?thesis" hence "a'_m < a_m" by simp hence "(\isnd (lnth E' i)\\<^bsub>o\<^esub>) = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + (\i = a'_m..snd (lnth E' i)\\<^bsub>o\<^esub>)" by(simp add: sum_upto_add_nat) hence "a' - n < a - n" using \a'_m < a_m\ a'_n E'_a'_m unfolding a_conv a'_conv by(subst (asm) sum.atLeast_Suc_lessThan) simp_all with a_a' show False by simp qed have a'_less: "a' - n < (a - n) - a_n + length \ta_a\\<^bsub>o\<^esub>" proof(rule ccontr) assume "\ ?thesis" hence a'_greater: "(a - n) - a_n + length \ta_a\\<^bsub>o\<^esub> \ a' - n" by simp have "a_m < a'_m" proof(rule ccontr) assume "\ ?thesis" with \a_m \ a'_m\ have "a_m = a'_m" by simp with a'_greater a_n a'_n E'_a'_m E'_a_m show False unfolding a_conv a'_conv by simp qed hence a'_m_a_m: "enat (a'_m - Suc a_m) < llength ?E'" using a'_m by(cases "llength E'") simp_all from \a_m < a'_m\ a'_m E'_a'_m have E'_a'_m': "lnth ?E' (a'_m - Suc a_m) = (t_a', ta_a')" by simp have E'_unfold': "?E' = lappend (ltake (enat (a'_m - Suc a_m)) ?E') (LCons (lnth ?E' (a'_m - Suc a_m)) (ldropn (Suc (a'_m - Suc a_m)) ?E'))" unfolding ldropn_Suc_conv_ldropn[OF a'_m_a_m] lappend_ltake_enat_ldropn .. hence "mthr.if.mthr.Runs \'' (lappend (ltake (enat (a'_m - Suc a_m)) ?E') (LCons (lnth ?E' (a'_m - Suc a_m)) (ldropn (Suc (a'_m - Suc a_m)) ?E')))" using \Runs'' by simp then obtain \''' where \''_\''': "mthr.if.mthr.Trsys \'' (list_of (ltake (enat (a'_m - Suc a_m)) ?E')) \'''" and \Runs''': "mthr.if.mthr.Runs \''' (LCons (lnth ?E' (a'_m - Suc a_m)) (ldropn (Suc (a'_m - Suc a_m)) ?E'))" by(rule mthr.if.mthr.Runs_lappendE) simp from \Runs''' obtain \'''' where red_a': "mthr.if.redT \''' (t_a', ta_a') \''''" and \Runs'''': "mthr.if.mthr.Runs \'''' (ldropn (Suc (a'_m - Suc a_m)) ?E')" unfolding E'_a'_m' by cases from New_ta_a' a'_n have "NormalAction (NewHeapElem ad CTn') \ set \ta_a'\\<^bsub>o\<^esub>" unfolding in_set_conv_nth by blast with red_a' obtain x_a' x'_a' m'_a' where red'_a': "mthr.init_fin t_a' (x_a', shr \''') ta_a' (x'_a', m'_a')" and \'''''': "redT_upd \''' t_a' ta_a' x'_a' m'_a' \''''" and ts_t_a': "thr \''' t_a' = \(x_a', no_wait_locks)\" by cases auto from red'_a' \NormalAction (NewHeapElem ad CTn') \ set \ta_a'\\<^bsub>o\<^esub>\ obtain ta'_a' X_a' X'_a' where x_a': "x_a' = (Running, X_a')" and x'_a': "x'_a' = (Running, X'_a')" and ta_a': "ta_a' = convert_TA_initial (convert_obs_initial ta'_a')" and red''_a': "t_a' \ (X_a', shr \''') -ta'_a'\ (X'_a', m'_a')" by cases fastforce+ from ta_a' New_ta_a' a'_n have New_ta'_a': "\ta'_a'\\<^bsub>o\<^esub> ! a'_n = NewHeapElem ad CTn'" and a'_n': "a'_n < length \ta'_a'\\<^bsub>o\<^esub>" by auto hence "NewHeapElem ad CTn' \ set \ta'_a'\\<^bsub>o\<^esub>" unfolding in_set_conv_nth by blast with red''_a' have allocated_ad': "ad \ allocated (shr \''')" by(auto dest: red_New_allocatedD) have "allocated m'_a = allocated (shr \'')" using \''' by auto also have "\ \ allocated (shr \''')" using \''_\''' unfolding mthr.if.RedT_def[symmetric] by(rule init_fin_RedT_allocated_mono) finally have "ad \ allocated (shr \''')" using allocated_m'_a_ad by blast with allocated_ad' show False by contradiction qed from \a_m \ a'_m\ have [simp]: "a_m = a'_m" proof(rule le_antisym) show "a'_m \ a_m" proof(rule ccontr) assume "\ ?thesis" hence "a_m < a'_m" by simp hence "(\isnd (lnth E' i)\\<^bsub>o\<^esub>) = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + (\i = a_m..snd (lnth E' i)\\<^bsub>o\<^esub>)" by(simp add: sum_upto_add_nat) with a'_less \a_m < a'_m\ E'_a_m a_n a'_n show False unfolding a'_conv a_conv by(subst (asm) sum.atLeast_Suc_lessThan) simp_all qed qed with E'_a_m E'_a'_m have [simp]: "t_a' = t_a" "ta_a' = ta_a" by simp_all from New_ta_a' a'_n ta_a have a'_n': "a'_n < length \ta'_a\\<^bsub>o\<^esub>" and New_ta'_a': "\ta'_a\\<^bsub>o\<^esub> ! a'_n = NewHeapElem ad CTn'" by auto with red''_a New_ta'_a a_n' have "a'_n = a_n" by(auto dest: red_New_same_addr_same) with \a_m = a'_m\ have "a - n = a' - n" unfolding a_conv a'_conv by simp thus ?thesis using \n \ a\ \n \ a'\ by simp qed qed qed end text \Knowledge of addresses of a multithreaded state\ fun ka_Val :: "'addr val \ 'addr set" where "ka_Val (Addr a) = {a}" | "ka_Val _ = {}" fun new_obs_addr :: "('addr, 'thread_id) obs_event \ 'addr set" where "new_obs_addr (ReadMem ad al (Addr ad')) = {ad'}" | "new_obs_addr (NewHeapElem ad hT) = {ad}" | "new_obs_addr _ = {}" lemma new_obs_addr_cases[consumes 1, case_names ReadMem NewHeapElem, cases set]: assumes "ad \ new_obs_addr ob" obtains ad' al where "ob = ReadMem ad' al (Addr ad)" | CTn where "ob = NewHeapElem ad CTn" using assms by(cases ob rule: new_obs_addr.cases) auto definition new_obs_addrs :: "('addr, 'thread_id) obs_event list \ 'addr set" where "new_obs_addrs obs = \(new_obs_addr ` set obs)" fun new_obs_addr_if :: "('addr, 'thread_id) obs_event action \ 'addr set" where "new_obs_addr_if (NormalAction a) = new_obs_addr a" | "new_obs_addr_if _ = {}" definition new_obs_addrs_if :: "('addr, 'thread_id) obs_event action list \ 'addr set" where "new_obs_addrs_if obs = \(new_obs_addr_if ` set obs)" lemma ka_Val_subset_new_obs_Addr_ReadMem: "ka_Val v \ new_obs_addr (ReadMem ad al v)" by(cases v) simp_all lemma typeof_ka: "typeof v \ None \ ka_Val v = {}" by(cases v) simp_all lemma ka_Val_undefined_value [simp]: "ka_Val undefined_value = {}" apply(cases "undefined_value :: 'a val") apply(bestsimp simp add: undefined_value_not_Addr dest: subst)+ done locale known_addrs_base = fixes known_addrs :: "'t \ 'x \ 'addr set" begin definition known_addrs_thr :: "('l, 't, 'x) thread_info \ 'addr set" where "known_addrs_thr ts = (\t \ dom ts. known_addrs t (fst (the (ts t))))" definition known_addrs_state :: "('l,'t,'x,'m,'w) state \ 'addr set" where "known_addrs_state s = known_addrs_thr (thr s)" lemma known_addrs_state_simps [simp]: "known_addrs_state (ls, (ts, m), ws) = known_addrs_thr ts" by(simp add: known_addrs_state_def) lemma known_addrs_thr_cases[consumes 1, case_names known_addrs, cases set: known_addrs_thr]: assumes "ad \ known_addrs_thr ts" and "\t x ln. \ ts t = \(x, ln)\; ad \ known_addrs t x \ \ thesis" shows thesis using assms by(auto simp add: known_addrs_thr_def ran_def) lemma known_addrs_stateI: "\ln. \ ad \ known_addrs t x; thr s t = \(x, ln)\ \ \ ad \ known_addrs_state s" by(fastforce simp add: known_addrs_state_def known_addrs_thr_def intro: rev_bexI) fun known_addrs_if :: "'t \ status \ 'x \ 'addr set" where "known_addrs_if t (s, x) = known_addrs t x" end locale if_known_addrs_base = known_addrs_base known_addrs + multithreaded_base final r convert_RA for known_addrs :: "'t \ 'x \ 'addr set" and final :: "'x \ bool" and r :: "('addr, 't, 'x, 'heap, 'addr, 'obs) semantics" ("_ \ _ -_\ _" [50,0,0,50] 80) and convert_RA :: "'addr released_locks \ 'obs list" sublocale if_known_addrs_base < "if": known_addrs_base known_addrs_if . locale known_addrs = allocated_multithreaded (* Check why all the heap operations are necessary in this locale! *) addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write allocated final r P + if_known_addrs_base known_addrs final r convert_RA for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and allocated :: "'heap \ 'addr set" and known_addrs :: "'thread_id \ 'x \ 'addr set" and final :: "'x \ bool" and r :: "('addr, 'thread_id, 'x, 'heap, 'addr, ('addr, 'thread_id) obs_event) semantics" ("_ \ _ -_\ _" [50,0,0,50] 80) and P :: "'md prog" + assumes red_known_addrs_new: "t \ \x, m\ -ta\ \x', m'\ \ known_addrs t x' \ known_addrs t x \ new_obs_addrs \ta\\<^bsub>o\<^esub>" and red_known_addrs_new_thread: "\ t \ \x, m\ -ta\ \x', m'\; NewThread t' x'' m'' \ set \ta\\<^bsub>t\<^esub> \ \ known_addrs t' x'' \ known_addrs t x" and red_read_knows_addr: "\ t \ \x, m\ -ta\ \x', m'\; ReadMem ad al v \ set \ta\\<^bsub>o\<^esub> \ \ ad \ known_addrs t x" and red_write_knows_addr: "\ t \ \x, m\ -ta\ \x', m'\; \ta\\<^bsub>o\<^esub> ! n = WriteMem ad al (Addr ad'); n < length \ta\\<^bsub>o\<^esub> \ \ ad' \ known_addrs t x \ ad' \ new_obs_addrs (take n \ta\\<^bsub>o\<^esub>)" \ \second possibility necessary for @{term heap_clone}\ begin notation mthr.redT_syntax1 ("_ -_\_\ _" [50,0,0,50] 80) lemma if_red_known_addrs_new: assumes "t \ (x, m) -ta\i (x', m')" shows "known_addrs_if t x' \ known_addrs_if t x \ new_obs_addrs_if \ta\\<^bsub>o\<^esub>" using assms by cases(auto dest!: red_known_addrs_new simp add: new_obs_addrs_if_def new_obs_addrs_def) lemma if_red_known_addrs_new_thread: assumes "t \ (x, m) -ta\i (x', m')" "NewThread t' x'' m'' \ set \ta\\<^bsub>t\<^esub>" shows "known_addrs_if t' x'' \ known_addrs_if t x" using assms by cases(fastforce dest: red_known_addrs_new_thread)+ lemma if_red_read_knows_addr: assumes "t \ (x, m) -ta\i (x', m')" "NormalAction (ReadMem ad al v) \ set \ta\\<^bsub>o\<^esub>" shows "ad \ known_addrs_if t x" using assms by cases(fastforce dest: red_read_knows_addr)+ lemma if_red_write_knows_addr: assumes "t \ (x, m) -ta\i (x', m')" and "\ta\\<^bsub>o\<^esub> ! n = NormalAction (WriteMem ad al (Addr ad'))" "n < length \ta\\<^bsub>o\<^esub>" shows "ad' \ known_addrs_if t x \ ad' \ new_obs_addrs_if (take n \ta\\<^bsub>o\<^esub>)" using assms by cases(auto dest: red_write_knows_addr simp add: new_obs_addrs_if_def new_obs_addrs_def take_map) lemma if_redT_known_addrs_new: assumes redT: "mthr.if.redT s (t, ta) s'" shows "if.known_addrs_state s' \ if.known_addrs_state s \ new_obs_addrs_if \ta\\<^bsub>o\<^esub>" using redT proof(cases) case redT_acquire thus ?thesis by(cases s)(fastforce simp add: if.known_addrs_thr_def split: if_split_asm intro: rev_bexI) next case (redT_normal x x' m) note red = \t \ (x, shr s) -ta\i (x', m)\ show ?thesis proof fix ad assume "ad \ if.known_addrs_state s'" hence "ad \ if.known_addrs_thr (thr s')" by(simp add: if.known_addrs_state_def) then obtain t' x'' ln'' where ts't': "thr s' t' = \(x'', ln'')\" and ad: "ad \ known_addrs_if t' x''" by(rule if.known_addrs_thr_cases) show "ad \ if.known_addrs_state s \ new_obs_addrs_if \ta\\<^bsub>o\<^esub>" proof(cases "thr s t'") case None with redT_normal \thr s' t' = \(x'', ln'')\\ obtain m'' where "NewThread t' x'' m'' \ set \ta\\<^bsub>t\<^esub>" by(fastforce dest: redT_updTs_new_thread split: if_split_asm) with red have "known_addrs_if t' x'' \ known_addrs_if t x" by(rule if_red_known_addrs_new_thread) also have "\ \ known_addrs_if t x \ new_obs_addrs_if \ta\\<^bsub>o\<^esub>" by simp finally have "ad \ known_addrs_if t x \ new_obs_addrs_if \ta\\<^bsub>o\<^esub>" using ad by blast thus ?thesis using \thr s t = \(x, no_wait_locks)\\ by(blast intro: if.known_addrs_stateI) next case (Some xln) show ?thesis proof(cases "t = t'") case True with redT_normal ts't' if_red_known_addrs_new[OF red] ad have "ad \ known_addrs_if t x \ new_obs_addrs_if \ta\\<^bsub>o\<^esub>" by auto thus ?thesis using \thr s t = \(x, no_wait_locks)\\ by(blast intro: if.known_addrs_stateI) next case False with ts't' redT_normal ad Some show ?thesis by(fastforce dest: redT_updTs_Some[where ts="thr s" and t=t'] intro: if.known_addrs_stateI) qed qed qed qed lemma if_redT_read_knows_addr: assumes redT: "mthr.if.redT s (t, ta) s'" and read: "NormalAction (ReadMem ad al v) \ set \ta\\<^bsub>o\<^esub>" shows "ad \ if.known_addrs_state s" using redT proof(cases) case redT_acquire thus ?thesis using read by auto next case (redT_normal x x' m') with if_red_read_knows_addr[OF \t \ (x, shr s) -ta\i (x', m')\ read] show ?thesis by(auto simp add: if.known_addrs_state_def if.known_addrs_thr_def intro: bexI[where x=t]) qed lemma init_fin_redT_known_addrs_subset: assumes "mthr.if.redT s (t, ta) s'" shows "if.known_addrs_state s' \ if.known_addrs_state s \ known_addrs_if t (fst (the (thr s' t)))" using assms apply(cases) apply(rule subsetI) apply(clarsimp simp add: if.known_addrs_thr_def split: if_split_asm) apply(rename_tac status x status' x' m' a ws' t'' status'' x'' ln'') apply(case_tac "thr s t''") apply(drule (2) redT_updTs_new_thread) apply clarsimp apply(drule (1) if_red_known_addrs_new_thread) apply simp apply(drule (1) subsetD) apply(rule_tac x="(status, x)" in if.known_addrs_stateI) apply(simp) apply simp apply(frule_tac t="t''" in redT_updTs_Some, assumption) apply clarsimp apply(rule_tac x="(status'', x'')" in if.known_addrs_stateI) apply simp apply simp apply(auto simp add: if.known_addrs_state_def if.known_addrs_thr_def split: if_split_asm) done lemma w_values_no_write_unchanged: assumes no_write: "\w. \ w \ set obs; is_write_action w; adal \ action_loc_aux P w \ \ False" shows "w_values P vs obs adal = vs adal" using assms proof(induct obs arbitrary: vs) case Nil show ?case by simp next case (Cons ob obs) from Cons.prems[of ob] have "w_value P vs ob adal = vs adal" by(cases adal)(cases ob rule: w_value_cases, auto simp add: addr_locs_def split: htype.split_asm, blast+) moreover have "w_values P (w_value P vs ob) obs adal = w_value P vs ob adal" proof(rule Cons.hyps) fix w assume "w \ set obs" "is_write_action w" "adal \ action_loc_aux P w" with Cons.prems[of w] \w_value P vs ob adal = vs adal\ show "False" by simp qed ultimately show ?case by simp qed lemma redT_non_speculative_known_addrs_allocated: assumes red: "mthr.if.redT s (t, ta) s'" and tasc: "non_speculative P vs (llist_of \ta\\<^bsub>o\<^esub>)" and ka: "if.known_addrs_state s \ allocated (shr s)" and vs: "w_addrs vs \ allocated (shr s)" shows "if.known_addrs_state s' \ allocated (shr s')" (is "?thesis1") and "w_addrs (w_values P vs \ta\\<^bsub>o\<^esub>) \ allocated (shr s')" (is "?thesis2") proof - have "?thesis1 \ ?thesis2" using red proof(cases) case (redT_acquire x ln n) hence "if.known_addrs_state s' = if.known_addrs_state s" by(auto 4 4 simp add: if.known_addrs_state_def if.known_addrs_thr_def split: if_split_asm dest: bspec) also note ka also from redT_acquire have "shr s = shr s'" by simp finally have "if.known_addrs_state s' \ allocated (shr s')" . moreover have "w_values P vs \ta\\<^bsub>o\<^esub> = vs" using redT_acquire by(fastforce intro!: w_values_no_write_unchanged del: equalityI dest: convert_RA_not_write) ultimately show ?thesis using vs by(simp add: \shr s = shr s'\) next case (redT_normal x x' m') note red = \t \ (x, shr s) -ta\i (x', m')\ and tst = \thr s t = \(x, no_wait_locks)\\ have allocated_subset: "allocated (shr s) \ allocated (shr s')" using \mthr.if.redT s (t, ta) s'\ by(rule init_fin_redT_allocated_mono) with vs have vs': "w_addrs vs \ allocated (shr s')" by blast { fix obs obs' assume "\ta\\<^bsub>o\<^esub> = obs @ obs'" moreover with tasc have "non_speculative P vs (llist_of obs)" by(simp add: lappend_llist_of_llist_of[symmetric] non_speculative_lappend del: lappend_llist_of_llist_of) ultimately have "w_addrs (w_values P vs obs) \ new_obs_addrs_if obs \ allocated (shr s')" (is "?concl obs") proof(induct obs arbitrary: obs' rule: rev_induct) case Nil thus ?case using vs' by(simp add: new_obs_addrs_if_def) next case (snoc ob obs) note ta = \\ta\\<^bsub>o\<^esub> = (obs @ [ob]) @ obs'\ note tasc = \non_speculative P vs (llist_of (obs @ [ob]))\ from snoc have IH: "?concl obs" by(simp add: lappend_llist_of_llist_of[symmetric] non_speculative_lappend del: lappend_llist_of_llist_of) hence "?concl (obs @ [ob])" proof(cases "ob" rule: mrw_value_cases) case (1 ad' al v) note ob = \ob = NormalAction (WriteMem ad' al v)\ with ta have Write: "\ta\\<^bsub>o\<^esub> ! length obs = NormalAction (WriteMem ad' al v)" by simp show ?thesis proof fix ad'' assume "ad'' \ w_addrs (w_values P vs (obs @ [ob])) \ new_obs_addrs_if (obs @ [ob])" hence "ad'' \ w_addrs (w_values P vs obs) \ new_obs_addrs_if obs \ v = Addr ad''" by(auto simp add: ob w_addrs_def ran_def new_obs_addrs_if_def split: if_split_asm) thus "ad'' \ allocated (shr s')" proof assume "ad'' \ w_addrs (w_values P vs obs) \ new_obs_addrs_if obs" also note IH finally show ?thesis . next assume v: "v = Addr ad''" with Write have "\ta\\<^bsub>o\<^esub> ! length obs = NormalAction (WriteMem ad' al (Addr ad''))" by simp with red have "ad'' \ known_addrs_if t x \ ad'' \ new_obs_addrs_if (take (length obs) \ta\\<^bsub>o\<^esub>)" by(rule if_red_write_knows_addr)(simp add: ta) thus ?thesis proof assume "ad'' \ known_addrs_if t x" hence "ad'' \ if.known_addrs_state s" using tst by(rule if.known_addrs_stateI) with ka allocated_subset show ?thesis by blast next assume "ad'' \ new_obs_addrs_if (take (length obs) \ta\\<^bsub>o\<^esub>)" with ta have "ad'' \ new_obs_addrs_if obs" by simp with IH show ?thesis by blast qed qed qed next case (2 ad hT) hence ob: "ob = NormalAction (NewHeapElem ad hT)" by simp hence "w_addrs (w_values P vs (obs @ [ob])) \ w_addrs (w_values P vs obs)" by(cases hT)(auto simp add: w_addrs_def default_val_not_Addr Addr_not_default_val) moreover from ob ta have "NormalAction (NewHeapElem ad hT) \ set \ta\\<^bsub>o\<^esub>" by simp from init_fin_red_New_allocatedD[OF red this] have "ad \ allocated m'" .. with redT_normal have "ad \ allocated (shr s')" by auto ultimately show ?thesis using IH ob by(auto simp add: new_obs_addrs_if_def) next case (4 ad al v) note ob = \ob = NormalAction (ReadMem ad al v)\ { fix ad' assume v: "v = Addr ad'" with tasc ob have mrw: "Addr ad' \ w_values P vs obs (ad, al)" by(auto simp add: lappend_llist_of_llist_of[symmetric] non_speculative_lappend simp del: lappend_llist_of_llist_of) hence "ad' \ w_addrs (w_values P vs obs)" by(auto simp add: w_addrs_def) with IH have "ad' \ allocated (shr s')" by blast } with ob IH show ?thesis by(cases v)(simp_all add: new_obs_addrs_if_def) qed(simp_all add: new_obs_addrs_if_def) thus ?case by simp qed } note this[of "\ta\\<^bsub>o\<^esub>" "[]"] moreover have "if.known_addrs_state s' \ if.known_addrs_state s \ new_obs_addrs_if \ta\\<^bsub>o\<^esub>" using \mthr.if.redT s (t, ta) s'\ by(rule if_redT_known_addrs_new) ultimately show ?thesis using ka allocated_subset by blast qed thus ?thesis1 ?thesis2 by simp_all qed lemma RedT_non_speculative_known_addrs_allocated: assumes red: "mthr.if.RedT s ttas s'" and tasc: "non_speculative P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" and ka: "if.known_addrs_state s \ allocated (shr s)" and vs: "w_addrs vs \ allocated (shr s)" shows "if.known_addrs_state s' \ allocated (shr s')" (is "?thesis1 s'") and "w_addrs (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) \ allocated (shr s')" (is "?thesis2 s' ttas") proof - from red tasc have "?thesis1 s' \ ?thesis2 s' ttas" proof(induct rule: mthr.if.RedT_induct') case refl thus ?case using ka vs by simp next case (step ttas s' t ta s'') hence "non_speculative P (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of \ta\\<^bsub>o\<^esub>)" and "?thesis1 s'" "?thesis2 s' ttas" by(simp_all add: lappend_llist_of_llist_of[symmetric] non_speculative_lappend del: lappend_llist_of_llist_of) from redT_non_speculative_known_addrs_allocated[OF \mthr.if.redT s' (t, ta) s''\ this] show ?case by simp qed thus "?thesis1 s'" "?thesis2 s' ttas" by simp_all qed lemma read_ex_NewHeapElem [consumes 5, case_names start Red]: assumes RedT: "mthr.if.RedT (init_fin_lift_state status (start_state f P C M vs)) ttas s" and red: "mthr.if.redT s (t, ta) s'" and read: "NormalAction (ReadMem ad al v) \ set \ta\\<^bsub>o\<^esub>" and sc: "non_speculative P (\_. {}) (llist_of (map snd (lift_start_obs start_tid start_heap_obs) @ concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" and known: "known_addrs start_tid (f (fst (method P C M)) M (fst (snd (method P C M))) (fst (snd (snd (method P C M)))) (the (snd (snd (snd (method P C M))))) vs) \ allocated start_heap" obtains (start) CTn where "NewHeapElem ad CTn \ set start_heap_obs" | (Red) ttas' s'' t' ta' s''' ttas'' CTn where "mthr.if.RedT (init_fin_lift_state status (start_state f P C M vs)) ttas' s''" and "mthr.if.redT s'' (t', ta') s'''" and "mthr.if.RedT s''' ttas'' s" and "ttas = ttas' @ (t', ta') # ttas''" and "NormalAction (NewHeapElem ad CTn) \ set \ta'\\<^bsub>o\<^esub>" proof - let ?start_state = "init_fin_lift_state status (start_state f P C M vs)" let ?obs_prefix = "lift_start_obs start_tid start_heap_obs" let ?vs_start = "w_values P (\_. {}) (map snd ?obs_prefix)" from sc have "non_speculative P (w_values P (\_. {}) (map snd (lift_start_obs start_tid start_heap_obs))) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" by(simp add: non_speculative_lappend lappend_llist_of_llist_of[symmetric] del: lappend_llist_of_llist_of) with RedT have "if.known_addrs_state s \ allocated (shr s)" proof(rule RedT_non_speculative_known_addrs_allocated) show "if.known_addrs_state ?start_state \ allocated (shr ?start_state)" using known by(auto simp add: if.known_addrs_state_def if.known_addrs_thr_def start_state_def init_fin_lift_state_def split_beta split: if_split_asm) have "w_addrs ?vs_start \ w_addrs (\_. {})" by(rule w_addrs_lift_start_heap_obs) thus "w_addrs ?vs_start \ allocated (shr ?start_state)" by simp qed also from red read obtain x_ra x'_ra m'_ra where red'_ra: "t \ (x_ra, shr s) -ta\i (x'_ra, m'_ra)" and s': "redT_upd s t ta x'_ra m'_ra s'" and ts_t: "thr s t = \(x_ra, no_wait_locks)\" by cases auto from red'_ra read have "ad \ known_addrs_if t x_ra" by(rule if_red_read_knows_addr) hence "ad \ if.known_addrs_state s" using ts_t by(rule if.known_addrs_stateI) finally have "ad \ allocated (shr s)" . show ?thesis proof(cases "ad \ allocated start_heap") case True then obtain CTn where "NewHeapElem ad CTn \ set start_heap_obs" unfolding start_addrs_allocated by(blast dest: start_addrs_NewHeapElem_start_heap_obsD) thus ?thesis by(rule start) next case False hence "ad \ allocated (shr ?start_state)" by(simp add: start_state_def split_beta shr_init_fin_lift_state) with RedT \ad \ allocated (shr s)\ obtain t' ta' CTn where tta: "(t', ta') \ set ttas" and new: "NormalAction (NewHeapElem ad CTn) \ set \ta'\\<^bsub>o\<^esub>" by(blast dest: init_fin_RedT_allocated_NewHeapElemD) from tta obtain ttas' ttas'' where ttas: "ttas = ttas' @ (t', ta') # ttas''" by(auto dest: split_list) with RedT obtain s'' s''' where "mthr.if.RedT ?start_state ttas' s''" and "mthr.if.redT s'' (t', ta') s'''" and "mthr.if.RedT s''' ttas'' s" unfolding mthr.if.RedT_def by(auto elim!: rtrancl3p_appendE dest!: converse_rtrancl3p_step) thus thesis using ttas new by(rule Red) qed qed end locale known_addrs_typing = known_addrs addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write allocated known_addrs final r P for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and allocated :: "'heap \ 'addr set" and known_addrs :: "'thread_id \ 'x \ 'addr set" and final :: "'x \ bool" and r :: "('addr, 'thread_id, 'x, 'heap, 'addr, ('addr, 'thread_id) obs_event) semantics" ("_ \ _ -_\ _" [50,0,0,50] 80) and wfx :: "'thread_id \ 'x \ 'heap \ bool" and P :: "'md prog" + assumes wfs_non_speculative_invar: "\ t \ (x, m) -ta\ (x', m'); wfx t x m; vs_conf P m vs; non_speculative P vs (llist_of (map NormalAction \ta\\<^bsub>o\<^esub>)) \ \ wfx t x' m'" and wfs_non_speculative_spawn: "\ t \ (x, m) -ta\ (x', m'); wfx t x m; vs_conf P m vs; non_speculative P vs (llist_of (map NormalAction \ta\\<^bsub>o\<^esub>)); NewThread t'' x'' m'' \ set \ta\\<^bsub>t\<^esub> \ \ wfx t'' x'' m''" and wfs_non_speculative_other: "\ t \ (x, m) -ta\ (x', m'); wfx t x m; vs_conf P m vs; non_speculative P vs (llist_of (map NormalAction \ta\\<^bsub>o\<^esub>)); wfx t'' x'' m \ \ wfx t'' x'' m'" and wfs_non_speculative_vs_conf: "\ t \ (x, m) -ta\ (x', m'); wfx t x m; vs_conf P m vs; non_speculative P vs (llist_of (take n (map NormalAction \ta\\<^bsub>o\<^esub>))) \ \ vs_conf P m' (w_values P vs (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" and red_read_typeable: "\ t \ (x, m) -ta\ (x', m'); wfx t x m; ReadMem ad al v \ set \ta\\<^bsub>o\<^esub> \ \ \T. P,m \ ad@al : T" and red_NewHeapElemD: "\ t \ (x, m) -ta\ (x', m'); wfx t x m; NewHeapElem ad hT \ set \ta\\<^bsub>o\<^esub> \ \ typeof_addr m' ad = \hT\" and red_hext_incr: "\ t \ (x, m) -ta\ (x', m'); wfx t x m; vs_conf P m vs; non_speculative P vs (llist_of (map NormalAction \ta\\<^bsub>o\<^esub>)) \ \ m \ m'" begin lemma redT_wfs_non_speculative_invar: assumes redT: "mthr.redT s (t, ta) s'" and wfx: "ts_ok wfx (thr s) (shr s)" and vs: "vs_conf P (shr s) vs" and ns: "non_speculative P vs (llist_of (map NormalAction \ta\\<^bsub>o\<^esub>))" shows "ts_ok wfx (thr s') (shr s')" using redT proof(cases) case (redT_normal x x' m') with vs wfx ns show ?thesis apply(clarsimp intro!: ts_okI split: if_split_asm) apply(erule wfs_non_speculative_invar, auto dest: ts_okD) apply(rename_tac t' x' ln ws') apply(case_tac "thr s t'") apply(frule (2) redT_updTs_new_thread, clarify) apply(frule (1) mthr.new_thread_memory) apply(auto intro: wfs_non_speculative_other wfs_non_speculative_spawn dest: ts_okD simp add: redT_updTs_Some) done next case (redT_acquire x ln n) thus ?thesis using wfx by(auto intro!: ts_okI dest: ts_okD split: if_split_asm) qed lemma redT_wfs_non_speculative_vs_conf: assumes redT: "mthr.redT s (t, ta) s'" and wfx: "ts_ok wfx (thr s) (shr s)" and conf: "vs_conf P (shr s) vs" and ns: "non_speculative P vs (llist_of (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" shows "vs_conf P (shr s') (w_values P vs (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" using redT proof(cases) case (redT_normal x x' m') thus ?thesis using ns conf wfx by(auto dest: wfs_non_speculative_vs_conf ts_okD) next case (redT_acquire x l ln) have "w_values P vs (take n (map NormalAction (convert_RA ln :: ('addr, 'thread_id) obs_event list))) = vs" by(fastforce dest: in_set_takeD simp add: convert_RA_not_write intro!: w_values_no_write_unchanged del: equalityI) thus ?thesis using conf redT_acquire by(auto) qed lemma if_redT_non_speculative_invar: assumes red: "mthr.if.redT s (t, ta) s'" and ts_ok: "ts_ok (init_fin_lift wfx) (thr s) (shr s)" and sc: "non_speculative P vs (llist_of \ta\\<^bsub>o\<^esub>)" and vs: "vs_conf P (shr s) vs" shows "ts_ok (init_fin_lift wfx) (thr s') (shr s')" proof - let ?s = "\s. (locks s, (\t. map_option (\((status, x), ln). (x, ln)) (thr s t), shr s), wset s, interrupts s)" from ts_ok have ts_ok': "ts_ok wfx (thr (?s s)) (shr (?s s))" by(auto intro!: ts_okI dest: ts_okD) from vs have vs': "vs_conf P (shr (?s s)) vs" by simp from red show ?thesis proof(cases) case (redT_normal x x' m) note tst = \thr s t = \(x, no_wait_locks)\\ from \t \ (x, shr s) -ta\i (x', m)\ show ?thesis proof(cases) case (NormalAction X TA X') from \ta = convert_TA_initial (convert_obs_initial TA)\ \mthr.if.actions_ok s t ta\ have "mthr.actions_ok (?s s) t TA" by(auto elim: rev_iffD1[OF _ thread_oks_ts_change] cond_action_oks_final_change) with tst NormalAction \redT_upd s t ta x' m s'\ have "mthr.redT (?s s) (t, TA) (?s s')" using map_redT_updTs[of snd "thr s" "\ta\\<^bsub>t\<^esub>"] by(auto intro!: mthr.redT.intros simp add: split_def map_prod_def o_def fun_eq_iff) moreover note ts_ok' vs' moreover from \ta = convert_TA_initial (convert_obs_initial TA)\ have "\ta\\<^bsub>o\<^esub> = map NormalAction \TA\\<^bsub>o\<^esub>" by(auto) with sc have "non_speculative P vs (llist_of (map NormalAction \TA\\<^bsub>o\<^esub>))" by simp ultimately have "ts_ok wfx (thr (?s s')) (shr (?s s'))" by(auto dest: redT_wfs_non_speculative_invar) thus ?thesis using \\ta\\<^bsub>o\<^esub> = map NormalAction \TA\\<^bsub>o\<^esub>\ by(auto intro!: ts_okI dest: ts_okD) next case InitialThreadAction with redT_normal ts_ok' vs show ?thesis by(auto 4 3 intro!: ts_okI dest: ts_okD split: if_split_asm) next case ThreadFinishAction with redT_normal ts_ok' vs show ?thesis by(auto 4 3 intro!: ts_okI dest: ts_okD split: if_split_asm) qed next case (redT_acquire x ln l) thus ?thesis using vs ts_ok by(auto 4 3 intro!: ts_okI dest: ts_okD split: if_split_asm) qed qed lemma if_redT_non_speculative_vs_conf: assumes red: "mthr.if.redT s (t, ta) s'" and ts_ok: "ts_ok (init_fin_lift wfx) (thr s) (shr s)" and sc: "non_speculative P vs (llist_of (take n \ta\\<^bsub>o\<^esub>))" and vs: "vs_conf P (shr s) vs" shows "vs_conf P (shr s') (w_values P vs (take n \ta\\<^bsub>o\<^esub>))" proof - let ?s = "\s. (locks s, (\t. map_option (\((status, x), ln). (x, ln)) (thr s t), shr s), wset s, interrupts s)" from ts_ok have ts_ok': "ts_ok wfx (thr (?s s)) (shr (?s s))" by(auto intro!: ts_okI dest: ts_okD) from vs have vs': "vs_conf P (shr (?s s)) vs" by simp from red show ?thesis proof(cases) case (redT_normal x x' m) note tst = \thr s t = \(x, no_wait_locks)\\ from \t \ (x, shr s) -ta\i (x', m)\ show ?thesis proof(cases) case (NormalAction X TA X') from \ta = convert_TA_initial (convert_obs_initial TA)\ \mthr.if.actions_ok s t ta\ have "mthr.actions_ok (?s s) t TA" by(auto elim: rev_iffD1[OF _ thread_oks_ts_change] cond_action_oks_final_change) with tst NormalAction \redT_upd s t ta x' m s'\ have "mthr.redT (?s s) (t, TA) (?s s')" using map_redT_updTs[of snd "thr s" "\ta\\<^bsub>t\<^esub>"] by(auto intro!: mthr.redT.intros simp add: split_def map_prod_def o_def fun_eq_iff) moreover note ts_ok' vs' moreover from \ta = convert_TA_initial (convert_obs_initial TA)\ have "\ta\\<^bsub>o\<^esub> = map NormalAction \TA\\<^bsub>o\<^esub>" by(auto) with sc have "non_speculative P vs (llist_of (take n (map NormalAction \TA\\<^bsub>o\<^esub>)))" by simp ultimately have "vs_conf P (shr (?s s')) (w_values P vs (take n (map NormalAction \TA\\<^bsub>o\<^esub>)))" by(auto dest: redT_wfs_non_speculative_vs_conf) thus ?thesis using \\ta\\<^bsub>o\<^esub> = map NormalAction \TA\\<^bsub>o\<^esub>\ by(auto) next case InitialThreadAction with redT_normal vs show ?thesis by(auto simp add: take_Cons') next case ThreadFinishAction with redT_normal vs show ?thesis by(auto simp add: take_Cons') qed next case (redT_acquire x l ln) have "w_values P vs (take n (map NormalAction (convert_RA ln :: ('addr, 'thread_id) obs_event list))) = vs" by(fastforce simp add: convert_RA_not_write take_Cons' dest: in_set_takeD intro!: w_values_no_write_unchanged del: equalityI) thus ?thesis using vs redT_acquire by auto qed qed lemma if_RedT_non_speculative_invar: assumes red: "mthr.if.RedT s ttas s'" and tsok: "ts_ok (init_fin_lift wfx) (thr s) (shr s)" and sc: "non_speculative P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" and vs: "vs_conf P (shr s) vs" shows "ts_ok (init_fin_lift wfx) (thr s') (shr s')" (is ?thesis1) and "vs_conf P (shr s') (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" (is ?thesis2) using red tsok sc vs unfolding mthr.if.RedT_def proof(induct arbitrary: vs rule: rtrancl3p_converse_induct') case refl case 1 thus ?case by - case 2 thus ?case by simp next case (step s tta s' ttas) obtain t ta where tta: "tta = (t, ta)" by(cases tta) case 1 hence sc1: "non_speculative P vs (llist_of \ta\\<^bsub>o\<^esub>)" and sc2: "non_speculative P (w_values P vs \ta\\<^bsub>o\<^esub>) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" unfolding lconcat_llist_of[symmetric] lmap_llist_of[symmetric] llist.map_comp o_def llist_of.simps llist.map(2) lconcat_LCons tta by(simp_all add: non_speculative_lappend list_of_lconcat o_def) from if_redT_non_speculative_invar[OF step(2)[unfolded tta] _ sc1] if_redT_non_speculative_vs_conf[OF step(2)[unfolded tta], where vs = vs and n="length \ta\\<^bsub>o\<^esub>"] 1 step.hyps(3)[of "w_values P vs \ta\\<^bsub>o\<^esub>"] sc2 sc1 show ?case by simp case 2 hence sc1: "non_speculative P vs (llist_of \ta\\<^bsub>o\<^esub>)" and sc2: "non_speculative P (w_values P vs \ta\\<^bsub>o\<^esub>) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" unfolding lconcat_llist_of[symmetric] lmap_llist_of[symmetric] llist.map_comp o_def llist_of.simps llist.map(2) lconcat_LCons tta by(simp_all add: non_speculative_lappend list_of_lconcat o_def) from if_redT_non_speculative_invar[OF step(2)[unfolded tta] _ sc1] if_redT_non_speculative_vs_conf[OF step(2)[unfolded tta], where vs = vs and n="length \ta\\<^bsub>o\<^esub>"] 2 step.hyps(4)[of "w_values P vs \ta\\<^bsub>o\<^esub>"] sc2 sc1 show ?case by(simp add: tta o_def) qed lemma init_fin_hext_incr: assumes "t \ (x, m) -ta\i (x', m')" and "init_fin_lift wfx t x m" and "non_speculative P vs (llist_of \ta\\<^bsub>o\<^esub>)" and "vs_conf P m vs" shows "m \ m'" using assms by(cases)(auto intro: red_hext_incr) lemma init_fin_redT_hext_incr: assumes "mthr.if.redT s (t, ta) s'" and "ts_ok (init_fin_lift wfx) (thr s) (shr s)" and "non_speculative P vs (llist_of \ta\\<^bsub>o\<^esub>)" and "vs_conf P (shr s) vs" shows "shr s \ shr s'" using assms by(cases)(auto dest: init_fin_hext_incr ts_okD) lemma init_fin_RedT_hext_incr: assumes "mthr.if.RedT s ttas s'" and "ts_ok (init_fin_lift wfx) (thr s) (shr s)" and sc: "non_speculative P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" and vs: "vs_conf P (shr s) vs" shows "shr s \ shr s'" using assms proof(induction rule: mthr.if.RedT_induct') case refl thus ?case by simp next case (step ttas s' t ta s'') note ts_ok = \ts_ok (init_fin_lift wfx) (thr s) (shr s)\ from \non_speculative P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (ttas @ [(t, ta)]))))\ have ns: "non_speculative P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" and ns': "non_speculative P (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of \ta\\<^bsub>o\<^esub>)" by(simp_all add: lappend_llist_of_llist_of[symmetric] non_speculative_lappend del: lappend_llist_of_llist_of) from ts_ok ns have "shr s \ shr s'" using \vs_conf P (shr s) vs\ by(rule step.IH) also have "ts_ok (init_fin_lift wfx) (thr s') (shr s')" using \mthr.if.RedT s ttas s'\ ts_ok ns \vs_conf P (shr s) vs\ by(rule if_RedT_non_speculative_invar) with \mthr.if.redT s' (t, ta) s''\ have "\ \ shr s''" using ns' proof(rule init_fin_redT_hext_incr) from \mthr.if.RedT s ttas s'\ ts_ok ns \vs_conf P (shr s) vs\ show "vs_conf P (shr s') (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" by(rule if_RedT_non_speculative_invar) qed finally show ?case . qed lemma init_fin_red_read_typeable: assumes "t \ (x, m) -ta\i (x', m')" and "init_fin_lift wfx t x m" "NormalAction (ReadMem ad al v) \ set \ta\\<^bsub>o\<^esub>" shows "\T. P,m \ ad@al : T" using assms by cases(auto dest: red_read_typeable) lemma Ex_new_action_for: assumes wf: "wf_syscls P" and wfx_start: "ts_ok wfx (thr (start_state f P C M vs)) start_heap" and ka: "known_addrs start_tid (f (fst (method P C M)) M (fst (snd (method P C M))) (fst (snd (snd (method P C M)))) (the (snd (snd (snd (method P C M))))) vs) \ allocated start_heap" and E: "E \ \_start f P C M vs status" and read: "ra \ read_actions E" and aloc: "adal \ action_loc P E ra" and sc: "non_speculative P (\_. {}) (ltake (enat ra) (lmap snd E))" shows "\wa. wa \ new_actions_for P E adal \ wa < ra" proof - let ?obs_prefix = "lift_start_obs start_tid start_heap_obs" let ?start_state = "init_fin_lift_state status (start_state f P C M vs)" from start_state_vs_conf[OF wf] have vs_conf_start: "vs_conf P start_heap (w_values P (\_. {}) (map NormalAction start_heap_obs))" by(simp add: lift_start_obs_def o_def) obtain ad al where adal: "adal = (ad, al)" by(cases adal) with read aloc obtain v where ra: "action_obs E ra = NormalAction (ReadMem ad al v)" and ra_len: "enat ra < llength E" by(cases "lnth E ra")(auto elim!: read_actions.cases actionsE) from E obtain E'' where E: "E = lappend (llist_of ?obs_prefix) E''" and E'': "E'' \ mthr.if.\ ?start_state" by(auto) from E'' obtain E' where E': "E'' = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E')" and \Runs: "mthr.if.mthr.Runs ?start_state E'" by(rule mthr.if.\.cases) have ra_len': "length ?obs_prefix \ ra" proof(rule ccontr) assume "\ ?thesis" hence "ra < length ?obs_prefix" by simp moreover with ra ra_len E obtain ra' ad al v where "start_heap_obs ! ra' = ReadMem ad al v" "ra' < length start_heap_obs" by(cases ra)(auto simp add: lnth_LCons lnth_lappend1 action_obs_def lift_start_obs_def) ultimately have "ReadMem ad al v \ set start_heap_obs" unfolding in_set_conv_nth by blast thus False by(simp add: start_heap_obs_not_Read) qed let ?n = "length ?obs_prefix" from ra ra_len ra_len' E have "enat (ra - ?n) < llength E''" and ra_obs: "action_obs E'' (ra - ?n) = NormalAction (ReadMem ad al v)" by(cases "llength E''", auto simp add: action_obs_def lnth_lappend2) from \Runs \enat (ra - ?n) < llength E''\ obtain ra_m ra_n t_ra ta_ra where E_ra: "lnth E'' (ra - ?n) = (t_ra, \ta_ra\\<^bsub>o\<^esub> ! ra_n)" and ra_n: "ra_n < length \ta_ra\\<^bsub>o\<^esub>" and ra_m: "enat ra_m < llength E'" and ra_conv: "ra - ?n = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + ra_n" and E'_ra_m: "lnth E' ra_m = (t_ra, ta_ra)" unfolding E' by(rule mthr.if.actions_\E_aux) let ?E' = "ldropn (Suc ra_m) E'" have E'_unfold: "E' = lappend (ltake (enat ra_m) E') (LCons (lnth E' ra_m) ?E')" unfolding ldropn_Suc_conv_ldropn[OF ra_m] by simp hence "mthr.if.mthr.Runs ?start_state (lappend (ltake (enat ra_m) E') (LCons (lnth E' ra_m) ?E'))" using \Runs by simp then obtain \' where \_\': "mthr.if.mthr.Trsys ?start_state (list_of (ltake (enat ra_m) E')) \'" and \Runs': "mthr.if.mthr.Runs \' (LCons (lnth E' ra_m) ?E')" by(rule mthr.if.mthr.Runs_lappendE) simp from \Runs' obtain \'' where red_ra: "mthr.if.redT \' (t_ra, ta_ra) \''" and \Runs'': "mthr.if.mthr.Runs \'' ?E'" unfolding E'_ra_m by cases from E_ra ra_n ra_obs have "NormalAction (ReadMem ad al v) \ set \ta_ra\\<^bsub>o\<^esub>" by(auto simp add: action_obs_def in_set_conv_nth) with red_ra obtain x_ra x'_ra m'_ra where red'_ra: "mthr.init_fin t_ra (x_ra, shr \') ta_ra (x'_ra, m'_ra)" and \'': "redT_upd \' t_ra ta_ra x'_ra m'_ra \''" and ts_t_a: "thr \' t_ra = \(x_ra, no_wait_locks)\" by cases auto from red'_ra \NormalAction (ReadMem ad al v) \ set \ta_ra\\<^bsub>o\<^esub>\ obtain ta'_ra X_ra X'_ra where x_ra: "x_ra = (Running, X_ra)" and x'_ra: "x'_ra = (Running, X'_ra)" and ta_ra: "ta_ra = convert_TA_initial (convert_obs_initial ta'_ra)" and red''_ra: "t_ra \ (X_ra, shr \') -ta'_ra\ (X'_ra, m'_ra)" by cases fastforce+ from \NormalAction (ReadMem ad al v) \ set \ta_ra\\<^bsub>o\<^esub>\ ta_ra have "ReadMem ad al v \ set \ta'_ra\\<^bsub>o\<^esub>" by auto from wfx_start have wfx_start: "ts_ok (init_fin_lift wfx) (thr ?start_state) (shr ?start_state)" by(simp add: start_state_def split_beta) from sc ra_len' have "non_speculative P (w_values P (\_. {}) (map snd ?obs_prefix)) (lmap snd (ltake (enat (ra - ?n)) (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E'))))" unfolding E E' by(simp add: ltake_lappend2 lmap_lappend_distrib non_speculative_lappend) also note ra_conv also note plus_enat_simps(1)[symmetric] also have "enat (\isnd (lnth E' i)\\<^bsub>o\<^esub>) = (\isnd (lnth E' i)\\<^bsub>o\<^esub>))" - by(subst sum_hom[symmetric])(simp_all add: zero_enat_def) + by(subst sum_comp_morphism[symmetric])(simp_all add: zero_enat_def) also have "\ = (\i(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E') i))" using ra_m by-(rule sum.cong[OF refl], simp add: le_less_trans[where y="enat ra_m"] split_beta) also note ltake_plus_conv_lappend also note lconcat_ltake[symmetric] also note lmap_lappend_distrib also note non_speculative_lappend finally have "non_speculative P (w_values P (\_. {}) (map snd ?obs_prefix)) (lmap snd (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (llist_of (list_of (ltake (enat ra_m) E'))))))" by(simp add: split_def) hence sc': "non_speculative P (w_values P (\_. {}) (map snd ?obs_prefix)) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat ra_m) E')))))" unfolding lmap_lconcat llist.map_comp o_def lconcat_llist_of[symmetric] lmap_llist_of[symmetric] by(simp add: split_beta o_def) from vs_conf_start have vs_conf_start: "vs_conf P (shr ?start_state) (w_values P (\_. {}) (map snd ?obs_prefix))" by(simp add:init_fin_lift_state_conv_simps start_state_def split_beta lift_start_obs_def o_def) with \_\' wfx_start sc' have "ts_ok (init_fin_lift wfx) (thr \') (shr \')" unfolding mthr.if.RedT_def[symmetric] by(rule if_RedT_non_speculative_invar) with ts_t_a have "wfx t_ra X_ra (shr \')" unfolding x_ra by(auto dest: ts_okD) with red''_ra \ReadMem ad al v \ set \ta'_ra\\<^bsub>o\<^esub>\ obtain T' where type_adal: "P,shr \' \ ad@al : T'" by(auto dest: red_read_typeable) from sc ra_len' have "non_speculative P (\_. {}) (llist_of (map snd ?obs_prefix))" unfolding E by(simp add: ltake_lappend2 lmap_lappend_distrib non_speculative_lappend) with sc' have sc'': "non_speculative P (\_. {}) (llist_of (map snd (lift_start_obs start_tid start_heap_obs) @ concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat ra_m) E')))))" by(simp add: lappend_llist_of_llist_of[symmetric] non_speculative_lappend del: lappend_llist_of_llist_of) from \_\' red_ra \NormalAction (ReadMem ad al v) \ set \ta_ra\\<^bsub>o\<^esub>\ sc'' ka show "\wa. wa \ new_actions_for P E adal \ wa < ra" unfolding mthr.if.RedT_def[symmetric] proof(cases rule: read_ex_NewHeapElem) case (start CTn) then obtain n where n: "start_heap_obs ! n = NewHeapElem ad CTn" and len: "n < length start_heap_obs" unfolding in_set_conv_nth by blast from len have "Suc n \ actions E" unfolding E by(simp add: actions_def enat_less_enat_plusI) moreover from \_\' have hext: "start_heap \ shr \'" unfolding mthr.if.RedT_def[symmetric] using wfx_start sc' vs_conf_start by(auto dest!: init_fin_RedT_hext_incr simp add: start_state_def split_beta init_fin_lift_state_conv_simps) from start have "typeof_addr start_heap ad = \CTn\" by(auto dest: NewHeapElem_start_heap_obsD[OF wf]) with hext have "typeof_addr (shr \') ad = \CTn\" by(rule typeof_addr_hext_mono) with type_adal have "adal \ action_loc P E (Suc n)" using n len unfolding E adal by cases(auto simp add: action_obs_def lnth_lappend1 lift_start_obs_def) moreover have "is_new_action (action_obs E (Suc n))" using n len unfolding E by(simp add: action_obs_def lnth_lappend1 lift_start_obs_def) ultimately have "Suc n \ new_actions_for P E adal" by(rule new_actionsI) moreover have "Suc n < ra" using ra_len' len by(simp) ultimately show ?thesis by blast next case (Red ttas' s'' t' ta' s''' ttas'' CTn) from \NormalAction (NewHeapElem ad CTn) \ set \ta'\\<^bsub>o\<^esub>\ obtain obs obs' where obs: "\ta'\\<^bsub>o\<^esub> = obs @ NormalAction (NewHeapElem ad CTn) # obs'" by(auto dest: split_list) let ?wa = "?n + length (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')) + length obs" have "enat (length (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')) + length obs) < enat (length (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (ttas' @ [(t', ta')]))))" using obs by simp also have "\ = llength (lconcat (lmap llist_of (lmap (\(t, ta). \ta\\<^bsub>o\<^esub>) (llist_of (ttas' @ [(t', ta')])))))" by(simp del: map_map map_append add: lconcat_llist_of) also have "\ \ llength (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) (llist_of (ttas' @ (t', ta') # ttas''))))" by(auto simp add: o_def split_def intro: lprefix_llist_ofI intro!: lprefix_lconcatI lprefix_llength_le) also note len_less = calculation have "\ \ (\i(t, ta). llist_of \ta\\<^bsub>o\<^esub>) E') i))" unfolding \list_of (ltake (enat ra_m) E') = ttas' @ (t', ta') # ttas''\[symmetric] by(simp add: ltake_lmap[symmetric] lconcat_ltake del: ltake_lmap) also have "\ = enat (\isnd (lnth E' i)\\<^bsub>o\<^esub>)" using ra_m - by(subst sum_hom[symmetric, where f="enat"])(auto intro: sum.cong simp add: zero_enat_def less_trans[where y="enat ra_m"] split_beta) + by(subst sum_comp_morphism[symmetric, where h="enat"])(auto intro: sum.cong simp add: zero_enat_def less_trans[where y="enat ra_m"] split_beta) also have "\ \ enat (ra - ?n)" unfolding ra_conv by simp finally have enat_length: "enat (length (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')) + length obs) < enat (ra - length (lift_start_obs start_tid start_heap_obs))" . then have wa_ra: "?wa < ra" by simp with ra_len have "?wa \ actions E" by(cases "llength E")(simp_all add: actions_def) moreover from \mthr.if.redT s'' (t', ta') s'''\ \NormalAction (NewHeapElem ad CTn) \ set \ta'\\<^bsub>o\<^esub>\ obtain x_wa x_wa' where ts''t': "thr s'' t' = \(x_wa, no_wait_locks)\" and red_wa: "mthr.init_fin t' (x_wa, shr s'') ta' (x_wa', shr s''')" by(cases) fastforce+ from sc' have ns: "non_speculative P (w_values P (\_. {}) (map snd ?obs_prefix)) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')))" and ns': "non_speculative P (w_values P (w_values P (\_. {}) (map snd ?obs_prefix)) (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas'))) (llist_of \ta'\\<^bsub>o\<^esub>)" and ns'': "non_speculative P (w_values P (w_values P (w_values P (\_. {}) (map snd ?obs_prefix)) (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas'))) \ta'\\<^bsub>o\<^esub>) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas'')))" unfolding \list_of (ltake (enat ra_m) E') = ttas' @ (t', ta') # ttas''\ by(simp_all add: lappend_llist_of_llist_of[symmetric] lmap_lappend_distrib non_speculative_lappend del: lappend_llist_of_llist_of) from \mthr.if.RedT ?start_state ttas' s''\ wfx_start ns have ts_ok'': "ts_ok (init_fin_lift wfx) (thr s'') (shr s'')" using vs_conf_start by(rule if_RedT_non_speculative_invar) with ts''t' have wfxt': "wfx t' (snd x_wa) (shr s'')" by(cases x_wa)(auto dest: ts_okD) { have "action_obs E ?wa = snd (lnth (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E')) (length (concat (map (\(t, y). \y\\<^bsub>o\<^esub>) ttas')) + length obs))" unfolding E E' by(simp add: action_obs_def lnth_lappend2) also from enat_length \enat (ra - ?n) < llength E''\ have "\ = lnth (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) E')) (length (concat (map (\(t, y). \y\\<^bsub>o\<^esub>) ttas')) + length obs)" unfolding E' by(subst lnth_lmap[symmetric, where f=snd])(erule (1) less_trans, simp add: lmap_lconcat llist.map_comp split_def o_def) also from len_less have "enat (length (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')) + length obs) < llength (lconcat (ltake (enat ra_m) (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) E')))" unfolding \list_of (ltake (enat ra_m) E') = ttas' @ (t', ta') # ttas''\[symmetric] by(simp add: ltake_lmap[symmetric] del: ltake_lmap) note lnth_lconcat_ltake[OF this, symmetric] also note ltake_lmap also have "ltake (enat ra_m) E' = llist_of (list_of (ltake (enat ra_m) E'))" by(simp) also note \list_of (ltake (enat ra_m) E') = ttas' @ (t', ta') # ttas''\ also note lmap_llist_of also have "(\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) = llist_of \ (\(t, ta). \ta\\<^bsub>o\<^esub>)" by(simp add: o_def split_def) also note map_map[symmetric] also note lconcat_llist_of also note lnth_llist_of also have "concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (ttas' @ (t', ta') # ttas'')) ! (length (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')) + length obs) = NormalAction (NewHeapElem ad CTn)" by(simp add: nth_append obs) finally have "action_obs E ?wa = NormalAction (NewHeapElem ad CTn)" . } note wa_obs = this from \mthr.if.RedT ?start_state ttas' s''\ wfx_start ns vs_conf_start have vs'': "vs_conf P (shr s'') (w_values P (w_values P (\_. {}) (map snd ?obs_prefix)) (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')))" by(rule if_RedT_non_speculative_invar) from if_redT_non_speculative_vs_conf[OF \mthr.if.redT s'' (t', ta') s'''\ ts_ok'' _ vs'', of "length \ta'\\<^bsub>o\<^esub>"] ns' have vs''': "vs_conf P (shr s''') (w_values P (w_values P (w_values P (\_. {}) (map snd ?obs_prefix)) (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas'))) \ta'\\<^bsub>o\<^esub>)" by simp from \mthr.if.redT s'' (t', ta') s'''\ ts_ok'' ns' vs'' have "ts_ok (init_fin_lift wfx) (thr s''') (shr s''')" by(rule if_redT_non_speculative_invar) with \mthr.if.RedT s''' ttas'' \'\ have hext: "shr s''' \ shr \'" using ns'' vs''' by(rule init_fin_RedT_hext_incr) from red_wa have "typeof_addr (shr s''') ad = \CTn\" using wfxt' \NormalAction (NewHeapElem ad CTn) \ set \ta'\\<^bsub>o\<^esub>\ by cases(auto dest: red_NewHeapElemD) with hext have "typeof_addr (shr \') ad = \CTn\" by(rule typeof_addr_hext_mono) with type_adal have "adal \ action_loc P E ?wa" using wa_obs unfolding E adal by cases (auto simp add: action_obs_def lnth_lappend1 lift_start_obs_def) moreover have "is_new_action (action_obs E ?wa)" using wa_obs by simp ultimately have "?wa \ new_actions_for P E adal" by(rule new_actionsI) thus ?thesis using wa_ra by blast qed qed lemma executions_sc_hb: assumes "wf_syscls P" and "ts_ok wfx (thr (start_state f P C M vs)) start_heap" and "known_addrs start_tid (f (fst (method P C M)) M (fst (snd (method P C M))) (fst (snd (snd (method P C M)))) (the (snd (snd (snd (method P C M))))) vs) \ allocated start_heap" shows "executions_sc_hb (\_start f P C M vs status) P" (is "executions_sc_hb ?E P") proof fix E a adal a' assume "E \ ?E" "a \ new_actions_for P E adal" "a' \ new_actions_for P E adal" thus "a = a'" by(rule \_new_actions_for_unique) next fix E ra adal assume "E \ ?E" "ra \ read_actions E" "adal \ action_loc P E ra" and "non_speculative P (\_. {}) (ltake (enat ra) (lmap snd E))" with assms show "\wa. wa \ new_actions_for P E adal \ wa < ra" by(rule Ex_new_action_for) qed lemma executions_aux: assumes wf: "wf_syscls P" and wfx_start: "ts_ok wfx (thr (start_state f P C M vs)) start_heap" (is "ts_ok wfx (thr ?start_state) _") and ka: "known_addrs start_tid (f (fst (method P C M)) M (fst (snd (method P C M))) (fst (snd (snd (method P C M)))) (the (snd (snd (snd (method P C M))))) vs) \ allocated start_heap" shows "executions_aux (\_start f P C M vs status) P" (is "executions_aux ?\ P") proof fix E a adal a' assume "E \ ?\" "a \ new_actions_for P E adal" "a' \ new_actions_for P E adal" thus "a = a'" by(rule \_new_actions_for_unique) next fix E ws r adal assume E: "E \ ?\" and wf_exec: "P \ (E, ws) \" and read: "r \ read_actions E" "adal \ action_loc P E r" and sc: "\a. \a < r; a \ read_actions E\ \ P,E \ a \mrw ws a" interpret jmm: executions_sc_hb ?\ P using wf wfx_start ka by(rule executions_sc_hb) from E wf_exec sc have "ta_seq_consist P Map.empty (ltake (enat r) (lmap snd E))" unfolding ltake_lmap by(rule jmm.ta_seq_consist_mrwI) simp hence "non_speculative P (\_. {}) (ltake (enat r) (lmap snd E))" by(rule ta_seq_consist_into_non_speculative) simp with wf wfx_start ka E read have "\i. i \ new_actions_for P E adal \ i < r" by(rule Ex_new_action_for) thus "\i new_actions_for P E adal" by blast qed lemma drf: assumes cut_and_update: "if.cut_and_update (init_fin_lift_state status (start_state f P C M vs)) (mrw_values P Map.empty (map snd (lift_start_obs start_tid start_heap_obs)))" (is "if.cut_and_update ?start_state (mrw_values _ _ (map _ ?start_heap_obs))") and wf: "wf_syscls P" and wfx_start: "ts_ok wfx (thr (start_state f P C M vs)) start_heap" and ka: "known_addrs start_tid (f (fst (method P C M)) M (fst (snd (method P C M))) (fst (snd (snd (method P C M)))) (the (snd (snd (snd (method P C M))))) vs) \ allocated start_heap" shows "drf (\_start f P C M vs status) P" (is "drf ?\ _") proof - interpret jmm: executions_sc_hb "?\" P using wf wfx_start ka by(rule executions_sc_hb) let ?n = "length ?start_heap_obs" let ?\' = "lappend (llist_of ?start_heap_obs) ` mthr.if.\ ?start_state" show ?thesis proof fix E ws r assume E: "E \ ?\'" and wf: "P \ (E, ws) \" and mrw: "\a. \ a < r; a \ read_actions E \ \ P,E \ a \mrw ws a" show "\E'\?\'. \ws'. P \ (E', ws') \ \ ltake (enat r) E = ltake (enat r) E' \ sequentially_consistent P (E', ws') \ action_tid E r = action_tid E' r \ action_obs E r \ action_obs E' r \ (r \ actions E \ r \ actions E')" proof(cases "\r'. r' \ read_actions E \ r \ r'") case False have "sequentially_consistent P (E, ws)" proof(rule sequentially_consistentI) fix a assume "a \ read_actions E" with False have "a < r" by auto thus "P,E \ a \mrw ws a" using \a \ read_actions E\ by(rule mrw) qed moreover have "action_obs E r \ action_obs E r" by(rule sim_action_refl) ultimately show ?thesis using wf E by blast next case True let ?P = "\r'. r' \ read_actions E \ r \ r'" let ?r = "Least ?P" from True obtain r' where r': "?P r'" by blast hence r: "?P ?r" by(rule LeastI) { fix a assume "a < ?r" "a \ read_actions E" have "P,E \ a \mrw ws a" proof(cases "a < r") case True thus ?thesis using \a \ read_actions E\ by(rule mrw) next case False with \a \ read_actions E\ have "?P a" by simp hence "?r \ a" by(rule Least_le) with \a < ?r\ have False by simp thus ?thesis .. qed } note mrw' = this from E obtain E'' where E: "E = lappend (llist_of ?start_heap_obs) E''" and E'': "E'' \ mthr.if.\ ?start_state" by auto from E'' obtain E' where E': "E'' = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E')" and \Runs: "mthr.if.mthr.Runs ?start_state E'" by(rule mthr.if.\.cases) have r_len: "length ?start_heap_obs \ ?r" proof(rule ccontr) assume "\ ?thesis" hence "?r < length ?start_heap_obs" by simp moreover with r E obtain t ad al v where "?start_heap_obs ! ?r = (t, NormalAction (ReadMem ad al v))" by(cases "?start_heap_obs ! ?r")(fastforce elim!: read_actions.cases simp add: actions_def action_obs_def lnth_lappend1) ultimately have "(t, NormalAction (ReadMem ad al v)) \ set ?start_heap_obs" unfolding in_set_conv_nth by blast thus False by(auto simp add: start_heap_obs_not_Read) qed let ?n = "length ?start_heap_obs" from r r_len E have r: "?r - ?n \ read_actions E''" by(fastforce elim!: read_actions.cases simp add: actions_lappend action_obs_def lnth_lappend2 elim: actionsE intro: read_actions.intros) from r have "?r - ?n \ actions E''" by(auto) hence "enat (?r - ?n) < llength E''" by(rule actionsE) with \Runs obtain r_m r_n t_r ta_r where E_r: "lnth E'' (?r - ?n) = (t_r, \ta_r\\<^bsub>o\<^esub> ! r_n)" and r_n: "r_n < length \ta_r\\<^bsub>o\<^esub>" and r_m: "enat r_m < llength E'" and r_conv: "?r - ?n = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + r_n" and E'_r_m: "lnth E' r_m = (t_r, ta_r)" unfolding E' by(rule mthr.if.actions_\E_aux) let ?E' = "ldropn (Suc r_m) E'" let ?r_m_E' = "ltake (enat r_m) E'" have E'_unfold: "E' = lappend (ltake (enat r_m) E') (LCons (lnth E' r_m) ?E')" unfolding ldropn_Suc_conv_ldropn[OF r_m] by simp hence "mthr.if.mthr.Runs ?start_state (lappend ?r_m_E' (LCons (lnth E' r_m) ?E'))" using \Runs by simp then obtain \' where \_\': "mthr.if.mthr.Trsys ?start_state (list_of ?r_m_E') \'" and \Runs': "mthr.if.mthr.Runs \' (LCons (lnth E' r_m) ?E')" by(rule mthr.if.mthr.Runs_lappendE) simp from \Runs' obtain \''' where red_ra: "mthr.if.redT \' (t_r, ta_r) \'''" and \Runs'': "mthr.if.mthr.Runs \''' ?E'" unfolding E'_r_m by cases let ?vs = "mrw_values P Map.empty (map snd ?start_heap_obs)" { fix a assume "enat a < enat ?r" and "a \ read_actions E" have "a < r" proof(rule ccontr) assume "\ a < r" with \a \ read_actions E\ have "?P a" by simp hence "?r \ a" by(rule Least_le) with \enat a < enat ?r\ show False by simp qed hence "P,E \ a \mrw ws a" using \a \ read_actions E\ by(rule mrw) } with \E \ ?\'\ wf have "ta_seq_consist P Map.empty (lmap snd (ltake (enat ?r) E))" by(rule jmm.ta_seq_consist_mrwI) hence start_sc: "ta_seq_consist P Map.empty (llist_of (map snd ?start_heap_obs))" and "ta_seq_consist P ?vs (lmap snd (ltake (enat (?r - ?n)) E''))" using \?n \ ?r\ unfolding E ltake_lappend lmap_lappend_distrib by(simp_all add: ta_seq_consist_lappend o_def) note this(2) also from r_m have r_m_sum_len_eq: "(\i(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E') i)) = enat (\isnd (lnth E' i)\\<^bsub>o\<^esub>)" - by(subst sum_hom[symmetric, where f=enat])(auto simp add: zero_enat_def split_def less_trans[where y="enat r_m"] intro: sum.cong) + by(subst sum_comp_morphism[symmetric, where h=enat])(auto simp add: zero_enat_def split_def less_trans[where y="enat r_m"] intro: sum.cong) hence "ltake (enat (?r - ?n)) E'' = lappend (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?r_m_E')) (ltake (enat r_n) (ldrop (enat (\isnd (lnth E' i)\\<^bsub>o\<^esub>)) E''))" unfolding ltake_lmap[symmetric] lconcat_ltake r_conv plus_enat_simps(1)[symmetric] ltake_plus_conv_lappend unfolding E' by simp finally have "ta_seq_consist P ?vs (lmap snd (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?r_m_E')))" and sc_ta_r: "ta_seq_consist P (mrw_values P ?vs (map snd (list_of (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?r_m_E'))))) (lmap snd (ltake (enat r_n) (ldropn (\isnd (lnth E' i)\\<^bsub>o\<^esub>) E'')))" unfolding lmap_lappend_distrib by(simp_all add: ta_seq_consist_lappend split_def ldrop_enat) note this(1) also have "lmap snd (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (ltake (enat r_m) E'))) = llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of ?r_m_E')))" unfolding lmap_lconcat llist.map_comp o_def split_def lconcat_llist_of[symmetric] map_map lmap_llist_of[symmetric] by simp finally have "ta_seq_consist P ?vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of ?r_m_E'))))" . from if.sequential_completion[OF cut_and_update ta_seq_consist_convert_RA \_\'[folded mthr.if.RedT_def] this red_ra] obtain ta' ttas' where "mthr.if.mthr.Runs \' (LCons (t_r, ta') ttas')" and sc: "ta_seq_consist P (mrw_values P Map.empty (map snd ?start_heap_obs)) (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) (lappend (llist_of (list_of ?r_m_E')) (LCons (t_r, ta') ttas'))))" and eq_ta: "eq_upto_seq_inconsist P \ta_r\\<^bsub>o\<^esub> \ta'\\<^bsub>o\<^esub> (mrw_values P ?vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of ?r_m_E'))))" by blast let ?E_sc' = "lappend (llist_of (list_of ?r_m_E')) (LCons (t_r, ta') ttas')" let ?E_sc'' = "lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?E_sc')" let ?E_sc = "lappend (llist_of ?start_heap_obs) ?E_sc''" from \_\' \mthr.if.mthr.Runs \' (LCons (t_r, ta') ttas')\ have "mthr.if.mthr.Runs ?start_state ?E_sc'" by(rule mthr.if.mthr.Trsys_into_Runs) hence "?E_sc'' \ mthr.if.\ ?start_state" by(rule mthr.if.\.intros) hence "?E_sc \ ?\" by(rule imageI) moreover from \?E_sc'' \ mthr.if.\ ?start_state\ have tsa_ok: "thread_start_actions_ok ?E_sc" by(rule thread_start_actions_ok_init_fin) from sc have "ta_seq_consist P Map.empty (lmap snd ?E_sc)" by(simp add: lmap_lappend_distrib o_def lmap_lconcat llist.map_comp split_def ta_seq_consist_lappend start_sc) from ta_seq_consist_imp_sequentially_consistent[OF tsa_ok jmm.\_new_actions_for_fun[OF \?E_sc \ ?\\] this] obtain ws_sc where "sequentially_consistent P (?E_sc, ws_sc)" and "P \ (?E_sc, ws_sc) \" unfolding start_heap_obs_def[symmetric] by iprover moreover { have enat_sum_r_m_eq: "enat (\isnd (lnth E' i)\\<^bsub>o\<^esub>) = llength (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?r_m_E'))" - by(auto intro: sum.cong simp add: less_trans[OF _ r_m] lnth_ltake llength_lconcat_lfinite_conv_sum sum_hom[symmetric, where f=enat] zero_enat_def[symmetric] split_beta) + by(auto intro: sum.cong simp add: less_trans[OF _ r_m] lnth_ltake llength_lconcat_lfinite_conv_sum sum_comp_morphism[symmetric, where h=enat] zero_enat_def[symmetric] split_beta) also have "\ \ llength E''" unfolding E' by(blast intro: lprefix_llength_le lprefix_lconcatI lmap_lprefix) finally have r_m_E: "ltake (enat (?n + (\isnd (lnth E' i)\\<^bsub>o\<^esub>))) E = ltake (enat (?n + (\isnd (lnth E' i)\\<^bsub>o\<^esub>))) ?E_sc" by(simp add: ltake_lappend lappend_eq_lappend_conv lmap_lappend_distrib r_m_sum_len_eq ltake_lmap[symmetric] min_def zero_enat_def[symmetric] E E' lconcat_ltake ltake_all del: ltake_lmap) have drop_r_m_E: "ldropn (?n + (\isnd (lnth E' i)\\<^bsub>o\<^esub>)) E = lappend (llist_of (map (Pair t_r) \ta_r\\<^bsub>o\<^esub>)) (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (ldropn (Suc r_m) E')))" (is "_ = ?drop_r_m_E") using E'_r_m unfolding E E' by(subst (2) E'_unfold)(simp add: ldropn_lappend2 lmap_lappend_distrib enat_sum_r_m_eq[symmetric]) have drop_r_m_E_sc: "ldropn (?n + (\isnd (lnth E' i)\\<^bsub>o\<^esub>)) ?E_sc = lappend (llist_of (map (Pair t_r) \ta'\\<^bsub>o\<^esub>)) (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ttas'))" by(simp add: ldropn_lappend2 lmap_lappend_distrib enat_sum_r_m_eq[symmetric]) let ?vs_r_m = "mrw_values P ?vs (map snd (list_of (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?r_m_E'))))" note sc_ta_r also from drop_r_m_E have "ldropn (\isnd (lnth E' i)\\<^bsub>o\<^esub>) E'' = ?drop_r_m_E" unfolding E by(simp add: ldropn_lappend2) also have "lmap snd (ltake (enat r_n) \) = llist_of (take r_n \ta_r\\<^bsub>o\<^esub>)" using r_n by(simp add: ltake_lappend lmap_lappend_distrib ltake_lmap[symmetric] take_map o_def zero_enat_def[symmetric] del: ltake_lmap) finally have sc_ta_r: "ta_seq_consist P ?vs_r_m (llist_of (take r_n \ta_r\\<^bsub>o\<^esub>))" . note eq_ta also have "\ta_r\\<^bsub>o\<^esub> = take r_n \ta_r\\<^bsub>o\<^esub> @ drop r_n \ta_r\\<^bsub>o\<^esub>" by simp finally have "eq_upto_seq_inconsist P (take r_n \ta_r\\<^bsub>o\<^esub> @ drop r_n \ta_r\\<^bsub>o\<^esub>) \ta'\\<^bsub>o\<^esub> ?vs_r_m" by(simp add: list_of_lconcat split_def o_def map_concat) from eq_upto_seq_inconsist_appendD[OF this sc_ta_r] have r_n': "r_n \ length \ta'\\<^bsub>o\<^esub>" and take_r_n_eq: "take r_n \ta'\\<^bsub>o\<^esub> = take r_n \ta_r\\<^bsub>o\<^esub>" and eq_r_n: "eq_upto_seq_inconsist P (drop r_n \ta_r\\<^bsub>o\<^esub>) (drop r_n \ta'\\<^bsub>o\<^esub>) (mrw_values P ?vs_r_m (take r_n \ta_r\\<^bsub>o\<^esub>))" using r_n by(simp_all add: min_def) from r_conv \?n \ ?r\ have r_conv': "?r = (?n + (\isnd (lnth E' i)\\<^bsub>o\<^esub>)) + r_n" by simp from r_n' r_n take_r_n_eq r_m_E drop_r_m_E drop_r_m_E_sc have take_r'_eq: "ltake (enat ?r) E = ltake (enat ?r) ?E_sc" unfolding r_conv' apply(subst (1 2) plus_enat_simps(1)[symmetric]) apply(subst (1 2) ltake_plus_conv_lappend) apply(simp add: lappend_eq_lappend_conv ltake_lappend1 ldrop_enat take_map) done hence take_r_eq: "ltake (enat r) E = ltake (enat r) ?E_sc" by(rule ltake_eq_ltake_antimono)(simp add: \?P ?r\) from eq_r_n Cons_nth_drop_Suc[OF r_n, symmetric] have "drop r_n \ta'\\<^bsub>o\<^esub> \ []" by(auto simp add: eq_upto_seq_inconsist_simps) hence r_n': "r_n < length \ta'\\<^bsub>o\<^esub>" by simp hence eq_r_n: "\ta_r\\<^bsub>o\<^esub> ! r_n \ \ta'\\<^bsub>o\<^esub> ! r_n" using eq_r_n Cons_nth_drop_Suc[OF r_n, symmetric] Cons_nth_drop_Suc[OF r_n', symmetric] by(simp add: eq_upto_seq_inconsist_simps split: action.split_asm obs_event.split_asm if_split_asm) obtain tid_eq: "action_tid E r = action_tid ?E_sc r" and obs_eq: "action_obs E r \ action_obs ?E_sc r" proof(cases "r < ?r") case True { from True have "action_tid E r = action_tid (ltake (enat ?r) E) r" by(simp add: action_tid_def lnth_ltake) also note take_r'_eq also have "action_tid (ltake (enat ?r) ?E_sc) r = action_tid ?E_sc r" using True by(simp add: action_tid_def lnth_ltake) finally have "action_tid E r = action_tid ?E_sc r" . } moreover { from True have "action_obs E r = action_obs (ltake (enat ?r) E) r" by(simp add: action_obs_def lnth_ltake) also note take_r'_eq also have "action_obs (ltake (enat ?r) ?E_sc) r = action_obs ?E_sc r" using True by(simp add: action_obs_def lnth_ltake) finally have "action_obs E r \ action_obs ?E_sc r" by simp } ultimately show thesis by(rule that) next case False with \?P ?r\ have r_eq: "r = ?r" by simp hence "lnth E r = (t_r, \ta_r\\<^bsub>o\<^esub> ! r_n)" using E_r r_conv' E by(simp add: lnth_lappend2) moreover have "lnth ?E_sc r = (t_r, \ta'\\<^bsub>o\<^esub> ! r_n)" using \?n \ ?r\ r_n' by(subst r_eq)(simp add: r_conv lnth_lappend2 lmap_lappend_distrib enat_sum_r_m_eq[symmetric] lnth_lappend1 del: length_lift_start_obs) ultimately have "action_tid E r = action_tid ?E_sc r" "action_obs E r \ action_obs ?E_sc r" using eq_r_n by(simp_all add: action_tid_def action_obs_def) thus thesis by(rule that) qed have "enat r < enat ?n + llength (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (lappend ?r_m_E' (LCons (t_r, ta') LNil))))" using \?P ?r\ r_n' unfolding lmap_lappend_distrib by(simp add: enat_sum_r_m_eq[symmetric] r_conv') also have "llength (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (lappend ?r_m_E' (LCons (t_r, ta') LNil)))) \ llength ?E_sc''" by(rule lprefix_llength_le[OF lprefix_lconcatI])(simp add: lmap_lprefix) finally have "r \ actions ?E_sc" by(simp add: actions_def add_left_mono) note this tid_eq obs_eq take_r_eq } ultimately show ?thesis by blast qed qed(rule \_new_actions_for_unique) qed lemma sc_legal: assumes hb_completion: "if.hb_completion (init_fin_lift_state status (start_state f P C M vs)) (lift_start_obs start_tid start_heap_obs)" (is "if.hb_completion ?start_state ?start_heap_obs") and wf: "wf_syscls P" and wfx_start: "ts_ok wfx (thr (start_state f P C M vs)) start_heap" and ka: "known_addrs start_tid (f (fst (method P C M)) M (fst (snd (method P C M))) (fst (snd (snd (method P C M)))) (the (snd (snd (snd (method P C M))))) vs) \ allocated start_heap" shows "sc_legal (\_start f P C M vs status) P" (is "sc_legal ?\ P") proof - interpret jmm: executions_sc_hb ?\ P using wf wfx_start ka by(rule executions_sc_hb) interpret jmm: executions_aux ?\ P using wf wfx_start ka by(rule executions_aux) show ?thesis proof fix E ws r assume E: "E \ ?\" and wf_exec: "P \ (E, ws) \" and mrw: "\a. \a < r; a \ read_actions E\ \ P,E \ a \mrw ws a" from E obtain E'' where E: "E = lappend (llist_of ?start_heap_obs) E''" and E'': "E'' \ mthr.if.\ ?start_state" by auto from E'' obtain E' where E': "E'' = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E')" and \Runs: "mthr.if.mthr.Runs ?start_state E'" by(rule mthr.if.\.cases) show "\E'\?\. \ws'. P \ (E', ws') \ \ ltake (enat r) E = ltake (enat r) E' \ (\a\read_actions E'. if a < r then ws' a = ws a else P,E' \ ws' a \hb a) \ action_tid E' r = action_tid E r \ (if r \ read_actions E then sim_action else (=)) (action_obs E' r) (action_obs E r) \ (r \ actions E \ r \ actions E')" (is "\E'\?\. \ws'. _ \ ?same E' \ ?read E' ws' \ ?tid E' \ ?obs E' \ ?actions E'") proof(cases "r < length ?start_heap_obs") case True from if.hb_completion_Runs[OF hb_completion ta_hb_consistent_convert_RA] obtain ttas where Runs: "mthr.if.mthr.Runs ?start_state ttas" and hb: "ta_hb_consistent P ?start_heap_obs (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ttas))" by blast from Runs have \: "lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ttas) \ mthr.if.\ ?start_state" by(rule mthr.if.\.intros) let ?E = "lappend (llist_of ?start_heap_obs) (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ttas))" from \ have E': "?E \ ?\" by blast from \ have tsa: "thread_start_actions_ok ?E" by(rule thread_start_actions_ok_init_fin) from start_heap_obs_not_Read have ws: "is_write_seen P (llist_of (lift_start_obs start_tid start_heap_obs)) ws" by(unfold in_set_conv_nth)(rule is_write_seenI, auto simp add: action_obs_def actions_def lift_start_obs_def lnth_LCons elim!: read_actions.cases split: nat.split_asm) with hb tsa have "\ws'. P \ (?E, ws') \ \ (\n. n \ read_actions ?E \ length ?start_heap_obs \ n \ P,?E \ ws' n \hb n) \ (\n_new_actions_for_fun[OF E']) then obtain ws' where wf_exec': "P \ (?E, ws') \" and read_hb: "\n. \ n \ read_actions ?E; length ?start_heap_obs \ n \ \ P,?E \ ws' n \hb n" and same: "\n. n ws' n = ws n" by blast from True have "?same ?E" unfolding E by(simp add: ltake_lappend1) moreover { fix a assume a: "a \ read_actions ?E" have "if a < r then ws' a = ws a else P,?E \ ws' a \hb a" proof(cases "a < length ?start_heap_obs") case True with a have False using start_heap_obs_not_Read by cases(auto simp add: action_obs_def actions_def lnth_lappend1 lift_start_obs_def lnth_LCons in_set_conv_nth split: nat.split_asm) thus ?thesis .. next case False with read_hb[of a] True a show ?thesis by auto qed } hence "?read ?E ws'" by blast moreover from True E have "?tid ?E" by(simp add: action_tid_def lnth_lappend1) moreover from True E have "?obs ?E" by(simp add: action_obs_def lnth_lappend1) moreover from True have "?actions ?E" by(simp add: actions_def enat_less_enat_plusI) ultimately show ?thesis using E' wf_exec' by blast next case False hence r: "length ?start_heap_obs \ r" by simp show ?thesis proof(cases "enat r < llength E") case False then obtain "?same E" "?read E ws" "?tid E" "?obs E" "?actions E" by(cases "llength E")(fastforce elim!: read_actions.cases simp add: actions_def split: if_split_asm)+ with wf_exec \E \ ?\\ show ?thesis by blast next case True note r' = this let ?r = "r - length ?start_heap_obs" from E r r' have "enat ?r < llength E''" by(cases "llength E''")(auto) with \Runs obtain r_m r_n t_r ta_r where E_r: "lnth E'' ?r = (t_r, \ta_r\\<^bsub>o\<^esub> ! r_n)" and r_n: "r_n < length \ta_r\\<^bsub>o\<^esub>" and r_m: "enat r_m < llength E'" and r_conv: "?r = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + r_n" and E'_r_m: "lnth E' r_m = (t_r, ta_r)" unfolding E' by(rule mthr.if.actions_\E_aux) let ?E' = "ldropn (Suc r_m) E'" let ?r_m_E' = "ltake (enat r_m) E'" have E'_unfold: "E' = lappend (ltake (enat r_m) E') (LCons (lnth E' r_m) ?E')" unfolding ldropn_Suc_conv_ldropn[OF r_m] by simp hence "mthr.if.mthr.Runs ?start_state (lappend ?r_m_E' (LCons (lnth E' r_m) ?E'))" using \Runs by simp then obtain \' where \_\': "mthr.if.mthr.Trsys ?start_state (list_of ?r_m_E') \'" and \Runs': "mthr.if.mthr.Runs \' (LCons (lnth E' r_m) ?E')" by(rule mthr.if.mthr.Runs_lappendE) simp from \Runs' obtain \''' where red_ra: "mthr.if.redT \' (t_r, ta_r) \'''" and \Runs'': "mthr.if.mthr.Runs \''' ?E'" unfolding E'_r_m by cases let ?vs = "mrw_values P Map.empty (map snd ?start_heap_obs)" from \E \ ?\\ wf_exec have "ta_seq_consist P Map.empty (lmap snd (ltake (enat r) E))" by(rule jmm.ta_seq_consist_mrwI)(simp add: mrw) hence ns: "non_speculative P (\_. {}) (lmap snd (ltake (enat r) E))" by(rule ta_seq_consist_into_non_speculative) simp also note E also note ltake_lappend2 also note E' also note E'_unfold also note lmap_lappend_distrib also note lmap_lappend_distrib also note lconcat_lappend also note llist.map(2) also note E'_r_m also note prod.simps(2) also note ltake_lappend2 also note lconcat_LCons also note ltake_lappend1 also note non_speculative_lappend also note lmap_lappend_distrib also note non_speculative_lappend also have "lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (ltake (enat r_m) E')) = llist_of (concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E'))))" by(simp add: lconcat_llist_of[symmetric] lmap_llist_of[symmetric] llist.map_comp o_def split_def del: lmap_llist_of) ultimately have "non_speculative P (\_. {}) (lmap snd (llist_of ?start_heap_obs))" and "non_speculative P (w_values P (\_. {}) (map snd ?start_heap_obs)) (lmap snd (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (ltake (enat r_m) E'))))" and ns': "non_speculative P (w_values P (w_values P (\_. {}) (map snd ?start_heap_obs)) (map snd (concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E')))))) (lmap snd (ltake (enat r_n) (llist_of (map (Pair t_r) \ta_r\\<^bsub>o\<^esub>))))" using r r_conv r_m r_n by(simp_all add: length_concat o_def split_def sum_list_sum_nth length_list_of_conv_the_enat less_min_eq1 atLeast0LessThan lnth_ltake split: if_split_asm cong: sum.cong_simp) hence ns: "non_speculative P (w_values P (\_. {}) (map snd ?start_heap_obs)) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E')))))" unfolding lconcat_llist_of[symmetric] lmap_lconcat lmap_llist_of[symmetric] llist.map_comp o_def split_def by(simp) from ns' have ns': "non_speculative P (w_values P (w_values P (\_. {}) (map snd ?start_heap_obs)) (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E'))))) (llist_of (take r_n \ta_r\\<^bsub>o\<^esub>))" unfolding map_concat map_map by(simp add: take_map[symmetric] o_def split_def) let ?hb = "\ta'_r :: ('addr, 'thread_id, status \ 'x, 'heap, 'addr, ('addr, 'thread_id) obs_event action) thread_action. ta_hb_consistent P (?start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E'))) @ map (Pair t_r) (take r_n \ta_r\\<^bsub>o\<^esub>)) (llist_of (map (Pair t_r) (drop r_n \ta'_r\\<^bsub>o\<^esub>)))" let ?sim = "\ta'_r. (if \ad al v. \ta_r\\<^bsub>o\<^esub> ! r_n = NormalAction (ReadMem ad al v) then sim_action else (=)) (\ta_r\\<^bsub>o\<^esub> ! r_n) (\ta'_r\\<^bsub>o\<^esub> ! r_n)" from red_ra obtain ta'_r \'''' where red_ra': "mthr.if.redT \' (t_r, ta'_r) \''''" and eq: "take r_n \ta'_r\\<^bsub>o\<^esub> = take r_n \ta_r\\<^bsub>o\<^esub>" and hb: "?hb ta'_r" and r_n': "r_n < length \ta'_r\\<^bsub>o\<^esub>" and sim: "?sim ta'_r" proof(cases) case (redT_normal x x' m') note tst = \thr \' t_r = \(x, no_wait_locks)\\ and red = \t_r \ (x, shr \') -ta_r\i (x', m')\ and aok = \mthr.if.actions_ok \' t_r ta_r\ and \''' = \redT_upd \' t_r ta_r x' m' \'''\ from if.hb_completionD[OF hb_completion \_\'[folded mthr.if.RedT_def] ns tst red aok ns'] r_n obtain ta'_r x'' m'' where red': "t_r \ (x, shr \') -ta'_r\i (x'', m'')" and aok': "mthr.if.actions_ok \' t_r ta'_r" and eq': "take r_n \ta'_r\\<^bsub>o\<^esub> = take r_n \ta_r\\<^bsub>o\<^esub>" and hb: "?hb ta'_r" and r_n': "r_n < length \ta'_r\\<^bsub>o\<^esub>" and sim: "?sim ta'_r" by blast from redT_updWs_total[of t_r "wset \'" "\ta'_r\\<^bsub>w\<^esub>"] obtain \'''' where "redT_upd \' t_r ta'_r x'' m'' \''''" by fastforce with red' tst aok' have "mthr.if.redT \' (t_r, ta'_r) \''''" .. thus thesis using eq' hb r_n' sim by(rule that) next case (redT_acquire x n ln) hence "?hb ta_r" using set_convert_RA_not_Read[where ln=ln] by -(rule ta_hb_consistent_not_ReadI, fastforce simp del: set_convert_RA_not_Read dest!: in_set_dropD) with red_ra r_n show ?thesis by(auto intro: that) qed from hb have "non_speculative P (w_values P (\_. {}) (map snd (?start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E'))) @ map (Pair t_r) (take r_n \ta_r\\<^bsub>o\<^esub>)))) (lmap snd (llist_of (map (Pair t_r) (drop r_n \ta'_r\\<^bsub>o\<^esub>))))" by(rule ta_hb_consistent_into_non_speculative) with ns' eq[symmetric] have "non_speculative P (w_values P (\_. {}) (map snd (?start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E')))))) (llist_of (map snd (map (Pair t_r) \ta'_r\\<^bsub>o\<^esub>)))" by(subst append_take_drop_id[where xs="\ta'_r\\<^bsub>o\<^esub>" and n=r_n, symmetric])(simp add: o_def map_concat split_def lappend_llist_of_llist_of[symmetric] non_speculative_lappend del: append_take_drop_id lappend_llist_of_llist_of) with ns have ns'': "non_speculative P (w_values P (\_. {}) (map snd ?start_heap_obs)) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E') @ [(t_r, ta'_r)]))))" unfolding lconcat_llist_of[symmetric] map_append lappend_llist_of_llist_of[symmetric] lmap_llist_of[symmetric] llist.map_comp by(simp add: o_def split_def non_speculative_lappend list_of_lconcat map_concat) from \_\' red_ra' have "mthr.if.RedT ?start_state (list_of ?r_m_E' @ [(t_r, ta'_r)]) \''''" unfolding mthr.if.RedT_def .. with hb_completion have hb_completion': "if.hb_completion \'''' (?start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E') @ [(t_r, ta'_r)])))" using ns'' by(rule if.hb_completion_shift) from if.hb_completion_Runs[OF hb_completion' ta_hb_consistent_convert_RA] obtain ttas' where Runs': "mthr.if.mthr.Runs \'''' ttas'" and hb': "ta_hb_consistent P (?start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E') @ [(t_r, ta'_r)]))) (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ttas'))" by blast let ?E = "lappend (llist_of ?start_heap_obs) (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (lappend (ltake (enat r_m) E') (LCons (t_r, ta'_r) ttas'))))" have \: "lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (lappend (ltake (enat r_m) E') (LCons (t_r, ta'_r) ttas'))) \ mthr.if.\ ?start_state" by(subst (4) llist_of_list_of[symmetric])(simp, blast intro: mthr.if.\.intros mthr.if.mthr.Trsys_into_Runs \_\' mthr.if.mthr.Runs.Step red_ra' Runs') hence \': "?E \ ?\" by blast from \ have tsa: "thread_start_actions_ok ?E" by(rule thread_start_actions_ok_init_fin) also let ?E' = "lappend (llist_of (lift_start_obs start_tid start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E'))) @ map (Pair t_r) (take r_n \ta_r\\<^bsub>o\<^esub>))) (lappend (llist_of (map (Pair t_r) (drop r_n \ta'_r\\<^bsub>o\<^esub>))) (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ttas')))" have "?E = ?E'" using eq[symmetric] by(simp add: lmap_lappend_distrib lappend_assoc lappend_llist_of_llist_of[symmetric] lconcat_llist_of[symmetric] lmap_llist_of[symmetric] llist.map_comp o_def split_def del: lmap_llist_of)(simp add: lappend_assoc[symmetric] lmap_lappend_distrib[symmetric] map_append[symmetric] lappend_llist_of_llist_of del: map_append) finally have tsa': "thread_start_actions_ok ?E'" . from hb hb' eq[symmetric] have HB: "ta_hb_consistent P (?start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E'))) @ map (Pair t_r) (take r_n \ta_r\\<^bsub>o\<^esub>)) (lappend (llist_of (map (Pair t_r) (drop r_n \ta'_r\\<^bsub>o\<^esub>))) (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ttas')))" by -(rule ta_hb_consistent_lappendI, simp_all add: take_map[symmetric] drop_map[symmetric]) define EE where "EE = llist_of (?start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E'))) @ map (Pair t_r) (take r_n \ta_r\\<^bsub>o\<^esub>))" from r r_conv have r_conv': "r = (\isnd (lnth E' i)\\<^bsub>o\<^esub>) + r_n + length ?start_heap_obs" by auto hence len_EE: "llength EE = enat r" using r_m r_n by(auto simp add: EE_def length_concat sum_list_sum_nth atLeast0LessThan lnth_ltake less_min_eq1 split_def min_def length_list_of_conv_the_enat cong: sum.cong_simp) from r_conv r_m have r_conv3: "llength (lconcat (lmap (\x. llist_of (map (Pair (fst x)) \snd x\\<^bsub>o\<^esub>)) (ltake (enat r_m) E'))) = enat (r - Suc (length start_heap_obs) - r_n)" apply(simp add: llength_lconcat_lfinite_conv_sum lnth_ltake cong: sum.cong_simp conj_cong) - apply(auto simp add: sum_hom[where f=enat, symmetric] zero_enat_def less_trans[where y="enat r_m"] intro: sum.cong) + apply(auto simp add: sum_comp_morphism[where h=enat, symmetric] zero_enat_def less_trans[where y="enat r_m"] intro: sum.cong) done have is_ws: "is_write_seen P EE ws" proof(rule is_write_seenI) fix a ad al v assume a: "a \ read_actions EE" and a_obs: "action_obs EE a = NormalAction (ReadMem ad al v)" from a have a_r: "a < r" by cases(simp add: len_EE actions_def) from r E'_r_m r_m r_n r_conv3 have eq: "ltake (enat r) EE = ltake (enat r) E" unfolding E E' EE_def apply(subst (2) E'_unfold) apply(simp add: ltake_lappend2 lappend_llist_of_llist_of[symmetric] lappend_eq_lappend_conv lmap_lappend_distrib lconcat_llist_of[symmetric] o_def split_def lmap_llist_of[symmetric] del: lappend_llist_of_llist_of lmap_llist_of) apply(subst ltake_lappend1) defer apply(simp add: ltake_lmap[symmetric] take_map[symmetric] ltake_llist_of[symmetric] del: ltake_lmap ltake_llist_of) apply(auto simp add: min_def) done hence sim: "ltake (enat r) EE [\] ltake (enat r) E" by(rule eq_into_sim_actions) from a sim have a': "a \ read_actions E" by(rule read_actions_change_prefix)(simp add: a_r) from action_obs_change_prefix_eq[OF eq, of a] a_r a_obs have a_obs': "action_obs E a = NormalAction (ReadMem ad al v)" by simp have a_mrw: "P,E \ a \mrw ws a" using a_r a' by(rule mrw) with \E \ ?\\ wf_exec have ws_a_a: "ws a < a" by(rule jmm.mrw_before)(auto intro: a_r less_trans mrw) hence [simp]: "ws a < r" using a_r by simp from wf_exec have ws: "is_write_seen P E ws" by(rule wf_exec_is_write_seenD) from is_write_seenD[OF this a' a_obs'] have "ws a \ write_actions E" and "(ad, al) \ action_loc P E (ws a)" and "value_written P E (ws a) (ad, al) = v" and "\ P,E \ a \hb ws a" and "is_volatile P al \ \ P,E \ a \so ws a" and between: "\a'. \ a' \ write_actions E; (ad, al) \ action_loc P E a'; P,E \ ws a \hb a' \ P,E \ a' \hb a \ is_volatile P al \ P,E \ ws a \so a' \ P,E \ a' \so a \ \ a' = ws a" by simp_all from \ws a \ write_actions E\ sim[symmetric] show "ws a \ write_actions EE" by(rule write_actions_change_prefix) simp from action_loc_change_prefix[OF sim, of "ws a" P] \(ad, al) \ action_loc P E (ws a)\ show "(ad, al) \ action_loc P EE (ws a)" by(simp) from value_written_change_prefix[OF eq, of "ws a" P] \value_written P E (ws a) (ad, al) = v\ show "value_written P EE (ws a) (ad, al) = v" by simp from wf_exec have tsa_E: "thread_start_actions_ok E" by(rule wf_exec_thread_start_actions_okD) from \\ P,E \ a \hb ws a\ show "\ P,EE \ a \hb ws a" proof(rule contrapos_nn) assume "P,EE \ a \hb ws a" thus "P,E \ a \hb ws a" using tsa_E sim by(rule happens_before_change_prefix)(simp_all add: a_r) qed { assume "is_volatile P al" hence "\ P,E \ a \so ws a" by fact thus "\ P,EE \ a \so ws a" by(rule contrapos_nn)(rule sync_order_change_prefix[OF _ sim], simp_all add: a_r) } fix a' assume "a' \ write_actions EE" "(ad, al) \ action_loc P EE a'" moreover hence [simp]: "a' < r" by cases(simp add: actions_def len_EE) ultimately have a': "a' \ write_actions E" "(ad, al) \ action_loc P E a'" using sim action_loc_change_prefix[OF sim, of a' P] by(auto intro: write_actions_change_prefix) { assume "P,EE \ ws a \hb a'" "P,EE \ a' \hb a" hence "P,E \ ws a \hb a'" "P,E \ a' \hb a" using tsa_E sim a_r by(auto elim!: happens_before_change_prefix) with between[OF a'] show "a' = ws a" by simp } { assume "is_volatile P al " "P,EE \ ws a \so a'" "P,EE \ a' \so a" with sim a_r between[OF a'] show "a' = ws a" by(fastforce elim: sync_order_change_prefix intro!: disjI2 del: disjCI) } qed with HB tsa' have "\ws'. P \ (?E', ws') \ \ (\n. n \ read_actions ?E' \ length (?start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E'))) @ map (Pair t_r) (take r_n \ta_r\\<^bsub>o\<^esub>)) \ n \ P,?E' \ ws' n \hb n) \ (\n(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E'))) @ map (Pair t_r) (take r_n \ta_r\\<^bsub>o\<^esub>)). ws' n = ws n)" unfolding EE_def by(rule ta_hb_consistent_Read_hb)(rule jmm.\_new_actions_for_fun[OF \'[unfolded \?E = ?E'\]]) also have r_conv'': "length (?start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (list_of (ltake (enat r_m) E'))) @ map (Pair t_r) (take r_n \ta_r\\<^bsub>o\<^esub>)) = r" using r_n r_m unfolding r_conv' by(auto simp add: length_concat sum_list_sum_nth atLeast0LessThan lnth_ltake split_def o_def less_min_eq1 min_def length_list_of_conv_the_enat cong: sum.cong_simp) finally obtain ws' where wf_exec': "P \ (?E', ws') \" and read_hb: "\n. \ n \ read_actions ?E'; r \ n \ \ P,?E' \ ws' n \hb n" and read_same: "\n. n < r \ ws' n = ws n" by blast have "?same ?E'" apply(subst ltake_lappend1, simp add: r_conv''[symmetric] length_list_of_conv_the_enat) unfolding E E' lappend_llist_of_llist_of[symmetric] apply(subst (1 2) ltake_lappend2, simp add: r[simplified]) apply(subst lappend_eq_lappend_conv, simp) apply safe apply(subst E'_unfold) unfolding lmap_lappend_distrib apply(subst lconcat_lappend, simp) apply(subst lconcat_llist_of[symmetric]) apply(subst (3) lmap_llist_of[symmetric]) apply(subst (3) lmap_llist_of[symmetric]) apply(subst llist.map_comp) apply(simp only: split_def o_def) apply(subst llist_of_list_of, simp) apply(subst (1 2) ltake_lappend2, simp add: r_conv3) apply(subst lappend_eq_lappend_conv, simp) apply safe unfolding llist.map(2) lconcat_LCons E'_r_m snd_conv fst_conv take_map apply(subst ltake_lappend1) defer apply(subst append_take_drop_id[where xs="\ta_r\\<^bsub>o\<^esub>" and n=r_n, symmetric]) unfolding map_append lappend_llist_of_llist_of[symmetric] apply(subst ltake_lappend1) using r_n apply(simp add: min_def r_conv3) apply(rule refl) apply(simp add: r_conv3) using r_n by arith moreover { fix a assume "a \ read_actions ?E'" with read_hb[of a] read_same[of a] have "if a < r then ws' a = ws a else P,?E' \ ws' a \hb a" by simp } hence "?read ?E' ws'" by blast moreover from r_m r_n r_n' have E'_r: "lnth ?E' r = (t_r, \ta'_r\\<^bsub>o\<^esub> ! r_n)" unfolding r_conv' by(auto simp add: lnth_lappend nth_append length_concat sum_list_sum_nth atLeast0LessThan split_beta lnth_ltake less_min_eq1 length_list_of_conv_the_enat cong: sum.cong_simp) from E_r r have E_r: "lnth E r = (t_r, \ta_r\\<^bsub>o\<^esub> ! r_n)" unfolding E by(simp add: lnth_lappend) have "r \ read_actions E \ (\ad al v. \ta_r\\<^bsub>o\<^esub> ! r_n = NormalAction (ReadMem ad al v))" using True by(auto elim!: read_actions.cases simp add: action_obs_def E_r actions_def intro!: read_actions.intros) with sim E'_r E_r have "?tid ?E'" "?obs ?E'" by(auto simp add: action_tid_def action_obs_def) moreover have "?actions ?E'" using r_n r_m r_n' unfolding r_conv' by(cases "llength ?E'")(auto simp add: actions_def less_min_eq2 length_concat sum_list_sum_nth atLeast0LessThan split_beta lnth_ltake less_min_eq1 length_list_of_conv_the_enat enat_plus_eq_enat_conv cong: sum.cong_simp) ultimately show ?thesis using wf_exec' \' unfolding \?E = ?E'\ by blast qed qed qed qed end lemma w_value_mrw_value_conf: assumes "set_option (vs' adal) \ vs adal \ UNIV" shows "set_option (mrw_value P vs' ob adal) \ w_value P vs ob adal \ UNIV" using assms by(cases adal)(cases ob rule: w_value_cases, auto) lemma w_values_mrw_values_conf: assumes "set_option (vs' adal) \ vs adal \ UNIV" shows "set_option (mrw_values P vs' obs adal) \ w_values P vs obs adal \ UNIV" using assms by(induct obs arbitrary: vs' vs)(auto del: subsetI intro: w_value_mrw_value_conf) lemma w_value_mrw_value_dom_eq_preserve: assumes "dom vs' = {adal. vs adal \ {}}" shows "dom (mrw_value P vs' ob) = {adal. w_value P vs ob adal \ {}}" using assms apply(cases ob rule: w_value_cases) apply(simp_all add: dom_def split_beta del: not_None_eq) apply(blast elim: equalityE dest: subsetD)+ done lemma w_values_mrw_values_dom_eq_preserve: assumes "dom vs' = {adal. vs adal \ {}}" shows "dom (mrw_values P vs' obs) = {adal. w_values P vs obs adal \ {}}" using assms by(induct obs arbitrary: vs vs')(auto del: equalityI intro: w_value_mrw_value_dom_eq_preserve) context jmm_multithreaded begin definition non_speculative_read :: "nat \ ('l, 'thread_id, 'x, 'm, 'w) state \ ('addr \ addr_loc \ 'addr val set) \ bool" where "non_speculative_read n s vs \ (\ttas s' t x ta x' m' i ad al v v'. s -\ttas\* s' \ non_speculative P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) \ thr s' t = \(x, no_wait_locks)\ \ t \ (x, shr s') -ta\ (x', m') \ actions_ok s' t ta \ i < length \ta\\<^bsub>o\<^esub> \ non_speculative P (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of (take i \ta\\<^bsub>o\<^esub>)) \ \ta\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v) \ v' \ w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas) @ take i \ta\\<^bsub>o\<^esub>) (ad, al) \ (\ta' x'' m''. t \ (x, shr s') -ta'\ (x'', m'') \ actions_ok s' t ta' \ i < length \ta'\\<^bsub>o\<^esub> \ take i \ta'\\<^bsub>o\<^esub> = take i \ta\\<^bsub>o\<^esub> \ \ta'\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v') \ length \ta'\\<^bsub>o\<^esub> \ max n (length \ta\\<^bsub>o\<^esub>)))" lemma non_speculative_readI [intro?]: "(\ttas s' t x ta x' m' i ad al v v'. \ s -\ttas\* s'; non_speculative P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))); thr s' t = \(x, no_wait_locks)\; t \ (x, shr s') -ta\ (x', m'); actions_ok s' t ta; i < length \ta\\<^bsub>o\<^esub>; non_speculative P (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of (take i \ta\\<^bsub>o\<^esub>)); \ta\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v); v' \ w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas) @ take i \ta\\<^bsub>o\<^esub>) (ad, al) \ \ \ta' x'' m''. t \ (x, shr s') -ta'\ (x'', m'') \ actions_ok s' t ta' \ i < length \ta'\\<^bsub>o\<^esub> \ take i \ta'\\<^bsub>o\<^esub> = take i \ta\\<^bsub>o\<^esub> \ \ta'\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v') \ length \ta'\\<^bsub>o\<^esub> \ max n (length \ta\\<^bsub>o\<^esub>)) \ non_speculative_read n s vs" unfolding non_speculative_read_def by blast lemma non_speculative_readD: "\ non_speculative_read n s vs; s -\ttas\* s'; non_speculative P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))); thr s' t = \(x, no_wait_locks)\; t \ (x, shr s') -ta\ (x', m'); actions_ok s' t ta; i < length \ta\\<^bsub>o\<^esub>; non_speculative P (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of (take i \ta\\<^bsub>o\<^esub>)); \ta\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v); v' \ w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas) @ take i \ta\\<^bsub>o\<^esub>) (ad, al) \ \ \ta' x'' m''. t \ (x, shr s') -ta'\ (x'', m'') \ actions_ok s' t ta' \ i < length \ta'\\<^bsub>o\<^esub> \ take i \ta'\\<^bsub>o\<^esub> = take i \ta\\<^bsub>o\<^esub> \ \ta'\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v') \ length \ta'\\<^bsub>o\<^esub> \ max n (length \ta\\<^bsub>o\<^esub>)" unfolding non_speculative_read_def by blast end subsection \@{term "non_speculative"} generalises @{term "cut_and_update"} and @{term "ta_hb_consistent"}\ context known_addrs_typing begin lemma read_non_speculative_new_actions_for: fixes status f C M params E defines "E \ lift_start_obs start_tid start_heap_obs" and "vs \ w_values P (\_. {}) (map snd E)" and "s \ init_fin_lift_state status (start_state f P C M params)" assumes wf: "wf_syscls P" and RedT: "mthr.if.RedT s ttas s'" and redT: "mthr.if.redT s' (t, ta') s''" and read: "NormalAction (ReadMem ad al v) \ set \ta'\\<^bsub>o\<^esub>" and ns: "non_speculative P (\_. {}) (llist_of (map snd E @ concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" and ka: "known_addrs start_tid (f (fst (method P C M)) M (fst (snd (method P C M))) (fst (snd (snd (method P C M)))) (the (snd (snd (snd (method P C M))))) params) \ allocated start_heap" and wt: "ts_ok (init_fin_lift wfx) (thr s) (shr s)" and type_adal: "P,shr s' \ ad@al : T" shows "\w. w \ new_actions_for P (llist_of (E @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) ttas))) (ad, al)" (is "\w. ?new_w w") using RedT redT read ns[unfolded E_def] ka unfolding s_def proof(cases rule: read_ex_NewHeapElem) case (start CTn) then obtain n where n: "start_heap_obs ! n = NewHeapElem ad CTn" and len: "n < length start_heap_obs" unfolding in_set_conv_nth by blast from ns have "non_speculative P (w_values P (\_. {}) (map snd E)) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" unfolding lappend_llist_of_llist_of[symmetric] by(simp add: non_speculative_lappend del: lappend_llist_of_llist_of) with RedT wt have hext: "start_heap \ shr s'" unfolding s_def E_def using start_state_vs_conf[OF wf] by(auto dest!: init_fin_RedT_hext_incr simp add: start_state_def split_beta init_fin_lift_state_conv_simps) from start have "typeof_addr start_heap ad = \CTn\" by(auto dest: NewHeapElem_start_heap_obsD[OF wf]) with hext have "typeof_addr (shr s') ad = \CTn\" by(rule typeof_addr_hext_mono) with type_adal have "(ad, al) \ action_loc_aux P (NormalAction (NewHeapElem ad CTn))" using n len by cases (auto simp add: action_obs_def lnth_lappend1 lift_start_obs_def) with n len have "?new_w (Suc n)" by(simp add: new_actions_for_def actions_def E_def action_obs_def lift_start_obs_def nth_append) thus ?thesis .. next case (Red ttas' s'' t' ta' s''' ttas'' CTn) note ttas = \ttas = ttas' @ (t', ta') # ttas''\ from \NormalAction (NewHeapElem ad CTn) \ set \ta'\\<^bsub>o\<^esub>\ obtain obs obs' where obs: "\ta'\\<^bsub>o\<^esub> = obs @ NormalAction (NewHeapElem ad CTn) # obs'" by(auto dest: split_list) let ?n = "length (lift_start_obs start_tid start_heap_obs)" let ?wa = "?n + length (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')) + length obs" have "?wa = ?n + length (concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) ttas')) + length obs" by(simp add: length_concat o_def split_def) also have "\ < length (E @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) ttas))" using obs ttas by(simp add: E_def) also from ttas obs have "(E @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) ttas)) ! ?wa = (t', NormalAction (NewHeapElem ad CTn))" by(auto simp add: E_def lift_start_obs_def nth_append o_def split_def length_concat) moreover from \mthr.if.redT s'' (t', ta') s'''\ \NormalAction (NewHeapElem ad CTn) \ set \ta'\\<^bsub>o\<^esub>\ obtain x_wa x_wa' where ts''t': "thr s'' t' = \(x_wa, no_wait_locks)\" and red_wa: "mthr.init_fin t' (x_wa, shr s'') ta' (x_wa', shr s''')" by(cases) fastforce+ from start_state_vs_conf[OF wf] have vs: "vs_conf P (shr s) vs" unfolding vs_def E_def s_def by(simp add: init_fin_lift_state_conv_simps start_state_def split_def) from ns have ns: "non_speculative P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')))" and ns': "non_speculative P (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas'))) (llist_of \ta'\\<^bsub>o\<^esub>)" and ns'': "non_speculative P (w_values P (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas'))) \ta'\\<^bsub>o\<^esub>) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas'')))" unfolding ttas vs_def by(simp_all add: lappend_llist_of_llist_of[symmetric] non_speculative_lappend del: lappend_llist_of_llist_of) from \mthr.if.RedT (init_fin_lift_state status (start_state f P C M params)) ttas' s''\ wt ns have ts_ok'': "ts_ok (init_fin_lift wfx) (thr s'') (shr s'')" using vs unfolding vs_def s_def by(rule if_RedT_non_speculative_invar) with ts''t' have wfxt': "wfx t' (snd x_wa) (shr s'')" by(cases x_wa)(auto dest: ts_okD) from \mthr.if.RedT (init_fin_lift_state status (start_state f P C M params)) ttas' s''\ wt ns have vs'': "vs_conf P (shr s'') (w_values P (w_values P (\_. {}) (map snd E)) (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')))" unfolding s_def E_def vs_def by(rule if_RedT_non_speculative_invar)(simp add: start_state_def split_beta init_fin_lift_state_conv_simps start_state_vs_conf[OF wf]) from if_redT_non_speculative_vs_conf[OF \mthr.if.redT s'' (t', ta') s'''\ ts_ok'' _ vs'', of "length \ta'\\<^bsub>o\<^esub>"] ns' have vs''': "vs_conf P (shr s''') (w_values P (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas'))) \ta'\\<^bsub>o\<^esub>)" by(simp add: vs_def) from \mthr.if.redT s'' (t', ta') s'''\ ts_ok'' ns' vs'' have "ts_ok (init_fin_lift wfx) (thr s''') (shr s''')" unfolding vs_def by(rule if_redT_non_speculative_invar) with \mthr.if.RedT s''' ttas'' s'\ have hext: "shr s''' \ shr s'" using ns'' vs''' by(rule init_fin_RedT_hext_incr) from red_wa have "typeof_addr (shr s''') ad = \CTn\" using wfxt' \NormalAction (NewHeapElem ad CTn) \ set \ta'\\<^bsub>o\<^esub>\ by cases(auto dest: red_NewHeapElemD) with hext have "typeof_addr (shr s') ad = \CTn\" by(rule typeof_addr_hext_mono) with type_adal have "(ad, al) \ action_loc_aux P (NormalAction (NewHeapElem ad CTn))" by cases auto ultimately have "?new_w ?wa" by(simp add: new_actions_for_def actions_def action_obs_def) thus ?thesis .. qed lemma non_speculative_read_into_cut_and_update: fixes status f C M params E defines "E \ lift_start_obs start_tid start_heap_obs" and "vs \ w_values P (\_. {}) (map snd E)" and "s \ init_fin_lift_state status (start_state f P C M params)" and "vs' \ mrw_values P Map.empty (map snd E)" assumes wf: "wf_syscls P" and nsr: "if.non_speculative_read n s vs" and wt: "ts_ok (init_fin_lift wfx) (thr s) (shr s)" and ka: "known_addrs start_tid (f (fst (method P C M)) M (fst (snd (method P C M))) (fst (snd (snd (method P C M)))) (the (snd (snd (snd (method P C M))))) params) \ allocated start_heap" shows "if.cut_and_update s vs'" proof(rule if.cut_and_updateI) fix ttas s' t x ta x' m' assume Red: "mthr.if.RedT s ttas s'" and sc: "ta_seq_consist P vs' (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" and tst: "thr s' t = \(x, no_wait_locks)\" and red: "t \ (x, shr s') -ta\i (x', m')" and aok: "mthr.if.actions_ok s' t ta" let ?vs = "w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))" let ?vs' = "mrw_values P vs' (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))" from start_state_vs_conf[OF wf] have vs: "vs_conf P (shr s) vs" unfolding vs_def E_def s_def by(simp add: init_fin_lift_state_conv_simps start_state_def split_def) from sc have ns: "non_speculative P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" by(rule ta_seq_consist_into_non_speculative)(auto simp add: vs'_def vs_def del: subsetI intro: w_values_mrw_values_conf) from ns have ns': "non_speculative P (\_. {}) (llist_of (map snd (lift_start_obs start_tid start_heap_obs) @ concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" unfolding lappend_llist_of_llist_of[symmetric] vs_def by(simp add: non_speculative_lappend E_def non_speculative_start_heap_obs del: lappend_llist_of_llist_of) have vs_vs'': "\adal. set_option (?vs' adal) \ ?vs adal \ UNIV" by(rule w_values_mrw_values_conf)(auto simp add: vs'_def vs_def del: subsetI intro: w_values_mrw_values_conf) from Red wt ns vs have wt': "ts_ok (init_fin_lift wfx) (thr s') (shr s')" by(rule if_RedT_non_speculative_invar) hence wtt: "init_fin_lift wfx t x (shr s')" using tst by(rule ts_okD) { fix i have "\ta' x'' m''. t \ (x, shr s') -ta'\i (x'', m'') \ mthr.if.actions_ok s' t ta' \ length \ta'\\<^bsub>o\<^esub> \ max n (length \ta\\<^bsub>o\<^esub>) \ ta_seq_consist P ?vs' (llist_of (take i \ta'\\<^bsub>o\<^esub>)) \ eq_upto_seq_inconsist P (take i \ta\\<^bsub>o\<^esub>) (take i \ta'\\<^bsub>o\<^esub>) ?vs' \ (ta_seq_consist P ?vs' (llist_of (take i \ta\\<^bsub>o\<^esub>)) \ ta' = ta)" proof(induct i) case 0 show ?case using red aok by(auto simp del: split_paired_Ex simp add: eq_upto_seq_inconsist_simps) next case (Suc i) then obtain ta' x'' m'' where red': "t \ (x, shr s') -ta'\i (x'', m'')" and aok': "mthr.if.actions_ok s' t ta'" and len: "length \ta'\\<^bsub>o\<^esub> \ max n (length \ta\\<^bsub>o\<^esub>)" and sc': "ta_seq_consist P ?vs' (llist_of (take i \ta'\\<^bsub>o\<^esub>))" and eusi: "eq_upto_seq_inconsist P (take i \ta\\<^bsub>o\<^esub>) (take i \ta'\\<^bsub>o\<^esub>) ?vs'" and ta'_ta: "ta_seq_consist P ?vs' (llist_of (take i \ta\\<^bsub>o\<^esub>)) \ ta' = ta" by blast let ?vs'' = "mrw_values P ?vs' (take i \ta'\\<^bsub>o\<^esub>)" show ?case proof(cases "i < length \ta'\\<^bsub>o\<^esub> \ \ ta_seq_consist P ?vs' (llist_of (take (Suc i) \ta'\\<^bsub>o\<^esub>)) \ \ ta_seq_consist P ?vs' (llist_of (take (Suc i) \ta\\<^bsub>o\<^esub>))") case True hence i: "i < length \ta'\\<^bsub>o\<^esub>" and "\ ta_seq_consist P ?vs'' (LCons (\ta'\\<^bsub>o\<^esub> ! i) LNil)" using sc' by(auto simp add: take_Suc_conv_app_nth lappend_llist_of_llist_of[symmetric] ta_seq_consist_lappend simp del: lappend_llist_of_llist_of) then obtain ad al v where ta'_i: "\ta'\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v)" by(auto split: action.split_asm obs_event.split_asm) from ta'_i True have read: "NormalAction (ReadMem ad al v) \ set \ta'\\<^bsub>o\<^esub>" by(auto simp add: in_set_conv_nth) with red' have "ad \ known_addrs_if t x" by(rule if_red_read_knows_addr) hence "ad \ if.known_addrs_state s'" using tst by(rule if.known_addrs_stateI) moreover from init_fin_red_read_typeable[OF red' wtt read] obtain T where type_adal: "P,shr s' \ ad@al : T" .. from redT_updWs_total[of t "wset s'" "\ta'\\<^bsub>w\<^esub>"] red' tst aok' obtain s'' where redT': "mthr.if.redT s' (t, ta') s''" by(auto dest!: mthr.if.redT.redT_normal) with wf Red have "\w. w \ new_actions_for P (llist_of (E @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) ttas))) (ad, al)" (is "\w. ?new_w w") using read ns' ka wt type_adal unfolding s_def E_def by(rule read_non_speculative_new_actions_for) then obtain w where w: "?new_w w" .. have "(ad, al) \ dom ?vs'" proof(cases "w < length E") case True with w have "(ad, al) \ dom vs'" unfolding vs'_def new_actions_for_def by(clarsimp)(erule mrw_values_new_actionD[rotated 1], auto simp del: split_paired_Ex simp add: set_conv_nth action_obs_def nth_append intro!: exI[where x=w]) also have "dom vs' \ dom ?vs'" by(rule mrw_values_dom_mono) finally show ?thesis . next case False with w show ?thesis unfolding new_actions_for_def apply(clarsimp) apply(erule mrw_values_new_actionD[rotated 1]) apply(simp_all add: set_conv_nth action_obs_def nth_append actions_def) apply(rule exI[where x="w - length E"]) apply(subst nth_map[where f=snd, symmetric]) apply(simp_all add: length_concat o_def split_def map_concat) done qed hence "(ad, al) \ dom (mrw_values P ?vs' (take i \ta'\\<^bsub>o\<^esub>))" by(rule subsetD[OF mrw_values_dom_mono]) then obtain v' b where v': "mrw_values P ?vs' (take i \ta'\\<^bsub>o\<^esub>) (ad, al) = \(v', b)\" by auto moreover from vs_vs''[of "(ad, al)"] have "set_option (mrw_values P ?vs' (take i \ta'\\<^bsub>o\<^esub>) (ad, al)) \ w_values P ?vs (take i \ta'\\<^bsub>o\<^esub>) (ad, al) \ UNIV" by(rule w_values_mrw_values_conf) ultimately have "v' \ w_values P ?vs (take i \ta'\\<^bsub>o\<^esub>) (ad, al)" by simp moreover from sc' have "non_speculative P (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of (take i \ta'\\<^bsub>o\<^esub>))" by(blast intro: ta_seq_consist_into_non_speculative vs_vs'' del: subsetI) ultimately obtain ta'' x'' m'' where red'': "t \ (x, shr s') -ta''\i (x'', m'')" and aok'': "mthr.if.actions_ok s' t ta''" and i': "i < length \ta''\\<^bsub>o\<^esub>" and eq: "take i \ta''\\<^bsub>o\<^esub> = take i \ta'\\<^bsub>o\<^esub>" and ta''_i: "\ta''\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v')" and len': "length \ta''\\<^bsub>o\<^esub> \ max n (length \ta'\\<^bsub>o\<^esub>)" using if.non_speculative_readD[OF nsr Red ns tst red' aok' i _ ta'_i, of v'] by auto from len' len have "length \ta''\\<^bsub>o\<^esub> \ max n (length \ta\\<^bsub>o\<^esub>)" by simp moreover have "ta_seq_consist P ?vs' (llist_of (take (Suc i) \ta''\\<^bsub>o\<^esub>))" using eq sc' i' ta''_i v' by(simp add: take_Suc_conv_app_nth lappend_llist_of_llist_of[symmetric] ta_seq_consist_lappend del: lappend_llist_of_llist_of) moreover have eusi': "eq_upto_seq_inconsist P (take (Suc i) \ta\\<^bsub>o\<^esub>) (take (Suc i) \ta''\\<^bsub>o\<^esub>) ?vs'" proof(cases "i < length \ta\\<^bsub>o\<^esub>") case True with i' i len eq eusi ta'_i ta''_i v' show ?thesis by(auto simp add: take_Suc_conv_app_nth ta'_ta eq_upto_seq_inconsist_simps intro: eq_upto_seq_inconsist_appendI) next case False with i ta'_ta have "\ ta_seq_consist P ?vs' (llist_of (take i \ta\\<^bsub>o\<^esub>))" by auto then show ?thesis using False i' eq eusi by(simp add: take_Suc_conv_app_nth eq_upto_seq_inconsist_append2) qed moreover { assume "ta_seq_consist P ?vs' (llist_of (take (Suc i) \ta\\<^bsub>o\<^esub>))" with True have "ta'' = ta" by simp } ultimately show ?thesis using red'' aok'' True by blast next case False hence "ta_seq_consist P ?vs' (llist_of (take (Suc i) \ta\\<^bsub>o\<^esub>)) \ length \ta'\\<^bsub>o\<^esub> \ i \ ta_seq_consist P ?vs' (llist_of (take (Suc i) \ta'\\<^bsub>o\<^esub>))" (is "?case1 \ ?case2 \ ?case3") by auto thus ?thesis proof(elim disjCE) assume "?case1" moreover hence "eq_upto_seq_inconsist P (take (Suc i) \ta\\<^bsub>o\<^esub>) (take (Suc i) \ta\\<^bsub>o\<^esub>) ?vs'" by(rule ta_seq_consist_imp_eq_upto_seq_inconsist_refl) ultimately show ?thesis using red aok by fastforce next assume "?case2" and "\ ?case1" have "eq_upto_seq_inconsist P (take (Suc i) \ta\\<^bsub>o\<^esub>) (take (Suc i) \ta'\\<^bsub>o\<^esub>) ?vs'" proof(cases "i < length \ta\\<^bsub>o\<^esub>") case True from \?case2\ \\ ?case1\ have "\ ta_seq_consist P ?vs' (llist_of (take i \ta\\<^bsub>o\<^esub>))" by(auto simp add: ta'_ta) hence "eq_upto_seq_inconsist P (take i \ta\\<^bsub>o\<^esub> @ [\ta\\<^bsub>o\<^esub> ! i]) (take i \ta'\\<^bsub>o\<^esub> @ []) ?vs'" by(blast intro: eq_upto_seq_inconsist_appendI[OF eusi]) thus ?thesis using True \?case2\ by(simp add: take_Suc_conv_app_nth) next case False with eusi \?case2\ show ?thesis by simp qed with red' aok' len sc' eusi \?case2\ \\ ?case1\show ?thesis by (fastforce simp add: take_all simp del: split_paired_Ex) next assume "?case3" and "\ ?case1" and "\ ?case2" with len eusi ta'_ta have "eq_upto_seq_inconsist P (take (Suc i) \ta\\<^bsub>o\<^esub>) (take (Suc i) \ta'\\<^bsub>o\<^esub>) ?vs'" by(cases "i < length \ta\\<^bsub>o\<^esub>")(auto simp add: take_Suc_conv_app_nth lappend_llist_of_llist_of[symmetric] ta_seq_consist_lappend intro: eq_upto_seq_inconsist_appendI eq_upto_seq_inconsist_append2 cong: action.case_cong obs_event.case_cong) with red' aok' \?case3\ len \\ ?case1\ show ?thesis by blast qed qed qed } from this[of "max n (length \ta\\<^bsub>o\<^esub>)"] show "\ta' x'' m''. t \ (x, shr s') -ta'\i (x'', m'') \ mthr.if.actions_ok s' t ta' \ ta_seq_consist P ?vs' (llist_of \ta'\\<^bsub>o\<^esub>) \ eq_upto_seq_inconsist P \ta\\<^bsub>o\<^esub> \ta'\\<^bsub>o\<^esub> ?vs'" by(auto simp del: split_paired_Ex cong: conj_cong) qed lemma non_speculative_read_into_hb_completion: fixes status f C M params E defines "E \ lift_start_obs start_tid start_heap_obs" and "vs \ w_values P (\_. {}) (map snd E)" and "s \ init_fin_lift_state status (start_state f P C M params)" assumes wf: "wf_syscls P" and nsr: "if.non_speculative_read n s vs" and wt: "ts_ok (init_fin_lift wfx) (thr s) (shr s)" and ka: "known_addrs start_tid (f (fst (method P C M)) M (fst (snd (method P C M))) (fst (snd (snd (method P C M)))) (the (snd (snd (snd (method P C M))))) params) \ allocated start_heap" shows "if.hb_completion s E" proof fix ttas s' t x ta x' m' i assume Red: "mthr.if.RedT s ttas s'" and ns: "non_speculative P (w_values P (\_. {}) (map snd E)) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" and tst: "thr s' t = \(x, no_wait_locks)\" and red: "t \ (x, shr s') -ta\i (x', m')" and aok: "mthr.if.actions_ok s' t ta" and nsi: "non_speculative P (w_values P (w_values P (\_. {}) (map snd E)) (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of (take i \ta\\<^bsub>o\<^esub>))" let ?E = "E @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) ttas) @ map (Pair t) (take i \ta\\<^bsub>o\<^esub>)" let ?vs = "w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))" from ns have ns': "non_speculative P (\_. {}) (llist_of (map snd (lift_start_obs start_tid start_heap_obs) @ concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" unfolding lappend_llist_of_llist_of[symmetric] by(simp add: non_speculative_lappend E_def non_speculative_start_heap_obs del: lappend_llist_of_llist_of) from start_state_vs_conf[OF wf] have vs: "vs_conf P (shr s) vs" unfolding vs_def E_def s_def by(simp add: init_fin_lift_state_conv_simps start_state_def split_def) from Red wt ns vs have wt': "ts_ok (init_fin_lift wfx) (thr s') (shr s')" unfolding vs_def by(rule if_RedT_non_speculative_invar) hence wtt: "init_fin_lift wfx t x (shr s')" using tst by(rule ts_okD) { fix j have "\ta' x'' m''. t \ (x, shr s') -ta'\i (x'', m'') \ mthr.if.actions_ok s' t ta' \ length \ta'\\<^bsub>o\<^esub> \ max n (length \ta\\<^bsub>o\<^esub>) \ take i \ta'\\<^bsub>o\<^esub> = take i \ta\\<^bsub>o\<^esub> \ ta_hb_consistent P ?E (llist_of (map (Pair t) (take j (drop i \ta'\\<^bsub>o\<^esub>)))) \ (i < length \ta\\<^bsub>o\<^esub> \ i < length \ta'\\<^bsub>o\<^esub>) \ (if \ad al v. \ta\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v) then sim_action else (=)) (\ta\\<^bsub>o\<^esub> ! i) (\ta'\\<^bsub>o\<^esub> ! i)" proof(induct j) case 0 from red aok show ?case by(fastforce simp del: split_paired_Ex) next case (Suc j) then obtain ta' x'' m'' where red': "t \ (x, shr s') -ta'\i (x'', m'')" and aok': "mthr.if.actions_ok s' t ta'" and len: "length \ta'\\<^bsub>o\<^esub> \ max n (length \ta\\<^bsub>o\<^esub>)" and eq: "take i \ta'\\<^bsub>o\<^esub> = take i \ta\\<^bsub>o\<^esub>" and hb: "ta_hb_consistent P ?E (llist_of (map (Pair t) (take j (drop i \ta'\\<^bsub>o\<^esub>))))" and len_i: "i < length \ta\\<^bsub>o\<^esub> \ i < length \ta'\\<^bsub>o\<^esub>" and sim_i: "(if \ad al v. \ta\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v) then sim_action else (=)) (\ta\\<^bsub>o\<^esub> ! i) (\ta'\\<^bsub>o\<^esub> ! i)" by blast show ?case proof(cases "i + j < length \ta'\\<^bsub>o\<^esub>") case False with red' aok' len eq hb len_i sim_i show ?thesis by(fastforce simp del: split_paired_Ex) next case True note j = this show ?thesis proof(cases "\ad al v. \ta'\\<^bsub>o\<^esub> ! (i + j) = NormalAction (ReadMem ad al v)") case True then obtain ad al v where ta'_j: "\ta'\\<^bsub>o\<^esub> ! (i + j) = NormalAction (ReadMem ad al v)" by blast hence read: "NormalAction (ReadMem ad al v) \ set \ta'\\<^bsub>o\<^esub>" using j by(auto simp add: in_set_conv_nth) with red' have "ad \ known_addrs_if t x" by(rule if_red_read_knows_addr) hence "ad \ if.known_addrs_state s'" using tst by(rule if.known_addrs_stateI) from init_fin_red_read_typeable[OF red' wtt read] obtain T where type_adal: "P,shr s' \ ad@al : T" .. from redT_updWs_total[of t "wset s'" "\ta'\\<^bsub>w\<^esub>"] red' tst aok' obtain s'' where redT': "mthr.if.redT s' (t, ta') s''" by(auto dest!: mthr.if.redT.redT_normal) with wf Red have "\w. w \ new_actions_for P (llist_of (E @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) ttas))) (ad, al)" (is "\w. ?new_w w") using read ns' ka wt type_adal unfolding s_def E_def by(rule read_non_speculative_new_actions_for) then obtain w where w: "?new_w w" .. define E'' where "E'' = ?E @ map (Pair t) (take (Suc j) (drop i \ta'\\<^bsub>o\<^esub>))" from Red redT' have "mthr.if.RedT s (ttas @ [(t, ta')]) s''" unfolding mthr.if.RedT_def .. hence tsa: "thread_start_actions_ok (llist_of (lift_start_obs start_tid start_heap_obs @ concat (map (\(t, ta). map (Pair t) \ta\\<^bsub>o\<^esub>) (ttas @ [(t, ta')]))))" unfolding s_def by(rule thread_start_actions_ok_init_fin_RedT) hence "thread_start_actions_ok (llist_of E'')" unfolding E_def[symmetric] E''_def by(rule thread_start_actions_ok_prefix)(rule lprefix_llist_ofI, simp, metis append_take_drop_id eq map_append) moreover from w have "w \ actions (llist_of E'')" unfolding E''_def by(auto simp add: new_actions_for_def actions_def) moreover have "length ?E + j \ actions (llist_of E'')" using j by(auto simp add: E''_def actions_def) moreover from w have "is_new_action (action_obs (llist_of E'') w)" by(auto simp add: new_actions_for_def action_obs_def actions_def nth_append E''_def) moreover have "\ is_new_action (action_obs (llist_of E'') (length ?E + j))" using j ta'_j by(auto simp add: action_obs_def nth_append min_def E''_def)(subst (asm) nth_map, simp_all) ultimately have hb_w: "P,llist_of E'' \ w \hb length ?E + j" by(rule happens_before_new_not_new) define writes where "writes = {w. P,llist_of E'' \ w \hb length ?E + j \ w \ write_actions (llist_of E'') \ (ad, al) \ action_loc P (llist_of E'') w}" define w' where "w' = Max_torder (action_order (llist_of E'')) writes" have writes_actions: "writes \ actions (llist_of E'')" unfolding writes_def actions_def by(auto dest!: happens_before_into_action_order elim!: action_orderE simp add: actions_def) also have "finite \" by(simp add: actions_def) finally (finite_subset) have "finite writes" . moreover from hb_w w have w_writes: "w \ writes" by(auto 4 3 simp add: writes_def new_actions_for_def action_obs_def actions_def nth_append E''_def intro!: write_actions.intros elim!: is_new_action.cases) hence "writes \ {}" by auto with torder_action_order \finite writes\ have w'_writes: "w' \ writes" using writes_actions unfolding w'_def by(rule Max_torder_in_set) moreover { fix w'' assume "w'' \ writes" with torder_action_order \finite writes\ have "llist_of E'' \ w'' \a w'" using writes_actions unfolding w'_def by(rule Max_torder_above) } note w'_maximal = this define v' where "v' = value_written P (llist_of E'') w' (ad, al)" from nsi ta_hb_consistent_into_non_speculative[OF hb] have nsi': "non_speculative P (w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of (take (i + j) \ta'\\<^bsub>o\<^esub>))" unfolding take_add lappend_llist_of_llist_of[symmetric] non_speculative_lappend vs_def eq by(simp add: non_speculative_lappend o_def map_concat split_def del: lappend_llist_of_llist_of) from w'_writes have adal_w': "(ad, al) \ action_loc P (llist_of E'') w'" by(simp add: writes_def) from w'_writes have "w' \ write_actions (llist_of E'')" unfolding writes_def by blast then obtain "is_write_action (action_obs (llist_of E'') w')" and w'_actions: "w' \ actions (llist_of E'')" by cases hence "v' \ w_values P (\_. {}) (map snd E'') (ad, al)" proof cases case (NewHeapElem ad' CTn) hence "NormalAction (NewHeapElem ad' CTn) \ set (map snd E'')" using w'_actions unfolding in_set_conv_nth by(auto simp add: actions_def action_obs_def cong: conj_cong) moreover have "ad' = ad" and "(ad, al) \ action_loc_aux P (NormalAction (NewHeapElem ad CTn))" using adal_w' NewHeapElem by auto ultimately show ?thesis using NewHeapElem unfolding v'_def by(simp add: value_written.simps w_values_new_actionD) next case (WriteMem ad' al' v'') hence "NormalAction (WriteMem ad' al' v'') \ set (map snd E'')" using w'_actions unfolding in_set_conv_nth by(auto simp add: actions_def action_obs_def cong: conj_cong) moreover have "ad' = ad" "al' = al" using adal_w' WriteMem by auto ultimately show ?thesis using WriteMem unfolding v'_def by(simp add: value_written.simps w_values_WriteMemD) qed hence "v' \ w_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas) @ take (i + j) \ta'\\<^bsub>o\<^esub>) (ad, al)" using j ta'_j eq unfolding E''_def vs_def by(simp add: o_def split_def map_concat take_add take_Suc_conv_app_nth) from if.non_speculative_readD[OF nsr Red ns[folded vs_def] tst red' aok' j nsi' ta'_j this] obtain ta'' x'' m'' where red'': "t \ (x, shr s') -ta''\i (x'', m'')" and aok'': "mthr.if.actions_ok s' t ta''" and j': "i + j < length \ta''\\<^bsub>o\<^esub>" and eq': "take (i + j) \ta''\\<^bsub>o\<^esub> = take (i + j) \ta'\\<^bsub>o\<^esub>" and ta''_j: "\ta''\\<^bsub>o\<^esub> ! (i + j) = NormalAction (ReadMem ad al v')" and len': "length \ta''\\<^bsub>o\<^esub> \ max n (length \ta'\\<^bsub>o\<^esub>)" by blast define EE where "EE = ?E @ map (Pair t) (take j (drop i \ta''\\<^bsub>o\<^esub>))" define E' where "E' = ?E @ map (Pair t) (take j (drop i \ta''\\<^bsub>o\<^esub>)) @ [(t, NormalAction (ReadMem ad al v'))]" from len' len have "length \ta''\\<^bsub>o\<^esub> \ max n (length \ta\\<^bsub>o\<^esub>)" by simp moreover from eq' eq j j' have "take i \ta''\\<^bsub>o\<^esub> = take i \ta\\<^bsub>o\<^esub>" by(auto simp add: take_add min_def) moreover { note hb also have eq'': "take j (drop i \ta'\\<^bsub>o\<^esub>) = take j (drop i \ta''\\<^bsub>o\<^esub>)" using eq' j j' by(simp add: take_add min_def) also have "ta_hb_consistent P (?E @ list_of (llist_of (map (Pair t) (take j (drop i \ta''\\<^bsub>o\<^esub>))))) (llist_of [(t, \ta''\\<^bsub>o\<^esub> ! (i + j))])" unfolding llist_of.simps ta_hb_consistent_LCons ta_hb_consistent_LNil ta''_j prod.simps action.simps obs_event.simps list_of_llist_of append_assoc E'_def[symmetric, unfolded append_assoc] unfolding EE_def[symmetric, unfolded append_assoc] proof(intro conjI TrueI exI[where x=w'] strip) have "llist_of E'' [\] llist_of E'" using j len eq'' ta'_j unfolding E''_def E'_def by(auto simp add: sim_actions_def list_all2_append List.list_all2_refl split_beta take_Suc_conv_app_nth take_map[symmetric]) moreover have "length E'' = length E'" using j j' by(simp add: E''_def E'_def) ultimately have sim: "ltake (enat (length E')) (llist_of E'') [\] ltake (enat (length E')) (llist_of E')" by simp from w'_actions \length E'' = length E'\ have w'_len: "w' < length E'" by(simp add: actions_def) from \w' \ write_actions (llist_of E'')\ sim show "w' \ write_actions (llist_of E')" by(rule write_actions_change_prefix)(simp add: w'_len) from adal_w' action_loc_change_prefix[OF sim, of w' P] show "(ad, al) \ action_loc P (llist_of E') w'" by(simp add: w'_len) from ta'_j j have "length ?E + j \ read_actions (llist_of E'')" by(auto intro!: read_actions.intros simp add: action_obs_def actions_def E''_def min_def nth_append)(auto) hence "w' \ length ?E + j" using \w' \ write_actions (llist_of E'')\ by(auto dest: read_actions_not_write_actions) with w'_len have "w' < length ?E + j" by(simp add: E'_def) from j j' len' eq'' have "ltake (enat (length ?E + j)) (llist_of E'') = ltake (enat (length ?E + j)) (llist_of E')" by(auto simp add: E''_def E'_def min_def take_Suc_conv_app_nth) from value_written_change_prefix[OF this, of w' P] \w' < length ?E + j\ show "value_written P (llist_of E') w' (ad, al) = v'" unfolding v'_def by simp from \thread_start_actions_ok (llist_of E'')\ \llist_of E'' [\] llist_of E'\ have tsa'': "thread_start_actions_ok (llist_of E')" by(rule thread_start_actions_ok_change) from w'_writes j j' len len' have "P,llist_of E'' \ w' \hb length EE" by(auto simp add: EE_def writes_def min_def ac_simps) thus "P,llist_of E' \ w' \hb length EE" using tsa'' sim by(rule happens_before_change_prefix)(simp add: w'_len, simp add: EE_def E'_def) fix w'' assume w'': "w'' \ write_actions (llist_of E')" and adal_w'': "(ad, al) \ action_loc P (llist_of E') w''" from w'' have w''_len: "w'' < length E'" by(cases)(simp add: actions_def) from w'' sim[symmetric] have w'': "w'' \ write_actions (llist_of E'')" by(rule write_actions_change_prefix)(simp add: w''_len) from adal_w'' action_loc_change_prefix[OF sim[symmetric], of w'' P] w''_len have adal_w'': "(ad, al) \ action_loc P (llist_of E'') w''" by simp { presume w'_w'': "llist_of E' \ w' \a w''" and w''_hb: "P,llist_of E' \ w'' \hb length EE" from w''_hb \thread_start_actions_ok (llist_of E'')\ sim[symmetric] have "P,llist_of E'' \ w'' \hb length EE" by(rule happens_before_change_prefix)(simp add: w''_len, simp add: E'_def EE_def) with w'' adal_w'' j j' len len' have "w'' \ writes" by(auto simp add: writes_def EE_def min_def ac_simps split: if_split_asm) hence "llist_of E'' \ w'' \a w'" by(rule w'_maximal) hence "llist_of E' \ w'' \a w'" using sim by(rule action_order_change_prefix)(simp_all add: w'_len w''_len) thus "w'' = w'" "w'' = w'" using w'_w'' by(rule antisymPD[OF antisym_action_order])+ } { assume "P,llist_of E' \ w' \hb w'' \ P,llist_of E' \ w'' \hb length EE" thus "llist_of E' \ w' \a w''" "P,llist_of E' \ w'' \hb length EE" by(blast dest: happens_before_into_action_order)+ } { assume "is_volatile P al \ P,llist_of E' \ w' \so w'' \ P,llist_of E' \ w'' \so length EE" then obtain vol: "is_volatile P al" and so: "P,llist_of E' \ w' \so w''" and so': "P,llist_of E' \ w'' \so length EE" by blast from so show "llist_of E' \ w' \a w''" by(blast elim: sync_orderE) show "P,llist_of E' \ w'' \hb length EE" proof(cases "is_new_action (action_obs (llist_of E') w'')") case True with \w'' \ write_actions (llist_of E')\ ta''_j show ?thesis by cases(rule happens_before_new_not_new[OF tsa''], auto simp add: actions_def EE_def E'_def action_obs_def min_def nth_append) next case False with \w'' \ write_actions (llist_of E')\ \(ad, al) \ action_loc P (llist_of E') w''\ obtain v'' where "action_obs (llist_of E') w'' = NormalAction (WriteMem ad al v'')" by cases(auto elim: is_write_action.cases) with ta''_j w'' j j' len len' have "P \ (action_tid (llist_of E') w'', action_obs (llist_of E') w'') \sw (action_tid (llist_of E') (length EE), action_obs (llist_of E') (length EE))" by(auto simp add: E'_def EE_def action_obs_def min_def nth_append Volatile) with so' have "P,llist_of E' \ w'' \sw length EE" by(rule sync_withI) thus ?thesis unfolding po_sw_def [abs_def] by(blast intro: tranclp.r_into_trancl) qed } qed ultimately have "ta_hb_consistent P ?E (lappend (llist_of (map (Pair t) (take j (drop i \ta''\\<^bsub>o\<^esub>)))) (llist_of ([(t, \ta''\\<^bsub>o\<^esub> ! (i + j))])))" by(rule ta_hb_consistent_lappendI) simp hence "ta_hb_consistent P ?E (llist_of (map (Pair t) (take (Suc j) (drop i \ta''\\<^bsub>o\<^esub>))))" using j' unfolding lappend_llist_of_llist_of by(simp add: take_Suc_conv_app_nth) } moreover from len_i have "i < length \ta\\<^bsub>o\<^esub> \ i < length \ta''\\<^bsub>o\<^esub>" using eq' j' by auto moreover from sim_i eq' ta''_j ta'_j have "(if \ad al v. \ta\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v) then sim_action else (=)) (\ta\\<^bsub>o\<^esub> ! i) (\ta''\\<^bsub>o\<^esub> ! i)" by(cases "j = 0")(auto split: if_split_asm, (metis add_strict_left_mono add_0_right nth_take)+) ultimately show ?thesis using red'' aok'' by blast next case False hence "ta_hb_consistent P (?E @ list_of (llist_of (map (Pair t) (take j (drop i \ta'\\<^bsub>o\<^esub>))))) (llist_of [(t, \ta'\\<^bsub>o\<^esub> ! (i + j))])" by(simp add: ta_hb_consistent_LCons split: action.split obs_event.split) with hb have "ta_hb_consistent P ?E (lappend (llist_of (map (Pair t) (take j (drop i \ta'\\<^bsub>o\<^esub>)))) (llist_of ([(t, \ta'\\<^bsub>o\<^esub> ! (i + j))])))" by(rule ta_hb_consistent_lappendI) simp hence "ta_hb_consistent P ?E (llist_of (map (Pair t) (take (Suc j) (drop i \ta'\\<^bsub>o\<^esub>))))" using j unfolding lappend_llist_of_llist_of by(simp add: take_Suc_conv_app_nth) with red' aok' len eq len_i sim_i show ?thesis by blast qed qed qed } from this[of "max n (length \ta\\<^bsub>o\<^esub>)"] show "\ta' x'' m''. t \ (x, shr s') -ta'\i (x'', m'') \ mthr.if.actions_ok s' t ta' \ take i \ta'\\<^bsub>o\<^esub> = take i \ta\\<^bsub>o\<^esub> \ ta_hb_consistent P ?E (llist_of (map (Pair t) (drop i \ta'\\<^bsub>o\<^esub>))) \ (i < length \ta\\<^bsub>o\<^esub> \ i < length \ta'\\<^bsub>o\<^esub>) \ (if \ad al v. \ta\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v) then sim_action else (=)) (\ta\\<^bsub>o\<^esub> ! i) (\ta'\\<^bsub>o\<^esub> ! i)" by(simp del: split_paired_Ex cong: conj_cong split del: if_split) blast qed end end diff --git a/thys/JinjaThreads/MM/JMM_Spec.thy b/thys/JinjaThreads/MM/JMM_Spec.thy --- a/thys/JinjaThreads/MM/JMM_Spec.thy +++ b/thys/JinjaThreads/MM/JMM_Spec.thy @@ -1,2069 +1,2069 @@ (* Title: JinjaThreads/MM/JMM_Spec.thy Author: Andreas Lochbihler *) section \Axiomatic specification of the JMM\ theory JMM_Spec imports Orders "../Common/Observable_Events" Coinductive.Coinductive_List begin subsection \Definitions\ type_synonym JMM_action = nat type_synonym ('addr, 'thread_id) execution = "('thread_id \ ('addr, 'thread_id) obs_event action) llist" definition "actions" :: "('addr, 'thread_id) execution \ JMM_action set" where "actions E = {n. enat n < llength E}" definition action_tid :: "('addr, 'thread_id) execution \ JMM_action \ 'thread_id" where "action_tid E a = fst (lnth E a)" definition action_obs :: "('addr, 'thread_id) execution \ JMM_action \ ('addr, 'thread_id) obs_event action" where "action_obs E a = snd (lnth E a)" definition tactions :: "('addr, 'thread_id) execution \ 'thread_id \ JMM_action set" where "tactions E t = {a. a \ actions E \ action_tid E a = t}" inductive is_new_action :: "('addr, 'thread_id) obs_event action \ bool" where NewHeapElem: "is_new_action (NormalAction (NewHeapElem a hT))" inductive is_write_action :: "('addr, 'thread_id) obs_event action \ bool" where NewHeapElem: "is_write_action (NormalAction (NewHeapElem ad hT))" | WriteMem: "is_write_action (NormalAction (WriteMem ad al v))" text \ Initialisation actions are synchronisation actions iff they initialize volatile fields -- cf. JMM mailing list, message no. 62 (5 Nov. 2006). However, intuitively correct programs might not be correctly synchronized: \begin{verbatim} x = 0 --------------- r1 = x | r2 = x \end{verbatim} Here, if x is not volatile, the initial write can belong to at most one thread. Hence, it is happens-before to either r1 = x or r2 = x, but not both. In any sequentially consistent execution, both reads must read from the initilisation action x = 0, but it is not happens-before ordered to one of them. Moreover, if only volatile initialisations synchronize-with all thread-start actions, this breaks the proof of the DRF guarantee since it assumes that the happens-before relation $<=hb$ a for an action $a$ in a topologically sorted action sequence depends only on the actions before it. Counter example: (y is volatile) [(t1, start), (t1, init x), (t2, start), (t1, init y), ... Here, (t1, init x) $<=hb$ (t2, start) via: (t1, init x) $<=po$ (t1, init y) $<=sw$ (t2, start), but in [(t1, start), (t1, init x), (t2, start)], not (t1, init x) $<=hb$ (t2, start). Sevcik speculated that one might add an initialisation thread which performs all initialisation actions. All normal threads' start action would then synchronize on the final action of the initialisation thread. However, this contradicts the memory chain condition in the final field extension to the JMM (threads must read addresses of objects that they have not created themselves before they can access the fields of the object at that address) -- not modelled here. Instead, we leave every initialisation action in the thread it belongs to, but order it explicitly before the thread's start action and add synchronizes-with edges from \emph{all} initialisation actions to \emph{all} thread start actions. \ inductive saction :: "'m prog \ ('addr, 'thread_id) obs_event action \ bool" for P :: "'m prog" where NewHeapElem: "saction P (NormalAction (NewHeapElem a hT))" | Read: "is_volatile P al \ saction P (NormalAction (ReadMem a al v))" | Write: "is_volatile P al \ saction P (NormalAction (WriteMem a al v))" | ThreadStart: "saction P (NormalAction (ThreadStart t))" | ThreadJoin: "saction P (NormalAction (ThreadJoin t))" | SyncLock: "saction P (NormalAction (SyncLock a))" | SyncUnlock: "saction P (NormalAction (SyncUnlock a))" | ObsInterrupt: "saction P (NormalAction (ObsInterrupt t))" | ObsInterrupted: "saction P (NormalAction (ObsInterrupted t))" | InitialThreadAction: "saction P InitialThreadAction" | ThreadFinishAction: "saction P ThreadFinishAction" definition sactions :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action set" where "sactions P E = {a. a \ actions E \ saction P (action_obs E a)}" inductive_set write_actions :: "('addr, 'thread_id) execution \ JMM_action set" for E :: "('addr, 'thread_id) execution" where write_actionsI: "\ a \ actions E; is_write_action (action_obs E a) \ \ a \ write_actions E" text \ @{term NewObj} and @{term NewArr} actions only initialize those fields and array cells that are in fact in the object or array. Hence, they are not a write for - reads from addresses for which no object/array is created during the whole execution - reads from fields/cells that are not part of the object/array at the specified address. \ primrec addr_locs :: "'m prog \ htype \ addr_loc set" where "addr_locs P (Class_type C) = {CField D F|D F. \fm T. P \ C has F:T (fm) in D}" | "addr_locs P (Array_type T n) = ({ACell n'|n'. n' < n} \ {CField Object F|F. \fm T. P \ Object has F:T (fm) in Object})" text \ \action_loc_aux\ would naturally be an inductive set, but inductive\_set does not allow to pattern match on parameters. Hence, specify it using function and derive the setup manually. \ fun action_loc_aux :: "'m prog \ ('addr, 'thread_id) obs_event action \ ('addr \ addr_loc) set" where "action_loc_aux P (NormalAction (NewHeapElem ad (Class_type C))) = {(ad, CField D F)|D F T fm. P \ C has F:T (fm) in D}" | "action_loc_aux P (NormalAction (NewHeapElem ad (Array_type T n'))) = {(ad, ACell n)|n. n < n'} \ {(ad, CField D F)|D F T fm. P \ Object has F:T (fm) in D}" | "action_loc_aux P (NormalAction (WriteMem ad al v)) = {(ad, al)}" | "action_loc_aux P (NormalAction (ReadMem ad al v)) = {(ad, al)}" | "action_loc_aux _ _ = {}" lemma action_loc_aux_intros [intro?]: "P \ class_type_of hT has F:T (fm) in D \ (ad, CField D F) \ action_loc_aux P (NormalAction (NewHeapElem ad hT))" "n < n' \ (ad, ACell n) \ action_loc_aux P (NormalAction (NewHeapElem ad (Array_type T n')))" "(ad, al) \ action_loc_aux P (NormalAction (WriteMem ad al v))" "(ad, al) \ action_loc_aux P (NormalAction (ReadMem ad al v))" by(cases hT) auto lemma action_loc_aux_cases [elim?, cases set: action_loc_aux]: assumes "adal \ action_loc_aux P obs" obtains (NewHeapElem) hT F T fm D ad where "obs = NormalAction (NewHeapElem ad hT)" "adal = (ad, CField D F)" "P \ class_type_of hT has F:T (fm) in D" | (NewArr) n n' ad T where "obs = NormalAction (NewHeapElem ad (Array_type T n'))" "adal = (ad, ACell n)" "n < n'" | (WriteMem) ad al v where "obs = NormalAction (WriteMem ad al v)" "adal = (ad, al)" | (ReadMem) ad al v where "obs = NormalAction (ReadMem ad al v)" "adal = (ad, al)" using assms by(cases "(P, obs)" rule: action_loc_aux.cases) fastforce+ lemma action_loc_aux_simps [simp]: "(ad', al') \ action_loc_aux P (NormalAction (NewHeapElem ad hT)) \ (\D F T fm. ad = ad' \ al' = CField D F \ P \ class_type_of hT has F:T (fm) in D) \ (\n T n'. ad = ad' \ al' = ACell n \ hT = Array_type T n' \ n < n')" "(ad', al') \ action_loc_aux P (NormalAction (WriteMem ad al v)) \ ad = ad' \ al = al'" "(ad', al') \ action_loc_aux P (NormalAction (ReadMem ad al v)) \ ad = ad' \ al = al'" "(ad', al') \ action_loc_aux P InitialThreadAction" "(ad', al') \ action_loc_aux P ThreadFinishAction" "(ad', al') \ action_loc_aux P (NormalAction (ExternalCall a m vs v))" "(ad', al') \ action_loc_aux P (NormalAction (ThreadStart t))" "(ad', al') \ action_loc_aux P (NormalAction (ThreadJoin t))" "(ad', al') \ action_loc_aux P (NormalAction (SyncLock a))" "(ad', al') \ action_loc_aux P (NormalAction (SyncUnlock a))" "(ad', al') \ action_loc_aux P (NormalAction (ObsInterrupt t))" "(ad', al') \ action_loc_aux P (NormalAction (ObsInterrupted t))" by(cases hT) auto declare action_loc_aux.simps [simp del] abbreviation action_loc :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ ('addr \ addr_loc) set" where "action_loc P E a \ action_loc_aux P (action_obs E a)" inductive_set read_actions :: "('addr, 'thread_id) execution \ JMM_action set" for E :: "('addr, 'thread_id) execution" where ReadMem: "\ a \ actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ a \ read_actions E" fun addr_loc_default :: "'m prog \ htype \ addr_loc \ 'addr val" where "addr_loc_default P (Class_type C) (CField D F) = default_val (fst (the (map_of (fields P C) (F, D))))" | "addr_loc_default P (Array_type T n) (ACell n') = default_val T" | addr_loc_default_Array_CField: "addr_loc_default P (Array_type T n) (CField D F) = default_val (fst (the (map_of (fields P Object) (F, Object))))" | "addr_loc_default P _ _ = undefined" definition new_actions_for :: "'m prog \ ('addr, 'thread_id) execution \ ('addr \ addr_loc) \ JMM_action set" where "new_actions_for P E adal = {a. a \ actions E \ adal \ action_loc P E a \ is_new_action (action_obs E a)}" inductive_set external_actions :: "('addr, 'thread_id) execution \ JMM_action set" for E :: "('addr, 'thread_id) execution" where "\ a \ actions E; action_obs E a = NormalAction (ExternalCall ad M vs v) \ \ a \ external_actions E" fun value_written_aux :: "'m prog \ ('addr, 'thread_id) obs_event action \ addr_loc \ 'addr val" where "value_written_aux P (NormalAction (NewHeapElem ad' hT)) al = addr_loc_default P hT al" | value_written_aux_WriteMem': "value_written_aux P (NormalAction (WriteMem ad al' v)) al = (if al = al' then v else undefined)" | value_written_aux_undefined: "value_written_aux P _ al = undefined" primrec value_written :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ ('addr \ addr_loc) \ 'addr val" where "value_written P E a (ad, al) = value_written_aux P (action_obs E a) al" definition value_read :: "('addr, 'thread_id) execution \ JMM_action \ 'addr val" where "value_read E a = (case action_obs E a of NormalAction obs \ (case obs of ReadMem ad al v \ v | _ \ undefined) | _ \ undefined)" definition action_order :: "('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_ \ _ \a _" [51,0,50] 50) where "E \ a \a a' \ a \ actions E \ a' \ actions E \ (if is_new_action (action_obs E a) then is_new_action (action_obs E a') \ a \ a' else \ is_new_action (action_obs E a') \ a \ a')" definition program_order :: "('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_ \ _ \po _" [51,0,50] 50) where "E \ a \po a' \ E \ a \a a' \ action_tid E a = action_tid E a'" inductive synchronizes_with :: "'m prog \ ('thread_id \ ('addr, 'thread_id) obs_event action) \ ('thread_id \ ('addr, 'thread_id) obs_event action) \ bool" ("_ \ _ \sw _" [51, 51, 51] 50) for P :: "'m prog" where ThreadStart: "P \ (t, NormalAction (ThreadStart t')) \sw (t', InitialThreadAction)" | ThreadFinish: "P \ (t, ThreadFinishAction) \sw (t', NormalAction (ThreadJoin t))" | UnlockLock: "P \ (t, NormalAction (SyncUnlock a)) \sw (t', NormalAction (SyncLock a))" | \ \Only volatile writes synchronize with volatile reads. We could check volatility of @{term "al"} here, but this is checked by @{term "sactions"} in @{text sync_with} anyway.\ Volatile: "P \ (t, NormalAction (WriteMem a al v)) \sw (t', NormalAction (ReadMem a al v'))" | VolatileNew: " al \ addr_locs P hT \ P \ (t, NormalAction (NewHeapElem a hT)) \sw (t', NormalAction (ReadMem a al v))" | NewHeapElem: "P \ (t, NormalAction (NewHeapElem a hT)) \sw (t', InitialThreadAction)" | Interrupt: "P \ (t, NormalAction (ObsInterrupt t')) \sw (t'', NormalAction (ObsInterrupted t'))" definition sync_order :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_,_ \ _ \so _" [51,0,0,50] 50) where "P,E \ a \so a' \ a \ sactions P E \ a' \ sactions P E \ E \ a \a a'" definition sync_with :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_,_ \ _ \sw _" [51, 0, 0, 50] 50) where "P,E \ a \sw a' \ P,E \ a \so a' \ P \ (action_tid E a, action_obs E a) \sw (action_tid E a', action_obs E a')" definition po_sw :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" where "po_sw P E a a' \ E \ a \po a' \ P,E \ a \sw a'" abbreviation happens_before :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_,_ \ _ \hb _" [51, 0, 0, 50] 50) where "happens_before P E \ (po_sw P E)^++" type_synonym write_seen = "JMM_action \ JMM_action" definition is_write_seen :: "'m prog \ ('addr, 'thread_id) execution \ write_seen \ bool" where "is_write_seen P E ws \ (\a \ read_actions E. \ad al v. action_obs E a = NormalAction (ReadMem ad al v) \ ws a \ write_actions E \ (ad, al) \ action_loc P E (ws a) \ value_written P E (ws a) (ad, al) = v \ \ P,E \ a \hb ws a \ (is_volatile P al \ \ P,E \ a \so ws a) \ (\w' \ write_actions E. (ad, al) \ action_loc P E w' \ (P,E \ ws a \hb w' \ P,E \ w' \hb a \ is_volatile P al \ P,E \ ws a \so w' \ P,E \ w' \so a) \ w' = ws a))" definition thread_start_actions_ok :: "('addr, 'thread_id) execution \ bool" where "thread_start_actions_ok E \ (\a \ actions E. \ is_new_action (action_obs E a) \ (\i. i \ a \ action_obs E i = InitialThreadAction \ action_tid E i = action_tid E a))" primrec wf_exec :: "'m prog \ ('addr, 'thread_id) execution \ write_seen \ bool" ("_ \ _ \" [51, 50] 51) where "P \ (E, ws) \ \ is_write_seen P E ws \ thread_start_actions_ok E" inductive most_recent_write_for :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_,_ \ _ \mrw _" [50, 0, 51] 51) for P :: "'m prog" and E :: "('addr, 'thread_id) execution" and ra :: JMM_action and wa :: JMM_action where "\ ra \ read_actions E; adal \ action_loc P E ra; E \ wa \a ra; wa \ write_actions E; adal \ action_loc P E wa; \wa'. \ wa' \ write_actions E; adal \ action_loc P E wa' \ \ E \ wa' \a wa \ E \ ra \a wa' \ \ P,E \ ra \mrw wa" primrec sequentially_consistent :: "'m prog \ (('addr, 'thread_id) execution \ write_seen) \ bool" where "sequentially_consistent P (E, ws) \ (\r \ read_actions E. P,E \ r \mrw ws r)" subsection \Actions\ inductive_cases is_new_action_cases [elim!]: "is_new_action (NormalAction (ExternalCall a M vs v))" "is_new_action (NormalAction (ReadMem a al v))" "is_new_action (NormalAction (WriteMem a al v))" "is_new_action (NormalAction (NewHeapElem a hT))" "is_new_action (NormalAction (ThreadStart t))" "is_new_action (NormalAction (ThreadJoin t))" "is_new_action (NormalAction (SyncLock a))" "is_new_action (NormalAction (SyncUnlock a))" "is_new_action (NormalAction (ObsInterrupt t))" "is_new_action (NormalAction (ObsInterrupted t))" "is_new_action InitialThreadAction" "is_new_action ThreadFinishAction" inductive_simps is_new_action_simps [simp]: "is_new_action (NormalAction (NewHeapElem a hT))" "is_new_action (NormalAction (ExternalCall a M vs v))" "is_new_action (NormalAction (ReadMem a al v))" "is_new_action (NormalAction (WriteMem a al v))" "is_new_action (NormalAction (ThreadStart t))" "is_new_action (NormalAction (ThreadJoin t))" "is_new_action (NormalAction (SyncLock a))" "is_new_action (NormalAction (SyncUnlock a))" "is_new_action (NormalAction (ObsInterrupt t))" "is_new_action (NormalAction (ObsInterrupted t))" "is_new_action InitialThreadAction" "is_new_action ThreadFinishAction" lemmas is_new_action_iff = is_new_action.simps inductive_simps is_write_action_simps [simp]: "is_write_action InitialThreadAction" "is_write_action ThreadFinishAction" "is_write_action (NormalAction (ExternalCall a m vs v))" "is_write_action (NormalAction (ReadMem a al v))" "is_write_action (NormalAction (WriteMem a al v))" "is_write_action (NormalAction (NewHeapElem a hT))" "is_write_action (NormalAction (ThreadStart t))" "is_write_action (NormalAction (ThreadJoin t))" "is_write_action (NormalAction (SyncLock a))" "is_write_action (NormalAction (SyncUnlock a))" "is_write_action (NormalAction (ObsInterrupt t))" "is_write_action (NormalAction (ObsInterrupted t))" declare saction.intros [intro!] inductive_cases saction_cases [elim!]: "saction P (NormalAction (ExternalCall a M vs v))" "saction P (NormalAction (ReadMem a al v))" "saction P (NormalAction (WriteMem a al v))" "saction P (NormalAction (NewHeapElem a hT))" "saction P (NormalAction (ThreadStart t))" "saction P (NormalAction (ThreadJoin t))" "saction P (NormalAction (SyncLock a))" "saction P (NormalAction (SyncUnlock a))" "saction P (NormalAction (ObsInterrupt t))" "saction P (NormalAction (ObsInterrupted t))" "saction P InitialThreadAction" "saction P ThreadFinishAction" inductive_simps saction_simps [simp]: "saction P (NormalAction (ExternalCall a M vs v))" "saction P (NormalAction (ReadMem a al v))" "saction P (NormalAction (WriteMem a al v))" "saction P (NormalAction (NewHeapElem a hT))" "saction P (NormalAction (ThreadStart t))" "saction P (NormalAction (ThreadJoin t))" "saction P (NormalAction (SyncLock a))" "saction P (NormalAction (SyncUnlock a))" "saction P (NormalAction (ObsInterrupt t))" "saction P (NormalAction (ObsInterrupted t))" "saction P InitialThreadAction" "saction P ThreadFinishAction" lemma new_action_saction [simp, intro]: "is_new_action a \ saction P a" by(blast elim: is_new_action.cases) lemmas saction_iff = saction.simps lemma actionsD: "a \ actions E \ enat a < llength E" unfolding actions_def by blast lemma actionsE: assumes "a \ actions E" obtains "enat a < llength E" using assms unfolding actions_def by blast lemma actions_lappend: "llength xs = enat n \ actions (lappend xs ys) = actions xs \ ((+) n) ` actions ys" unfolding actions_def apply safe apply(erule contrapos_np) apply(rule_tac x="x - n" in image_eqI) apply simp_all apply(case_tac [!] "llength ys") apply simp_all done lemma tactionsE: assumes "a \ tactions E t" obtains obs where "a \ actions E" "action_tid E a = t" "action_obs E a = obs" using assms by(cases "lnth E a")(auto simp add: tactions_def action_tid_def action_obs_def) lemma sactionsI: "\ a \ actions E; saction P (action_obs E a) \ \ a \ sactions P E" unfolding sactions_def by blast lemma sactionsE: assumes "a \ sactions P E" obtains "a \ actions E" "saction P (action_obs E a)" using assms unfolding sactions_def by blast lemma sactions_actions [simp]: "a \ sactions P E \ a \ actions E" by(rule sactionsE) lemma value_written_aux_WriteMem [simp]: "value_written_aux P (NormalAction (WriteMem ad al v)) al = v" by simp declare value_written_aux_undefined [simp del] declare value_written_aux_WriteMem' [simp del] inductive_simps is_write_action_iff: "is_write_action a" inductive_simps write_actions_iff: "a \ write_actions E" lemma write_actions_actions [simp]: "a \ write_actions E \ a \ actions E" by(rule write_actions.induct) inductive_simps read_actions_iff: "a \ read_actions E" lemma read_actions_actions [simp]: "a \ read_actions E \ a \ actions E" by(rule read_actions.induct) lemma read_action_action_locE: assumes "r \ read_actions E" obtains ad al where "(ad, al) \ action_loc P E r" using assms by cases auto lemma read_actions_not_write_actions: "\ a \ read_actions E; a \ write_actions E \ \ False" by(auto elim!: read_actions.cases write_actions.cases) lemma read_actions_Int_write_actions [simp]: "read_actions E \ write_actions E = {}" "write_actions E \ read_actions E = {}" by(blast dest: read_actions_not_write_actions)+ lemma action_loc_addr_fun: "\ (ad, al) \ action_loc P E a; (ad', al') \ action_loc P E a \ \ ad = ad'" by(auto elim!: action_loc_aux_cases) lemma value_written_cong [cong]: "\ P = P'; a = a'; action_obs E a' = action_obs E' a' \ \ value_written P E a = value_written P' E' a'" by(rule ext)(auto split: action.splits) declare value_written.simps [simp del] lemma new_actionsI: "\ a \ actions E; adal \ action_loc P E a; is_new_action (action_obs E a) \ \ a \ new_actions_for P E adal" unfolding new_actions_for_def by blast lemma new_actionsE: assumes "a \ new_actions_for P E adal" obtains "a \ actions E" "adal \ action_loc P E a" "is_new_action (action_obs E a)" using assms unfolding new_actions_for_def by blast lemma action_loc_read_action_singleton: "\ r \ read_actions E; adal \ action_loc P E r; adal' \ action_loc P E r \ \ adal = adal'" by(cases adal, cases adal')(fastforce elim: read_actions.cases action_loc_aux_cases) lemma addr_locsI: "P \ class_type_of hT has F:T (fm) in D \ CField D F \ addr_locs P hT" "\ hT = Array_type T n; n' < n \ \ ACell n' \ addr_locs P hT" by(cases hT)(auto dest: has_field_decl_above) subsection \Orders\ subsection \Action order\ lemma action_orderI: assumes "a \ actions E" "a' \ actions E" and "\ is_new_action (action_obs E a); is_new_action (action_obs E a') \ \ a \ a'" and "\ is_new_action (action_obs E a) \ \ is_new_action (action_obs E a') \ a \ a'" shows "E \ a \a a'" using assms unfolding action_order_def by simp lemma action_orderE: assumes "E \ a \a a'" obtains "a \ actions E" "a' \ actions E" "is_new_action (action_obs E a)" "is_new_action (action_obs E a') \ a \ a'" | "a \ actions E" "a' \ actions E" "\ is_new_action (action_obs E a)" "\ is_new_action (action_obs E a')" "a \ a'" using assms unfolding action_order_def by(simp split: if_split_asm) lemma refl_action_order: "refl_onP (actions E) (action_order E)" by(rule refl_onPI)(auto elim: action_orderE intro: action_orderI) lemma antisym_action_order: "antisymp (action_order E)" by(rule antisympI)(auto elim!: action_orderE) lemma trans_action_order: "transp (action_order E)" by(rule transpI)(auto elim!: action_orderE intro: action_orderI) lemma porder_action_order: "porder_on (actions E) (action_order E)" by(blast intro: porder_onI refl_action_order antisym_action_order trans_action_order) lemma total_action_order: "total_onP (actions E) (action_order E)" by(rule total_onPI)(auto simp add: action_order_def) lemma torder_action_order: "torder_on (actions E) (action_order E)" by(blast intro: torder_onI total_action_order porder_action_order) lemma wf_action_order: "wfP (action_order E)\<^sup>\\<^sup>\" unfolding wfP_eq_minimal proof(intro strip) fix Q and x :: JMM_action assume "x \ Q" show "\z \ Q. \y. (action_order E)\<^sup>\\<^sup>\ y z \ y \ Q" proof(cases "\a \ Q. a \ actions E \ is_new_action (action_obs E a)") case True then obtain a where a: "a \ actions E \ is_new_action (action_obs E a) \ a \ Q" by blast define a' where "a' = (LEAST a'. a' \ actions E \ is_new_action (action_obs E a') \ a' \ Q)" from a have a': "a' \ actions E \ is_new_action (action_obs E a') \ a' \ Q" unfolding a'_def by(rule LeastI) { fix y assume y_le_a': "(action_order E)\<^sup>\\<^sup>\ y a'" have "y \ Q" proof assume "y \ Q" with y_le_a' a' have y: "y \ actions E \ is_new_action (action_obs E y) \ y \ Q" by(auto elim: action_orderE) hence "a' \ y" unfolding a'_def by(rule Least_le) with y_le_a' a' show False by(auto elim: action_orderE) qed } with a' show ?thesis by blast next case False hence not_new: "\a. \ a \ Q; a \ actions E \ \ \ is_new_action (action_obs E a)" by blast show ?thesis proof(cases "Q \ actions E = {}") case True with \x \ Q\ show ?thesis by(auto elim: action_orderE) next case False define a' where "a' = (LEAST a'. a' \ Q \ a' \ actions E \ \ is_new_action (action_obs E a'))" from False obtain a where "a \ Q" "a \ actions E" by blast with not_new[OF this] have "a \ Q \ a \ actions E \ \ is_new_action (action_obs E a)" by blast hence a': "a' \ Q \ a' \ actions E \ \ is_new_action (action_obs E a')" unfolding a'_def by(rule LeastI) { fix y assume y_le_a': "(action_order E)\<^sup>\\<^sup>\ y a'" hence "y \ actions E" by(auto elim: action_orderE) have "y \ Q" proof assume "y \ Q" hence y_not_new: "\ is_new_action (action_obs E y)" using \y \ actions E\ by(rule not_new) with \y \ Q\ \y \ actions E\ have "a' \ y" unfolding a'_def by -(rule Least_le, blast) with y_le_a' y_not_new show False by(auto elim: action_orderE) qed } with a' show ?thesis by blast qed qed qed lemma action_order_is_new_actionD: "\ E \ a \a a'; is_new_action (action_obs E a') \ \ is_new_action (action_obs E a)" by(auto elim: action_orderE) subsection \Program order\ lemma program_orderI: assumes "E \ a \a a'" and "action_tid E a = action_tid E a'" shows "E \ a \po a'" using assms unfolding program_order_def by auto lemma program_orderE: assumes "E \ a \po a'" obtains t obs obs' where "E \ a \a a'" and "action_tid E a = t" "action_obs E a = obs" and "action_tid E a' = t" "action_obs E a' = obs'" using assms unfolding program_order_def by(cases "lnth E a")(cases "lnth E a'", auto simp add: action_obs_def action_tid_def) lemma refl_on_program_order: "refl_onP (actions E) (program_order E)" by(rule refl_onPI)(auto elim: action_orderE program_orderE intro: program_orderI refl_onPD[OF refl_action_order]) lemma antisym_program_order: "antisymp (program_order E)" using antisympD[OF antisym_action_order] by(auto intro: antisympI elim!: program_orderE) lemma trans_program_order: "transp (program_order E)" by(rule transpI)(auto elim!: program_orderE intro: program_orderI dest: transPD[OF trans_action_order]) lemma porder_program_order: "porder_on (actions E) (program_order E)" by(blast intro: porder_onI refl_on_program_order antisym_program_order trans_program_order) lemma total_program_order_on_tactions: "total_onP (tactions E t) (program_order E)" by(rule total_onPI)(auto elim: tactionsE simp add: program_order_def dest: total_onD[OF total_action_order]) subsection \Synchronization order\ lemma sync_orderI: "\ E \ a \a a'; a \ sactions P E; a' \ sactions P E \ \ P,E \ a \so a'" unfolding sync_order_def by blast lemma sync_orderE: assumes "P,E \ a \so a'" obtains "a \ sactions P E" "a' \ sactions P E" "E \ a \a a'" using assms unfolding sync_order_def by blast lemma refl_on_sync_order: "refl_onP (sactions P E) (sync_order P E)" by(rule refl_onPI)(fastforce elim: sync_orderE intro: sync_orderI refl_onPD[OF refl_action_order])+ lemma antisym_sync_order: "antisymp (sync_order P E)" using antisympD[OF antisym_action_order] by(rule antisympI)(auto elim!: sync_orderE) lemma trans_sync_order: "transp (sync_order P E)" by(rule transpI)(auto elim!: sync_orderE intro: sync_orderI dest: transPD[OF trans_action_order]) lemma porder_sync_order: "porder_on (sactions P E) (sync_order P E)" by(blast intro: porder_onI refl_on_sync_order antisym_sync_order trans_sync_order) lemma total_sync_order: "total_onP (sactions P E) (sync_order P E)" apply(rule total_onPI) apply(simp add: sync_order_def) apply(rule total_onPD[OF total_action_order]) apply simp_all done lemma torder_sync_order: "torder_on (sactions P E) (sync_order P E)" by(blast intro: torder_onI porder_sync_order total_sync_order) subsection \Synchronizes with\ lemma sync_withI: "\ P,E \ a \so a'; P \ (action_tid E a, action_obs E a) \sw (action_tid E a', action_obs E a') \ \ P,E \ a \sw a'" unfolding sync_with_def by blast lemma sync_withE: assumes "P,E \ a \sw a'" obtains "P,E \ a \so a'" "P \ (action_tid E a, action_obs E a) \sw (action_tid E a', action_obs E a')" using assms unfolding sync_with_def by blast lemma irrefl_synchronizes_with: "irreflP (synchronizes_with P)" by(rule irreflPI)(auto elim: synchronizes_with.cases) lemma irrefl_sync_with: "irreflP (sync_with P E)" by(rule irreflPI)(auto elim: sync_withE intro: irreflPD[OF irrefl_synchronizes_with]) lemma anitsym_sync_with: "antisymp (sync_with P E)" using antisymPD[OF antisym_sync_order, of P E] by -(rule antisymPI, auto elim: sync_withE) lemma consistent_program_order_sync_order: "order_consistent (program_order E) (sync_order P E)" apply(rule order_consistent_subset) apply(rule antisym_order_consistent_self[OF antisym_action_order[of E]]) apply(blast elim: program_orderE sync_orderE)+ done lemma consistent_program_order_sync_with: "order_consistent (program_order E) (sync_with P E)" by(rule order_consistent_subset[OF consistent_program_order_sync_order])(blast elim: sync_withE)+ subsection \Happens before\ lemma porder_happens_before: "porder_on (actions E) (happens_before P E)" unfolding po_sw_def [abs_def] by(rule porder_on_sub_torder_on_tranclp_porder_onI[OF porder_program_order torder_sync_order consistent_program_order_sync_order])(auto elim: sync_withE) lemma porder_tranclp_po_so: "porder_on (actions E) (\a a'. program_order E a a' \ sync_order P E a a')^++" by(rule porder_on_torder_on_tranclp_porder_onI[OF porder_program_order torder_sync_order consistent_program_order_sync_order]) auto lemma happens_before_refl: assumes "a \ actions E" shows "P,E \ a \hb a" using porder_happens_before[of E P] by(rule porder_onE)(erule refl_onPD[OF _ assms]) lemma happens_before_into_po_so_tranclp: assumes "P,E \ a \hb a'" shows "(\a a'. E \ a \po a' \ P,E \ a \so a')^++ a a'" using assms unfolding po_sw_def [abs_def] by(induct)(blast elim: sync_withE intro: tranclp.trancl_into_trancl)+ lemma po_sw_into_action_order: "po_sw P E a a' \ E \ a \a a'" by(auto elim: program_orderE sync_withE sync_orderE simp add: po_sw_def) lemma happens_before_into_action_order: assumes "P,E \ a \hb a'" shows "E \ a \a a'" using assms by induct(blast intro: po_sw_into_action_order transPD[OF trans_action_order])+ lemma action_order_consistent_with_happens_before: "order_consistent (action_order E) (happens_before P E)" by(blast intro: order_consistent_subset antisym_order_consistent_self antisym_action_order happens_before_into_action_order) lemma happens_before_new_actionD: assumes hb: "P,E \ a \hb a'" and new: "is_new_action (action_obs E a')" shows "is_new_action (action_obs E a)" "action_tid E a = action_tid E a'" "a \ a'" using hb proof(induct rule: converse_tranclp_induct) case (base a) case 1 from new base show ?case by(auto dest: po_sw_into_action_order elim: action_orderE) case 2 from new base show ?case by(auto simp add: po_sw_def elim!: sync_withE elim: program_orderE synchronizes_with.cases) case 3 from new base show ?case by(auto dest: po_sw_into_action_order elim: action_orderE) next case (step a a'') note po_sw = \po_sw P E a a''\ and new = \is_new_action (action_obs E a'')\ and tid = \action_tid E a'' = action_tid E a'\ case 1 from new po_sw show ?case by(auto dest: po_sw_into_action_order elim: action_orderE) case 2 from new po_sw tid show ?case by(auto simp add: po_sw_def elim!: sync_withE elim: program_orderE synchronizes_with.cases) case 3 from new po_sw \a'' \ a'\ show ?case by(auto dest!: po_sw_into_action_order elim!: action_orderE) qed lemma external_actions_not_new: "\ a \ external_actions E; is_new_action (action_obs E a) \ \ False" by(erule external_actions.cases)(simp) subsection \Most recent writes and sequential consistency\ lemma most_recent_write_for_fun: "\ P,E \ ra \mrw wa; P,E \ ra \mrw wa' \ \ wa = wa'" apply(erule most_recent_write_for.cases)+ apply clarsimp apply(erule meta_allE)+ apply(erule meta_impE) apply(rotate_tac 3) apply assumption apply(erule (1) meta_impE) apply(frule (1) action_loc_read_action_singleton) apply(rotate_tac 1) apply assumption apply(fastforce dest: antisymPD[OF antisym_action_order] elim: write_actions.cases read_actions.cases) done lemma THE_most_recent_writeI: "P,E \ r \mrw w \ (THE w. P,E \ r \mrw w) = w" by(blast dest: most_recent_write_for_fun)+ lemma most_recent_write_for_write_actionsD: assumes "P,E \ ra \mrw wa" shows "wa \ write_actions E" using assms by cases lemma most_recent_write_recent: "\ P,E \ r \mrw w; adal \ action_loc P E r; w' \ write_actions E; adal \ action_loc P E w' \ \ E \ w' \a w \ E \ r \a w'" apply(erule most_recent_write_for.cases) apply(drule (1) action_loc_read_action_singleton) apply(rotate_tac 1) apply assumption apply clarsimp done lemma is_write_seenI: "\ \a ad al v. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ ws a \ write_actions E; \a ad al v. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ (ad, al) \ action_loc P E (ws a); \a ad al v. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ value_written P E (ws a) (ad, al) = v; \a ad al v. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ \ P,E \ a \hb ws a; \a ad al v. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v); is_volatile P al \ \ \ P,E \ a \so ws a; \a ad al v a'. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v); a' \ write_actions E; (ad, al) \ action_loc P E a'; P,E \ ws a \hb a'; P,E \ a' \hb a \ \ a' = ws a; \a ad al v a'. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v); a' \ write_actions E; (ad, al) \ action_loc P E a'; is_volatile P al; P,E \ ws a \so a'; P,E \ a' \so a \ \ a' = ws a \ \ is_write_seen P E ws" unfolding is_write_seen_def by(blast 30) lemma is_write_seenD: "\ is_write_seen P E ws; a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ ws a \ write_actions E \ (ad, al) \ action_loc P E (ws a) \ value_written P E (ws a) (ad, al) = v \ \ P,E \ a \hb ws a \ (is_volatile P al \ \ P,E \ a \so ws a) \ (\a' \ write_actions E. (ad, al) \ action_loc P E a' \ (P,E \ ws a \hb a' \ P,E \ a' \hb a \ is_volatile P al \ P,E \ ws a \so a' \ P,E \ a' \so a) \ a' = ws a)" unfolding is_write_seen_def by blast lemma thread_start_actions_okI: "(\a. \ a \ actions E; \ is_new_action (action_obs E a) \ \ \i. i \ a \ action_obs E i = InitialThreadAction \ action_tid E i = action_tid E a) \ thread_start_actions_ok E" unfolding thread_start_actions_ok_def by blast lemma thread_start_actions_okD: "\ thread_start_actions_ok E; a \ actions E; \ is_new_action (action_obs E a) \ \ \i. i \ a \ action_obs E i = InitialThreadAction \ action_tid E i = action_tid E a" unfolding thread_start_actions_ok_def by blast lemma thread_start_actions_ok_prefix: "\ thread_start_actions_ok E'; lprefix E E' \ \ thread_start_actions_ok E" apply(clarsimp simp add: lprefix_conv_lappend) apply(rule thread_start_actions_okI) apply(drule_tac a=a in thread_start_actions_okD) apply(simp add: actions_def) apply(auto simp add: action_obs_def lnth_lappend1 actions_def action_tid_def le_less_trans[where y="enat a" for a]) apply (metis add.right_neutral add_strict_mono not_gr_zero) done lemma wf_execI [intro?]: "\ is_write_seen P E ws; thread_start_actions_ok E \ \ P \ (E, ws) \" by simp lemma wf_exec_is_write_seenD: "P \ (E, ws) \ \ is_write_seen P E ws" by simp lemma wf_exec_thread_start_actions_okD: "P \ (E, ws) \ \ thread_start_actions_ok E" by simp lemma sequentially_consistentI: "(\r. r \ read_actions E \ P,E \ r \mrw ws r) \ sequentially_consistent P (E, ws)" by simp lemma sequentially_consistentE: assumes "sequentially_consistent P (E, ws)" "a \ read_actions E" obtains "P,E \ a \mrw ws a" using assms by simp declare sequentially_consistent.simps [simp del] subsection \Similar actions\ text \Similar actions differ only in the values written/read.\ inductive sim_action :: "('addr, 'thread_id) obs_event action \ ('addr, 'thread_id) obs_event action \ bool" ("_ \ _" [50, 50] 51) where InitialThreadAction: "InitialThreadAction \ InitialThreadAction" | ThreadFinishAction: "ThreadFinishAction \ ThreadFinishAction" | NewHeapElem: "NormalAction (NewHeapElem a hT) \ NormalAction (NewHeapElem a hT)" | ReadMem: "NormalAction (ReadMem ad al v) \ NormalAction (ReadMem ad al v')" | WriteMem: "NormalAction (WriteMem ad al v) \ NormalAction (WriteMem ad al v')" | ThreadStart: "NormalAction (ThreadStart t) \ NormalAction (ThreadStart t)" | ThreadJoin: "NormalAction (ThreadJoin t) \ NormalAction (ThreadJoin t)" | SyncLock: "NormalAction (SyncLock a) \ NormalAction (SyncLock a)" | SyncUnlock: "NormalAction (SyncUnlock a) \ NormalAction (SyncUnlock a)" | ExternalCall: "NormalAction (ExternalCall a M vs v) \ NormalAction (ExternalCall a M vs v)" | ObsInterrupt: "NormalAction (ObsInterrupt t) \ NormalAction (ObsInterrupt t)" | ObsInterrupted: "NormalAction (ObsInterrupted t) \ NormalAction (ObsInterrupted t)" definition sim_actions :: "('addr, 'thread_id) execution \ ('addr, 'thread_id) execution \ bool" ("_ [\] _" [51, 50] 51) where "sim_actions = llist_all2 (\(t, a) (t', a'). t = t' \ a \ a')" lemma sim_action_refl [intro!, simp]: "obs \ obs" apply(cases obs) apply(rename_tac obs') apply(case_tac "obs'") apply(auto intro: sim_action.intros) done inductive_cases sim_action_cases [elim!]: "InitialThreadAction \ obs" "ThreadFinishAction \ obs" "NormalAction (NewHeapElem a hT) \ obs" "NormalAction (ReadMem ad al v) \ obs" "NormalAction (WriteMem ad al v) \ obs" "NormalAction (ThreadStart t) \ obs" "NormalAction (ThreadJoin t) \ obs" "NormalAction (SyncLock a) \ obs" "NormalAction (SyncUnlock a) \ obs" "NormalAction (ObsInterrupt t) \ obs" "NormalAction (ObsInterrupted t) \ obs" "NormalAction (ExternalCall a M vs v) \ obs" "obs \ InitialThreadAction" "obs \ ThreadFinishAction" "obs \ NormalAction (NewHeapElem a hT)" "obs \ NormalAction (ReadMem ad al v')" "obs \ NormalAction (WriteMem ad al v')" "obs \ NormalAction (ThreadStart t)" "obs \ NormalAction (ThreadJoin t)" "obs \ NormalAction (SyncLock a)" "obs \ NormalAction (SyncUnlock a)" "obs \ NormalAction (ObsInterrupt t)" "obs \ NormalAction (ObsInterrupted t)" "obs \ NormalAction (ExternalCall a M vs v)" inductive_simps sim_action_simps [simp]: "InitialThreadAction \ obs" "ThreadFinishAction \ obs" "NormalAction (NewHeapElem a hT) \ obs" "NormalAction (ReadMem ad al v) \ obs" "NormalAction (WriteMem ad al v) \ obs" "NormalAction (ThreadStart t) \ obs" "NormalAction (ThreadJoin t) \ obs" "NormalAction (SyncLock a) \ obs" "NormalAction (SyncUnlock a) \ obs" "NormalAction (ObsInterrupt t) \ obs" "NormalAction (ObsInterrupted t) \ obs" "NormalAction (ExternalCall a M vs v) \ obs" "obs \ InitialThreadAction" "obs \ ThreadFinishAction" "obs \ NormalAction (NewHeapElem a hT)" "obs \ NormalAction (ReadMem ad al v')" "obs \ NormalAction (WriteMem ad al v')" "obs \ NormalAction (ThreadStart t)" "obs \ NormalAction (ThreadJoin t)" "obs \ NormalAction (SyncLock a)" "obs \ NormalAction (SyncUnlock a)" "obs \ NormalAction (ObsInterrupt t)" "obs \ NormalAction (ObsInterrupted t)" "obs \ NormalAction (ExternalCall a M vs v)" lemma sim_action_trans [trans]: "\ obs \ obs'; obs' \ obs'' \ \ obs \ obs''" by(erule sim_action.cases) auto lemma sim_action_sym [sym]: assumes "obs \ obs'" shows "obs' \ obs" using assms by cases simp_all lemma sim_actions_sym [sym]: "E [\] E' \ E' [\] E" unfolding sim_actions_def by(auto simp add: llist_all2_conv_all_lnth split_beta intro: sim_action_sym) lemma sim_actions_action_obsD: "E [\] E' \ action_obs E a \ action_obs E' a" unfolding sim_actions_def action_obs_def by(cases "enat a < llength E")(auto dest: llist_all2_lnthD llist_all2_llengthD simp add: split_beta lnth_beyond split: enat.split) lemma sim_actions_action_tidD: "E [\] E' \ action_tid E a = action_tid E' a" unfolding sim_actions_def action_tid_def by(cases "enat a < llength E")(auto dest: llist_all2_lnthD llist_all2_llengthD simp add: lnth_beyond split: enat.split) lemma action_loc_aux_sim_action: "a \ a' \ action_loc_aux P a = action_loc_aux P a'" by(auto elim!: action_loc_aux_cases intro: action_loc_aux_intros) lemma eq_into_sim_actions: assumes "E = E'" shows "E [\] E'" unfolding sim_actions_def assms by(rule llist_all2_reflI)(auto) subsection \Well-formedness conditions for execution sets\ locale executions_base = fixes \ :: "('addr, 'thread_id) execution set" and P :: "'m prog" locale drf = executions_base \ P for \ :: "('addr, 'thread_id) execution set" and P :: "'m prog" + assumes \_new_actions_for_fun: "\ E \ \; a \ new_actions_for P E adal; a' \ new_actions_for P E adal \ \ a = a'" and \_sequential_completion: "\ E \ \; P \ (E, ws) \; \a. \ a < r; a \ read_actions E \ \ P,E \ a \mrw ws a \ \ \E' \ \. \ws'. P \ (E', ws') \ \ ltake (enat r) E = ltake (enat r) E' \ sequentially_consistent P (E', ws') \ action_tid E r = action_tid E' r \ action_obs E r \ action_obs E' r \ (r \ actions E \ r \ actions E')" locale executions_aux = executions_base \ P for \ :: "('addr, 'thread_id) execution set" and P :: "'m prog" + assumes init_before_read: "\ E \ \; P \ (E, ws) \; r \ read_actions E; adal \ action_loc P E r; \a. \ a < r; a \ read_actions E \ \ P,E \ a \mrw ws a \ \ \i new_actions_for P E adal" and \_new_actions_for_fun: "\ E \ \; a \ new_actions_for P E adal; a' \ new_actions_for P E adal \ \ a = a'" locale sc_legal = executions_aux \ P for \ :: "('addr, 'thread_id) execution set" and P :: "'m prog" + assumes \_hb_completion: "\ E \ \; P \ (E, ws) \; \a. \ a < r; a \ read_actions E \ \ P,E \ a \mrw ws a \ \ \E' \ \. \ws'. P \ (E', ws') \ \ ltake (enat r) E = ltake (enat r) E' \ (\a \ read_actions E'. if a < r then ws' a = ws a else P,E' \ ws' a \hb a) \ action_tid E' r = action_tid E r \ (if r \ read_actions E then sim_action else (=)) (action_obs E' r) (action_obs E r) \ (r \ actions E \ r \ actions E')" locale jmm_consistent = drf?: drf \ P + sc_legal \ P for \ :: "('addr, 'thread_id) execution set" and P :: "'m prog" subsection \Legal executions\ record ('addr, 'thread_id) pre_justifying_execution = committed :: "JMM_action set" justifying_exec :: "('addr, 'thread_id) execution" justifying_ws :: "write_seen" record ('addr, 'thread_id) justifying_execution = "('addr, 'thread_id) pre_justifying_execution" + action_translation :: "JMM_action \ JMM_action" type_synonym ('addr, 'thread_id) justification = "nat \ ('addr, 'thread_id) justifying_execution" definition wf_action_translation_on :: "('addr, 'thread_id) execution \ ('addr, 'thread_id) execution \ JMM_action set \ (JMM_action \ JMM_action) \ bool" where "wf_action_translation_on E E' A f \ inj_on f (actions E) \ (\a \ A. action_tid E a = action_tid E' (f a) \ action_obs E a \ action_obs E' (f a))" abbreviation wf_action_translation :: "('addr, 'thread_id) execution \ ('addr, 'thread_id) justifying_execution \ bool" where "wf_action_translation E J \ wf_action_translation_on (justifying_exec J) E (committed J) (action_translation J)" context fixes P :: "'m prog" and E :: "('addr, 'thread_id) execution" and ws :: "write_seen" and J :: "('addr, 'thread_id) justification" begin text \ This context defines the causality constraints for the JMM. The weak versions are for the fixed JMM as presented by Sevcik and Aspinall at ECOOP 2008. \ text \Committed actions are an ascending chain with all actions of E as a limit\ definition is_commit_sequence :: bool where "is_commit_sequence \ committed (J 0) = {} \ (\n. action_translation (J n) ` committed (J n) \ action_translation (J (Suc n)) ` committed (J (Suc n))) \ actions E = (\n. action_translation (J n) ` committed (J n))" definition justification_well_formed :: bool where "justification_well_formed \ (\n. P \ (justifying_exec (J n), justifying_ws (J n)) \)" definition committed_subset_actions :: bool where \ \JMM constraint 1\ "committed_subset_actions \ (\n. committed (J n) \ actions (justifying_exec (J n)))" definition happens_before_committed :: bool where \ \JMM constraint 2\ "happens_before_committed \ (\n. happens_before P (justifying_exec (J n)) |` committed (J n) = inv_imageP (happens_before P E) (action_translation (J n)) |` committed (J n))" definition happens_before_committed_weak :: bool where \ \relaxed JMM constraint\ "happens_before_committed_weak \ (\n. \r \ read_actions (justifying_exec (J n)) \ committed (J n). let r' = action_translation (J n) r; w' = ws r'; w = inv_into (actions (justifying_exec (J n))) (action_translation (J n)) w' in (P,E \ w' \hb r' \ P,justifying_exec (J n) \ w \hb r) \ \ P,justifying_exec (J n) \ r \hb w)" definition sync_order_committed :: bool where \ \JMM constraint 3\ "sync_order_committed \ (\n. sync_order P (justifying_exec (J n)) |` committed (J n) = inv_imageP (sync_order P E) (action_translation (J n)) |` committed (J n))" definition value_written_committed :: bool where \ \JMM constraint 4\ "value_written_committed \ (\n. \w \ write_actions (justifying_exec (J n)) \ committed (J n). let w' = action_translation (J n) w in (\adal \ action_loc P E w'. value_written P (justifying_exec (J n)) w adal = value_written P E w' adal))" definition write_seen_committed :: bool where \ \JMM constraint 5\ "write_seen_committed \ (\n. \r' \ read_actions (justifying_exec (J n)) \ committed (J n). let r = action_translation (J n) r'; r'' = inv_into (actions (justifying_exec (J (Suc n)))) (action_translation (J (Suc n))) r in action_translation (J (Suc n)) (justifying_ws (J (Suc n)) r'') = ws r)" text \uncommited reads see writes that happen before them -- JMM constraint 6\ (* this constraint does not apply to the 0th justifying execution, so reads may see arbitrary writes, but this cannot be exploited because reads cannot be committed in the 1st justifying execution since no writes are committed in the 0th execution *) definition uncommitted_reads_see_hb :: bool where "uncommitted_reads_see_hb \ (\n. \r' \ read_actions (justifying_exec (J (Suc n))). action_translation (J (Suc n)) r' \ action_translation (J n) ` committed (J n) \ P,justifying_exec (J (Suc n)) \ justifying_ws (J (Suc n)) r' \hb r')" text \ newly committed reads see already committed writes and write-seen relationship must not change any more -- JMM constraint 7 \ definition committed_reads_see_committed_writes :: bool where "committed_reads_see_committed_writes \ (\n. \r' \ read_actions (justifying_exec (J (Suc n))) \ committed (J (Suc n)). let r = action_translation (J (Suc n)) r'; committed_n = action_translation (J n) ` committed (J n) in r \ committed_n \ (action_translation (J (Suc n)) (justifying_ws (J (Suc n)) r') \ committed_n \ ws r \ committed_n))" definition committed_reads_see_committed_writes_weak :: bool where "committed_reads_see_committed_writes_weak \ (\n. \r' \ read_actions (justifying_exec (J (Suc n))) \ committed (J (Suc n)). let r = action_translation (J (Suc n)) r'; committed_n = action_translation (J n) ` committed (J n) in r \ committed_n \ ws r \ committed_n)" text \external actions must be committed as soon as hb-subsequent actions are committed -- JMM constraint 9\ definition external_actions_committed :: bool where "external_actions_committed \ (\n. \a \ external_actions (justifying_exec (J n)). \a' \ committed (J n). P,justifying_exec (J n) \ a \hb a' \ a \ committed (J n))" text \well-formedness conditions for action translations\ definition wf_action_translations :: bool where "wf_action_translations \ (\n. wf_action_translation_on (justifying_exec (J n)) E (committed (J n)) (action_translation (J n)))" end text \ Rule 8 of the justification for the JMM is incorrect because there might be no transitive reduction of the happens-before relation for an infinite execution, if infinitely many initialisation actions have to be ordered before the start action of every thread. Hence, \is_justified_by\ omits this constraint. \ primrec is_justified_by :: "'m prog \ ('addr, 'thread_id) execution \ write_seen \ ('addr, 'thread_id) justification \ bool" ("_ \ _ justified'_by _" [51, 50, 50] 50) where "P \ (E, ws) justified_by J \ is_commit_sequence E J \ justification_well_formed P J \ committed_subset_actions J \ happens_before_committed P E J \ sync_order_committed P E J \ value_written_committed P E J \ write_seen_committed ws J \ uncommitted_reads_see_hb P J \ committed_reads_see_committed_writes ws J \ external_actions_committed P J \ wf_action_translations E J" text \ Sevcik requires in the fixed JMM that external actions may only be committed when everything that happens before has already been committed. On the level of legality, this constraint is vacuous because it is always possible to delay committing external actions, so we omit it here. \ primrec is_weakly_justified_by :: "'m prog \ ('addr, 'thread_id) execution \ write_seen \ ('addr, 'thread_id) justification \ bool" ("_ \ _ weakly'_justified'_by _" [51, 50, 50] 50) where "P \ (E, ws) weakly_justified_by J \ is_commit_sequence E J \ justification_well_formed P J \ committed_subset_actions J \ happens_before_committed_weak P E ws J \ \ \no \sync_order\ constraint\ value_written_committed P E J \ write_seen_committed ws J \ uncommitted_reads_see_hb P J \ committed_reads_see_committed_writes_weak ws J \ wf_action_translations E J" text \ Notion of conflict is strengthened to explicitly exclude volatile locations. Otherwise, the following program is not correctly synchronised: \begin{verbatim} volatile x = 0; --------------- r = x; | x = 1; \end{verbatim} because in the SC execution [Init x 0, (t1, Read x 0), (t2, Write x 1)], the read and write are unrelated in hb, because synchronises-with is asymmetric for volatiles. The JLS considers conflicting volatiles for data races, but this is only a remark on the DRF guarantee. See JMM mailing list posts \#2477 to 2488. \ (* Problem already exists in Sevcik's formalisation *) definition non_volatile_conflict :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_,_ \/(_)\(_)" [51,50,50,50] 51) where "P,E \ a \ a' \ (a \ read_actions E \ a' \ write_actions E \ a \ write_actions E \ a' \ read_actions E \ a \ write_actions E \ a' \ write_actions E) \ (\ad al. (ad, al) \ action_loc P E a \ action_loc P E a' \ \ is_volatile P al)" definition correctly_synchronized :: "'m prog \ ('addr, 'thread_id) execution set \ bool" where "correctly_synchronized P \ \ (\E \ \. \ws. P \ (E, ws) \ \ sequentially_consistent P (E, ws) \ (\a \ actions E. \a' \ actions E. P,E \ a \ a' \ P,E \ a \hb a' \ P,E \ a' \hb a))" primrec gen_legal_execution :: "('m prog \ ('addr, 'thread_id) execution \ write_seen \ ('addr, 'thread_id) justification \ bool) \ 'm prog \ ('addr, 'thread_id) execution set \ ('addr, 'thread_id) execution \ write_seen \ bool" where "gen_legal_execution is_justification P \ (E, ws) \ E \ \ \ P \ (E, ws) \ \ (\J. is_justification P (E, ws) J \ range (justifying_exec \ J) \ \)" abbreviation legal_execution :: "'m prog \ ('addr, 'thread_id) execution set \ ('addr, 'thread_id) execution \ write_seen \ bool" where "legal_execution \ gen_legal_execution is_justified_by" abbreviation weakly_legal_execution :: "'m prog \ ('addr, 'thread_id) execution set \ ('addr, 'thread_id) execution \ write_seen \ bool" where "weakly_legal_execution \ gen_legal_execution is_weakly_justified_by" declare gen_legal_execution.simps [simp del] lemma sym_non_volatile_conflict: "symP (non_volatile_conflict P E)" unfolding non_volatile_conflict_def by(rule symPI) blast lemma legal_executionI: "\ E \ \; P \ (E, ws) \; is_justification P (E, ws) J; range (justifying_exec \ J) \ \ \ \ gen_legal_execution is_justification P \ (E, ws)" unfolding gen_legal_execution.simps by blast lemma legal_executionE: assumes "gen_legal_execution is_justification P \ (E, ws)" obtains J where "E \ \" "P \ (E, ws) \" "is_justification P (E, ws) J" "range (justifying_exec \ J) \ \" using assms unfolding gen_legal_execution.simps by blast lemma legal_\D: "gen_legal_execution is_justification P \ (E, ws) \ E \ \" by(erule legal_executionE) lemma legal_wf_execD: "gen_legal_execution is_justification P \ Ews \ P \ Ews \" by(cases Ews)(auto elim: legal_executionE) lemma correctly_synchronizedD: "\ correctly_synchronized P \; E \ \; P \ (E, ws) \; sequentially_consistent P (E, ws) \ \ \a a'. a \ actions E \ a' \ actions E \ P,E \ a \ a' \ P,E \ a \hb a' \ P,E \ a' \hb a" unfolding correctly_synchronized_def by blast lemma wf_action_translation_on_actionD: "\ wf_action_translation_on E E' A f; a \ A \ \ action_tid E a = action_tid E' (f a) \ action_obs E a \ action_obs E' (f a)" unfolding wf_action_translation_on_def by blast lemma wf_action_translation_on_inj_onD: "wf_action_translation_on E E' A f \ inj_on f (actions E)" unfolding wf_action_translation_on_def by simp lemma wf_action_translation_on_action_locD: "\ wf_action_translation_on E E' A f; a \ A \ \ action_loc P E a = action_loc P E' (f a)" apply(drule (1) wf_action_translation_on_actionD) apply(cases "(P, action_obs E a)" rule: action_loc_aux.cases) apply auto done lemma weakly_justified_write_seen_hb_read_committed: assumes J: "P \ (E, ws) weakly_justified_by J" and r: "r \ read_actions (justifying_exec (J n))" "r \ committed (J n)" shows "ws (action_translation (J n) r) \ action_translation (J n) ` committed (J n)" using r proof(induct n arbitrary: r) case 0 from J have [simp]: "committed (J 0) = {}" by(simp add: is_commit_sequence_def) with 0 show ?case by simp next case (Suc n) let ?E = "\n. justifying_exec (J n)" and ?ws = "\n. justifying_ws (J n)" and ?C = "\n. committed (J n)" and ?\ = "\n. action_translation (J n)" note r = \r \ read_actions (?E (Suc n))\ hence "r \ actions (?E (Suc n))" by simp from J have wfan: "wf_action_translation_on (?E n) E (?C n) (?\ n)" and wfaSn: "wf_action_translation_on (?E (Suc n)) E (?C (Suc n)) (?\ (Suc n))" by(simp_all add: wf_action_translations_def) from wfaSn have injSn: "inj_on (?\ (Suc n)) (actions (?E (Suc n)))" by(rule wf_action_translation_on_inj_onD) from J have C_sub_A: "?C (Suc n) \ actions (?E (Suc n))" by(simp add: committed_subset_actions_def) from J have CnCSn: "?\ n ` ?C n \ ?\ (Suc n) ` ?C (Suc n)" by(simp add: is_commit_sequence_def) from J have wsSn: "is_write_seen P (?E (Suc n)) (?ws (Suc n))" by(simp add: justification_well_formed_def) from r obtain ad al v where "action_obs (?E (Suc n)) r = NormalAction (ReadMem ad al v)" by cases from is_write_seenD[OF wsSn r this] have wsSn: "?ws (Suc n) r \ actions (?E (Suc n))" by simp show ?case proof(cases "?\ (Suc n) r \ ?\ n ` ?C n") case True then obtain r' where r': "r' \ ?C n" and r_r': "?\ (Suc n) r = ?\ n r'" by(auto) from r' wfan have "action_tid (?E n) r' = action_tid E (?\ n r')" and "action_obs (?E n) r' \ action_obs E (?\ n r')" by(blast dest: wf_action_translation_on_actionD)+ moreover from r' CnCSn have "?\ (Suc n) r \ ?\ (Suc n) ` ?C (Suc n)" unfolding r_r' by auto hence "r \ ?C (Suc n)" - unfolding inj_on_image_mem_iff[OF injSn C_sub_A \r \ actions (?E (Suc n))\] . + unfolding inj_on_image_mem_iff[OF injSn \r \ actions (?E (Suc n))\ C_sub_A] . with wfaSn have "action_tid (?E (Suc n)) r = action_tid E (?\ (Suc n) r)" and "action_obs (?E (Suc n)) r \ action_obs E (?\ (Suc n) r)" by(blast dest: wf_action_translation_on_actionD)+ ultimately have tid: "action_tid (?E n) r' = action_tid (?E (Suc n)) r" and obs: "action_obs (?E n) r' \ action_obs (?E (Suc n)) r" unfolding r_r' by(auto intro: sim_action_trans sim_action_sym) from J have "?C n \ actions (?E n)" by(simp add: committed_subset_actions_def) with r' have "r' \ actions (?E n)" by blast with r obs have "r' \ read_actions (?E n)" by cases(auto intro: read_actions.intros) hence ws: "ws (?\ n r') \ ?\ n ` ?C n" using r' by(rule Suc) have r_conv_inv: "r = inv_into (actions (?E (Suc n))) (?\ (Suc n)) (?\ n r')" using \r \ actions (?E (Suc n))\ unfolding r_r'[symmetric] by(simp add: inv_into_f_f[OF injSn]) with \r' \ ?C n\ r J \r' \ read_actions (?E n)\ have ws_eq: "?\ (Suc n) (?ws (Suc n) r) = ws (?\ n r')" by(simp add: Let_def write_seen_committed_def) with ws CnCSn have "?\ (Suc n) (?ws (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" by auto hence "?ws (Suc n) r \ ?C (Suc n)" - by(subst (asm) inj_on_image_mem_iff[OF injSn C_sub_A wsSn]) + by(subst (asm) inj_on_image_mem_iff[OF injSn wsSn C_sub_A]) moreover from ws CnCSn have "ws (?\ (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" unfolding r_r' by auto ultimately show ?thesis by simp next case False with r \r \ ?C (Suc n)\ J have "ws (?\ (Suc n) r) \ ?\ n ` ?C n" unfolding is_weakly_justified_by.simps Let_def committed_reads_see_committed_writes_weak_def by blast hence "ws (?\ (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" using CnCSn by blast+ - thus ?thesis by(simp add: inj_on_image_mem_iff[OF injSn C_sub_A wsSn]) + thus ?thesis by(simp add: inj_on_image_mem_iff[OF injSn wsSn C_sub_A]) qed qed lemma justified_write_seen_hb_read_committed: assumes J: "P \ (E, ws) justified_by J" and r: "r \ read_actions (justifying_exec (J n))" "r \ committed (J n)" shows "justifying_ws (J n) r \ committed (J n)" (is ?thesis1) and "ws (action_translation (J n) r) \ action_translation (J n) ` committed (J n)" (is ?thesis2) proof - have "?thesis1 \ ?thesis2" using r proof(induct n arbitrary: r) case 0 from J have [simp]: "committed (J 0) = {}" by(simp add: is_commit_sequence_def) with 0 show ?case by simp next case (Suc n) let ?E = "\n. justifying_exec (J n)" and ?ws = "\n. justifying_ws (J n)" and ?C = "\n. committed (J n)" and ?\ = "\n. action_translation (J n)" note r = \r \ read_actions (?E (Suc n))\ hence "r \ actions (?E (Suc n))" by simp from J have wfan: "wf_action_translation_on (?E n) E (?C n) (?\ n)" and wfaSn: "wf_action_translation_on (?E (Suc n)) E (?C (Suc n)) (?\ (Suc n))" by(simp_all add: wf_action_translations_def) from wfaSn have injSn: "inj_on (?\ (Suc n)) (actions (?E (Suc n)))" by(rule wf_action_translation_on_inj_onD) from J have C_sub_A: "?C (Suc n) \ actions (?E (Suc n))" by(simp add: committed_subset_actions_def) from J have CnCSn: "?\ n ` ?C n \ ?\ (Suc n) ` ?C (Suc n)" by(simp add: is_commit_sequence_def) from J have wsSn: "is_write_seen P (?E (Suc n)) (?ws (Suc n))" by(simp add: justification_well_formed_def) from r obtain ad al v where "action_obs (?E (Suc n)) r = NormalAction (ReadMem ad al v)" by cases from is_write_seenD[OF wsSn r this] have wsSn: "?ws (Suc n) r \ actions (?E (Suc n))" by simp show ?case proof(cases "?\ (Suc n) r \ ?\ n ` ?C n") case True then obtain r' where r': "r' \ ?C n" and r_r': "?\ (Suc n) r = ?\ n r'" by(auto) from r' wfan have "action_tid (?E n) r' = action_tid E (?\ n r')" and "action_obs (?E n) r' \ action_obs E (?\ n r')" by(blast dest: wf_action_translation_on_actionD)+ moreover from r' CnCSn have "?\ (Suc n) r \ ?\ (Suc n) ` ?C (Suc n)" unfolding r_r' by auto hence "r \ ?C (Suc n)" - unfolding inj_on_image_mem_iff[OF injSn C_sub_A \r \ actions (?E (Suc n))\] . + unfolding inj_on_image_mem_iff[OF injSn \r \ actions (?E (Suc n))\ C_sub_A] . with wfaSn have "action_tid (?E (Suc n)) r = action_tid E (?\ (Suc n) r)" and "action_obs (?E (Suc n)) r \ action_obs E (?\ (Suc n) r)" by(blast dest: wf_action_translation_on_actionD)+ ultimately have tid: "action_tid (?E n) r' = action_tid (?E (Suc n)) r" and obs: "action_obs (?E n) r' \ action_obs (?E (Suc n)) r" unfolding r_r' by(auto intro: sim_action_trans sim_action_sym) from J have "?C n \ actions (?E n)" by(simp add: committed_subset_actions_def) with r' have "r' \ actions (?E n)" by blast with r obs have "r' \ read_actions (?E n)" by cases(auto intro: read_actions.intros) hence "?ws n r' \ ?C n \ ws (?\ n r') \ ?\ n ` ?C n" using r' by(rule Suc) then obtain ws: "ws (?\ n r') \ ?\ n ` ?C n" .. have r_conv_inv: "r = inv_into (actions (?E (Suc n))) (?\ (Suc n)) (?\ n r')" using \r \ actions (?E (Suc n))\ unfolding r_r'[symmetric] by(simp add: inv_into_f_f[OF injSn]) with \r' \ ?C n\ r J \r' \ read_actions (?E n)\ have ws_eq: "?\ (Suc n) (?ws (Suc n) r) = ws (?\ n r')" by(simp add: Let_def write_seen_committed_def) with ws CnCSn have "?\ (Suc n) (?ws (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" by auto hence "?ws (Suc n) r \ ?C (Suc n)" - by(subst (asm) inj_on_image_mem_iff[OF injSn C_sub_A wsSn]) + by(subst (asm) inj_on_image_mem_iff[OF injSn wsSn C_sub_A]) moreover from ws CnCSn have "ws (?\ (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" unfolding r_r' by auto ultimately show ?thesis by simp next case False with r \r \ ?C (Suc n)\ J have "?\ (Suc n) (?ws (Suc n) r) \ ?\ n ` ?C n" and "ws (?\ (Suc n) r) \ ?\ n ` ?C n" unfolding is_justified_by.simps Let_def committed_reads_see_committed_writes_def by blast+ hence "?\ (Suc n) (?ws (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" and "ws (?\ (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" using CnCSn by blast+ - thus ?thesis by(simp add: inj_on_image_mem_iff[OF injSn C_sub_A wsSn]) + thus ?thesis by(simp add: inj_on_image_mem_iff[OF injSn wsSn C_sub_A]) qed qed thus ?thesis1 ?thesis2 by simp_all qed lemma is_justified_by_imp_is_weakly_justified_by: assumes justified: "P \ (E, ws) justified_by J" and wf: "P \ (E, ws) \" shows "P \ (E, ws) weakly_justified_by J" unfolding is_weakly_justified_by.simps proof(intro conjI) let ?E = "\n. justifying_exec (J n)" and ?ws = "\n. justifying_ws (J n)" and ?C = "\n. committed (J n)" and ?\ = "\n. action_translation (J n)" from justified show "is_commit_sequence E J" "justification_well_formed P J" "committed_subset_actions J" "value_written_committed P E J" "write_seen_committed ws J" "uncommitted_reads_see_hb P J" "wf_action_translations E J" by(simp_all) show "happens_before_committed_weak P E ws J" unfolding happens_before_committed_weak_def Let_def proof(intro strip conjI) fix n r assume "r \ read_actions (?E n) \ ?C n" hence read: "r \ read_actions (?E n)" and committed: "r \ ?C n" by simp_all with justified have committed_ws: "?ws n r \ ?C n" and committed_ws': "ws (?\ n r) \ ?\ n ` ?C n" by(rule justified_write_seen_hb_read_committed)+ from committed_ws' obtain w where w: "ws (?\ n r) = ?\ n w" and committed_w: "w \ ?C n" by blast from committed_w justified have "w \ actions (?E n)" by(auto simp add: committed_subset_actions_def) moreover from justified have "inj_on (?\ n) (actions (?E n))" by(auto simp add: wf_action_translations_def dest: wf_action_translation_on_inj_onD) ultimately have w_def: "w = inv_into (actions (?E n)) (?\ n) (ws (?\ n r))" by(simp_all add: w) from committed committed_w have "P,?E n \ w \hb r \ (happens_before P (?E n) |` ?C n) w r" by auto also have "\ \ (inv_imageP (happens_before P E) (?\ n) |` ?C n) w r" using justified by(simp add: happens_before_committed_def) also have "\ \ P,E \ ?\ n w \hb ?\ n r" using committed committed_w by auto finally show "P,E \ ws (?\ n r) \hb ?\ n r \ P,?E n \ inv_into (actions (?E n)) (?\ n) (ws (?\ n r)) \hb r" unfolding w[symmetric] unfolding w_def .. have "P,?E n\ r \hb w \ (happens_before P (?E n) |` ?C n) r w" using committed committed_w by auto also have "\ \ (inv_imageP (happens_before P E) (?\ n) |` ?C n) r w" using justified by(simp add: happens_before_committed_def) also have "\ \ P,E \ ?\ n r \hb ws (?\ n r)" using w committed committed_w by auto also { from read obtain ad al v where "action_obs (?E n) r = NormalAction (ReadMem ad al v)" by cases auto with justified committed obtain v' where obs': "action_obs E (?\ n r) = NormalAction (ReadMem ad al v')" by(fastforce simp add: wf_action_translations_def dest!: wf_action_translation_on_actionD) moreover from committed justified have "?\ n r \ actions E" by(auto simp add: is_commit_sequence_def) ultimately have read': "?\ n r \ read_actions E" by(auto intro: read_actions.intros) from wf have "is_write_seen P E ws" by(rule wf_exec_is_write_seenD) from is_write_seenD[OF this read' obs'] have "\ P,E \ ?\ n r \hb ws (?\ n r)" by simp } ultimately show "\ P,?E n \ r \hb inv_into (actions (?E n)) (?\ n) (ws (?\ n r))" unfolding w_def by simp qed from justified have "committed_reads_see_committed_writes ws J" by simp thus "committed_reads_see_committed_writes_weak ws J" by(auto simp add: committed_reads_see_committed_writes_def committed_reads_see_committed_writes_weak_def) qed corollary legal_imp_weakly_legal_execution: "legal_execution P \ Ews \ weakly_legal_execution P \ Ews" by(cases Ews)(auto 4 4 simp add: gen_legal_execution.simps simp del: is_justified_by.simps is_weakly_justified_by.simps intro: is_justified_by_imp_is_weakly_justified_by) lemma drop_0th_justifying_exec: assumes "P \ (E, ws) justified_by J" and wf: "P \ (E', ws') \" shows "P \ (E, ws) justified_by (J(0 := \committed = {}, justifying_exec = E', justifying_ws = ws', action_translation = id\))" (is "_ \ _ justified_by ?J") using assms unfolding is_justified_by.simps is_commit_sequence_def justification_well_formed_def committed_subset_actions_def happens_before_committed_def sync_order_committed_def value_written_committed_def write_seen_committed_def uncommitted_reads_see_hb_def committed_reads_see_committed_writes_def external_actions_committed_def wf_action_translations_def proof(intro conjI strip) let ?E = "\n. justifying_exec (?J n)" and ?\ = "\n. action_translation (?J n)" and ?C = "\n. committed (?J n)" and ?ws = "\n. justifying_ws (?J n)" show "?C 0 = {}" by simp from assms have C_0: "committed (J 0) = {}" by(simp add: is_commit_sequence_def) hence "(\n. ?\ n ` ?C n) = (\n. action_translation (J n) ` committed (J n))" by -(rule SUP_cong, simp_all) also have "\ = actions E" using assms by(simp add: is_commit_sequence_def) finally show "actions E = (\n. ?\ n ` ?C n)" .. fix n { fix r' assume "r' \ read_actions (?E (Suc n))" thus "?\ (Suc n) r' \ ?\ n ` ?C n \ P,?E (Suc n) \ ?ws (Suc n) r' \hb r'" using assms by(auto dest!: bspec simp add: uncommitted_reads_see_hb_def is_commit_sequence_def) } { fix r' assume r': "r' \ read_actions (?E (Suc n)) \ ?C (Suc n)" have "n \ 0" proof assume "n = 0" hence "r' \ read_actions (justifying_exec (J 1)) \ committed (J 1)" using r' by simp hence "action_translation (J 1) r' \ action_translation (J 0) ` committed (J 0) \ ws (action_translation (J 1) r') \ action_translation (J 0) ` committed (J 0)" using assms unfolding One_nat_def is_justified_by.simps Let_def committed_reads_see_committed_writes_def by(metis (lifting)) thus False unfolding C_0 by simp qed thus "let r = ?\ (Suc n) r'; committed_n = ?\ n ` ?C n in r \ committed_n \ (?\ (Suc n) (?ws (Suc n) r') \ committed_n \ ws r \ committed_n)" using assms r' by(simp add: committed_reads_see_committed_writes_def) } { fix a a' assume "a \ external_actions (?E n)" and "a' \ ?C n" "P,?E n \ a \hb a'" moreover hence "n > 0" by(simp split: if_split_asm) ultimately show "a \ ?C n" using assms by(simp add: external_actions_committed_def) blast } from assms have "wf_action_translation E (?J 0)" by(simp add: wf_action_translations_def wf_action_translation_on_def) thus "wf_action_translation E (?J n)" using assms by(simp add: wf_action_translations_def) qed auto lemma drop_0th_weakly_justifying_exec: assumes "P \ (E, ws) weakly_justified_by J" and wf: "P \ (E', ws') \" shows "P \ (E, ws) weakly_justified_by (J(0 := \committed = {}, justifying_exec = E', justifying_ws = ws', action_translation = id\))" (is "_ \ _ weakly_justified_by ?J") using assms unfolding is_weakly_justified_by.simps is_commit_sequence_def justification_well_formed_def committed_subset_actions_def happens_before_committed_weak_def value_written_committed_def write_seen_committed_def uncommitted_reads_see_hb_def committed_reads_see_committed_writes_weak_def external_actions_committed_def wf_action_translations_def proof(intro conjI strip) let ?E = "\n. justifying_exec (?J n)" and ?\ = "\n. action_translation (?J n)" and ?C = "\n. committed (?J n)" and ?ws = "\n. justifying_ws (?J n)" show "?C 0 = {}" by simp from assms have C_0: "committed (J 0) = {}" by(simp add: is_commit_sequence_def) hence "(\n. ?\ n ` ?C n) = (\n. action_translation (J n) ` committed (J n))" by -(rule SUP_cong, simp_all) also have "\ = actions E" using assms by(simp add: is_commit_sequence_def) finally show "actions E = (\n. ?\ n ` ?C n)" .. fix n { fix r' assume "r' \ read_actions (?E (Suc n))" thus "?\ (Suc n) r' \ ?\ n ` ?C n \ P,?E (Suc n) \ ?ws (Suc n) r' \hb r'" using assms by(auto dest!: bspec simp add: uncommitted_reads_see_hb_def is_commit_sequence_def) } { fix r' assume r': "r' \ read_actions (?E (Suc n)) \ ?C (Suc n)" have "n \ 0" proof assume "n = 0" hence "r' \ read_actions (justifying_exec (J 1)) \ committed (J 1)" using r' by simp hence "action_translation (J 1) r' \ action_translation (J 0) ` committed (J 0) \ ws (action_translation (J 1) r') \ action_translation (J 0) ` committed (J 0)" using assms unfolding One_nat_def is_weakly_justified_by.simps Let_def committed_reads_see_committed_writes_weak_def by(metis (lifting)) thus False unfolding C_0 by simp qed thus "let r = ?\ (Suc n) r'; committed_n = ?\ n ` ?C n in r \ committed_n \ ws r \ committed_n" using assms r' by(simp add: committed_reads_see_committed_writes_weak_def) } from assms have "wf_action_translation E (?J 0)" by(simp add: wf_action_translations_def wf_action_translation_on_def) thus "wf_action_translation E (?J n)" using assms by(simp add: wf_action_translations_def) qed auto subsection \Executions with common prefix\ lemma actions_change_prefix: assumes read: "a \ actions E" and prefix: "ltake n E [\] ltake n E'" and rn: "enat a < n" shows "a \ actions E'" using llist_all2_llengthD[OF prefix[unfolded sim_actions_def]] read rn by(simp add: actions_def min_def split: if_split_asm) lemma action_obs_change_prefix: assumes prefix: "ltake n E [\] ltake n E'" and rn: "enat a < n" shows "action_obs E a \ action_obs E' a" proof - from rn have "action_obs E a = action_obs (ltake n E) a" by(simp add: action_obs_def lnth_ltake) also from prefix have "\ \ action_obs (ltake n E') a" by(rule sim_actions_action_obsD) also have "\ = action_obs E' a" using rn by(simp add: action_obs_def lnth_ltake) finally show ?thesis . qed lemma action_obs_change_prefix_eq: assumes prefix: "ltake n E = ltake n E'" and rn: "enat a < n" shows "action_obs E a = action_obs E' a" proof - from rn have "action_obs E a = action_obs (ltake n E) a" by(simp add: action_obs_def lnth_ltake) also from prefix have "\ = action_obs (ltake n E') a" by(simp add: action_obs_def) also have "\ = action_obs E' a" using rn by(simp add: action_obs_def lnth_ltake) finally show ?thesis . qed lemma read_actions_change_prefix: assumes read: "r \ read_actions E" and prefix: "ltake n E [\] ltake n E'" "enat r < n" shows "r \ read_actions E'" using read action_obs_change_prefix[OF prefix] actions_change_prefix[OF _ prefix] by(cases)(auto intro: read_actions.intros) lemma sim_action_is_write_action_eq: assumes "obs \ obs'" shows "is_write_action obs \ is_write_action obs'" using assms by cases simp_all lemma write_actions_change_prefix: assumes "write": "w \ write_actions E" and prefix: "ltake n E [\] ltake n E'" "enat w < n" shows "w \ write_actions E'" using "write" action_obs_change_prefix[OF prefix] actions_change_prefix[OF _ prefix] by(cases)(auto intro: write_actions.intros dest: sim_action_is_write_action_eq) lemma action_loc_change_prefix: assumes "ltake n E [\] ltake n E'" "enat a < n" shows "action_loc P E a = action_loc P E' a" using action_obs_change_prefix[OF assms] by(fastforce elim!: action_loc_aux_cases intro: action_loc_aux_intros) lemma sim_action_is_new_action_eq: assumes "obs \ obs'" shows "is_new_action obs = is_new_action obs'" using assms by cases auto lemma action_order_change_prefix: assumes ao: "E \ a \a a'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "E' \ a \a a'" using ao actions_change_prefix[OF _ prefix an] actions_change_prefix[OF _ prefix a'n] action_obs_change_prefix[OF prefix an] action_obs_change_prefix[OF prefix a'n] by(auto simp add: action_order_def split: if_split_asm dest: sim_action_is_new_action_eq) lemma value_written_change_prefix: assumes eq: "ltake n E = ltake n E'" and an: "enat a < n" shows "value_written P E a = value_written P E' a" using action_obs_change_prefix_eq[OF eq an] by(simp add: value_written_def fun_eq_iff) lemma action_tid_change_prefix: assumes prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" shows "action_tid E a = action_tid E' a" proof - from an have "action_tid E a = action_tid (ltake n E) a" by(simp add: action_tid_def lnth_ltake) also from prefix have "\ = action_tid (ltake n E') a" by(rule sim_actions_action_tidD) also from an have "\ = action_tid E' a" by(simp add: action_tid_def lnth_ltake) finally show ?thesis . qed lemma program_order_change_prefix: assumes po: "E \ a \po a'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "E' \ a \po a'" using po action_order_change_prefix[OF _ prefix an a'n] action_tid_change_prefix[OF prefix an] action_tid_change_prefix[OF prefix a'n] by(auto elim!: program_orderE intro: program_orderI) lemma sim_action_sactionD: assumes "obs \ obs'" shows "saction P obs \ saction P obs'" using assms by cases simp_all lemma sactions_change_prefix: assumes sync: "a \ sactions P E" and prefix: "ltake n E [\] ltake n E'" and rn: "enat a < n" shows "a \ sactions P E'" using sync action_obs_change_prefix[OF prefix rn] actions_change_prefix[OF _ prefix rn] unfolding sactions_def by(simp add: sim_action_sactionD) lemma sync_order_change_prefix: assumes so: "P,E \ a \so a'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "P,E' \ a \so a'" using so action_order_change_prefix[OF _ prefix an a'n] sactions_change_prefix[OF _ prefix an, of P] sactions_change_prefix[OF _ prefix a'n, of P] by(simp add: sync_order_def) lemma sim_action_synchronizes_withD: assumes "obs \ obs'" "obs'' \ obs'''" shows "P \ (t, obs) \sw (t', obs'') \ P \ (t, obs') \sw (t', obs''')" using assms by(auto elim!: sim_action.cases synchronizes_with.cases intro: synchronizes_with.intros) lemma sync_with_change_prefix: assumes sw: "P,E \ a \sw a'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "P,E' \ a \sw a'" using sw sync_order_change_prefix[OF _ prefix an a'n, of P] action_tid_change_prefix[OF prefix an] action_tid_change_prefix[OF prefix a'n] action_obs_change_prefix[OF prefix an] action_obs_change_prefix[OF prefix a'n] by(auto simp add: sync_with_def dest: sim_action_synchronizes_withD) lemma po_sw_change_prefix: assumes posw: "po_sw P E a a'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "po_sw P E' a a'" using posw sync_with_change_prefix[OF _ prefix an a'n, of P] program_order_change_prefix[OF _ prefix an a'n] by(auto simp add: po_sw_def) lemma happens_before_new_not_new: assumes tsa_ok: "thread_start_actions_ok E" and a: "a \ actions E" and a': "a' \ actions E" and new_a: "is_new_action (action_obs E a)" and new_a': "\ is_new_action (action_obs E a')" shows "P,E \ a \hb a'" proof - from thread_start_actions_okD[OF tsa_ok a' new_a'] obtain i where "i \ a'" and obs_i: "action_obs E i = InitialThreadAction" and "action_tid E i = action_tid E a'" by auto from \i \ a'\ a' have "i \ actions E" by(auto simp add: actions_def le_less_trans[where y="enat a'"]) with \i \ a'\ obs_i a' new_a' have "E \ i \a a'" by(simp add: action_order_def) hence "E \ i \po a'" using \action_tid E i = action_tid E a'\ by(rule program_orderI) moreover { from \i \ actions E\ obs_i have "i \ sactions P E" by(auto intro: sactionsI) from a \i \ actions E\ new_a obs_i have "E \ a \a i" by(simp add: action_order_def) moreover from a new_a have "a \ sactions P E" by(auto intro: sactionsI) ultimately have "P,E \ a \so i" using \i \ sactions P E\ by(rule sync_orderI) moreover from new_a obs_i have "P \ (action_tid E a, action_obs E a) \sw (action_tid E i, action_obs E i)" by cases(auto intro: synchronizes_with.intros) ultimately have "P,E \ a \sw i" by(rule sync_withI) } ultimately show ?thesis unfolding po_sw_def [abs_def] by(blast intro: tranclp.r_into_trancl tranclp_trans) qed lemma happens_before_change_prefix: assumes hb: "P,E \ a \hb a'" and tsa_ok: "thread_start_actions_ok E'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "P,E' \ a \hb a'" using hb an a'n proof induct case (base a') thus ?case by(rule tranclp.r_into_trancl[where r="po_sw P E'", OF po_sw_change_prefix[OF _ prefix]]) next case (step a' a'') show ?case proof(cases "is_new_action (action_obs E a') \ \ is_new_action (action_obs E a'')") case False from \po_sw P E a' a''\ have "E \ a' \a a''" by(rule po_sw_into_action_order) with \enat a'' < n\ False have "enat a' < n" by(safe elim!: action_orderE)(metis Suc_leI Suc_n_not_le_n enat_ord_simps(2) le_trans nat_neq_iff xtrans(10))+ with \enat a < n\ have "P,E' \ a \hb a'" by(rule step) moreover from \po_sw P E a' a''\ prefix \enat a' < n\ \enat a'' < n\ have "po_sw P E' a' a''" by(rule po_sw_change_prefix) ultimately show ?thesis .. next case True then obtain new_a': "is_new_action (action_obs E a')" and "\ is_new_action (action_obs E a'')" .. from \P,E \ a \hb a'\ new_a' have new_a: "is_new_action (action_obs E a)" and tid: "action_tid E a = action_tid E a'" and "a \ a'" by(rule happens_before_new_actionD)+ note tsa_ok moreover from porder_happens_before[of E P] have "a \ actions E" by(rule porder_onE)(erule refl_onPD1, rule \P,E \ a \hb a'\) hence "a \ actions E'" using an by(rule actions_change_prefix[OF _ prefix]) moreover from \po_sw P E a' a''\ refl_on_program_order[of E] refl_on_sync_order[of P E] have "a'' \ actions E" unfolding po_sw_def by(auto dest: refl_onPD2 elim!: sync_withE) hence "a'' \ actions E'" using \enat a'' < n\ by(rule actions_change_prefix[OF _ prefix]) moreover from new_a action_obs_change_prefix[OF prefix an] have "is_new_action (action_obs E' a)" by(cases) auto moreover from \\ is_new_action (action_obs E a'')\ action_obs_change_prefix[OF prefix \enat a'' < n\] have "\ is_new_action (action_obs E' a'')" by(auto elim: is_new_action.cases) ultimately show "P,E' \ a \hb a''" by(rule happens_before_new_not_new) qed qed lemma thread_start_actions_ok_change: assumes tsa: "thread_start_actions_ok E" and sim: "E [\] E'" shows "thread_start_actions_ok E'" proof(rule thread_start_actions_okI) fix a assume "a \ actions E'" "\ is_new_action (action_obs E' a)" from sim have len_eq: "llength E = llength E'" by(simp add: sim_actions_def)(rule llist_all2_llengthD) with sim have sim': "ltake (llength E) E [\] ltake (llength E) E'" by(simp add: ltake_all) from \a \ actions E'\ len_eq have "enat a < llength E" by(simp add: actions_def) with \a \ actions E'\ sim'[symmetric] have "a \ actions E" by(rule actions_change_prefix) moreover have "\ is_new_action (action_obs E a)" using action_obs_change_prefix[OF sim' \enat a < llength E\] \\ is_new_action (action_obs E' a)\ by(auto elim!: is_new_action.cases) ultimately obtain i where "i \ a" "action_obs E i = InitialThreadAction" "action_tid E i = action_tid E a" by(blast dest: thread_start_actions_okD[OF tsa]) thus "\i \ a. action_obs E' i = InitialThreadAction \ action_tid E' i = action_tid E' a" using action_tid_change_prefix[OF sim', of i] action_tid_change_prefix[OF sim', of a] \enat a < llength E\ action_obs_change_prefix[OF sim', of i] by(cases "llength E")(auto intro!: exI[where x=i]) qed context executions_aux begin lemma \_new_same_addr_singleton: assumes E: "E \ \" shows "\a. new_actions_for P E adal \ {a}" by(blast dest: \_new_actions_for_fun[OF E]) lemma new_action_before_read: assumes E: "E \ \" and wf: "P \ (E, ws) \" and ra: "ra \ read_actions E" and adal: "adal \ action_loc P E ra" and new: "wa \ new_actions_for P E adal" and sc: "\a. \ a < ra; a \ read_actions E \ \ P,E \ a \mrw ws a" shows "wa < ra" using \_new_same_addr_singleton[OF E, of adal] init_before_read[OF E wf ra adal sc] new by auto lemma mrw_before: assumes E: "E \ \" and wf: "P \ (E, ws) \" and mrw: "P,E \ r \mrw w" and sc: "\a. \ a < r; a \ read_actions E \ \ P,E \ a \mrw ws a" shows "w < r" using mrw read_actions_not_write_actions[of r E] apply cases apply(erule action_orderE) apply(erule (1) new_action_before_read[OF E wf]) apply(simp add: new_actions_for_def) apply(erule (1) sc) apply(cases "w = r") apply auto done lemma mrw_change_prefix: assumes E': "E' \ \" and mrw: "P,E \ r \mrw w" and tsa_ok: "thread_start_actions_ok E'" and prefix: "ltake n E [\] ltake n E'" and an: "enat r < n" and a'n: "enat w < n" shows "P,E' \ r \mrw w" using mrw proof cases fix adal assume r: "r \ read_actions E" and adal_r: "adal \ action_loc P E r" and war: "E \ w \a r" and w: "w \ write_actions E" and adal_w: "adal \ action_loc P E w" and mrw: "\wa'. \wa' \ write_actions E; adal \ action_loc P E wa'\ \ E \ wa' \a w \ E \ r \a wa'" show ?thesis proof(rule most_recent_write_for.intros) from r prefix an show r': "r \ read_actions E'" by(rule read_actions_change_prefix) from adal_r show "adal \ action_loc P E' r" by(simp add: action_loc_change_prefix[OF prefix[symmetric] an]) from war prefix a'n an show "E' \ w \a r" by(rule action_order_change_prefix) from w prefix a'n show w': "w \ write_actions E'" by(rule write_actions_change_prefix) from adal_w show adal_w': "adal \ action_loc P E' w" by(simp add: action_loc_change_prefix[OF prefix[symmetric] a'n]) fix wa' assume wa': "wa' \ write_actions E'" and adal_wa': "adal \ action_loc P E' wa'" show "E' \ wa' \a w \ E' \ r \a wa'" proof(cases "enat wa' < n") case True note wa'n = this with wa' prefix[symmetric] have "wa' \ write_actions E" by(rule write_actions_change_prefix) moreover from adal_wa' have "adal \ action_loc P E wa'" by(simp add: action_loc_change_prefix[OF prefix wa'n]) ultimately have "E \ wa' \a w \ E \ r \a wa'" by(rule mrw) thus ?thesis proof assume "E \ wa' \a w" hence "E' \ wa' \a w" using prefix wa'n a'n by(rule action_order_change_prefix) thus ?thesis .. next assume "E \ r \a wa'" hence "E' \ r \a wa'" using prefix an wa'n by(rule action_order_change_prefix) thus ?thesis .. qed next case False note wa'n = this show ?thesis proof(cases "is_new_action (action_obs E' wa')") case False hence "E' \ r \a wa'" using wa'n r' wa' an by(auto intro!: action_orderI) (metis enat_ord_code(1) linorder_le_cases order_le_less_trans) thus ?thesis .. next case True with wa' adal_wa' have new: "wa' \ new_actions_for P E' adal" by(simp add: new_actions_for_def) show ?thesis proof(cases "is_new_action (action_obs E' w)") case True with adal_w' a'n w' have "w \ new_actions_for P E' adal" by(simp add: new_actions_for_def) with E' new have "wa' = w" by(rule \_new_actions_for_fun) thus ?thesis using w' by(auto intro: refl_onPD[OF refl_action_order]) next case False with True wa' w' show ?thesis by(auto intro!: action_orderI) qed qed qed qed qed lemma action_order_read_before_write: assumes E: "E \ \" "P \ (E, ws) \" and ao: "E \ w \a r" and r: "r \ read_actions E" and w: "w \ write_actions E" and adal: "adal \ action_loc P E r" "adal \ action_loc P E w" and sc: "\a. \ a < r; a \ read_actions E \ \ P,E \ a \mrw ws a" shows "w < r" using ao proof(cases rule: action_orderE) case 1 from init_before_read[OF E r adal(1) sc] obtain i where "i < r" "i \ new_actions_for P E adal" by blast moreover from \is_new_action (action_obs E w)\ adal(2) \w \ actions E\ have "w \ new_actions_for P E adal" by(simp add: new_actions_for_def) ultimately show "w < r" using E by(auto dest: \_new_actions_for_fun) next case 2 with r w show ?thesis by(cases "w = r")(auto dest: read_actions_not_write_actions) qed end end diff --git a/thys/JinjaThreads/MM/JMM_Typesafe.thy b/thys/JinjaThreads/MM/JMM_Typesafe.thy --- a/thys/JinjaThreads/MM/JMM_Typesafe.thy +++ b/thys/JinjaThreads/MM/JMM_Typesafe.thy @@ -1,948 +1,948 @@ (* Title: JinjaThreads/MM/JMM_Typesafe.thy Author: Andreas Lochbihler *) section \Type-safety proof for the Java memory model\ theory JMM_Typesafe imports JMM_Framework begin text \ Create a dynamic list \heap_independent\ of theorems for replacing heap-dependent constants by heap-independent ones. \ ML \ structure Heap_Independent_Rules = Named_Thms ( val name = @{binding heap_independent} val description = "Simplification rules for heap-independent constants" ) \ setup \Heap_Independent_Rules.setup\ locale heap_base' = h: heap_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate "\_. typeof_addr" heap_read heap_write for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" begin definition typeof_h :: "'addr val \ ty option" where "typeof_h = h.typeof_h undefined" lemma typeof_h_conv_typeof_h [heap_independent, iff]: "h.typeof_h h = typeof_h" by(rule ext)(case_tac x, simp_all add: typeof_h_def) lemmas typeof_h_simps [simp] = h.typeof_h.simps [unfolded heap_independent] definition cname_of :: "'addr \ cname" where "cname_of = h.cname_of undefined" lemma cname_of_conv_cname_of [heap_independent, iff]: "h.cname_of h = cname_of" by(simp add: cname_of_def h.cname_of_def[abs_def]) definition addr_loc_type :: "'m prog \ 'addr \ addr_loc \ ty \ bool" where "addr_loc_type P = h.addr_loc_type P undefined" notation addr_loc_type ("_ \ _@_ : _" [50, 50, 50, 50] 51) lemma addr_loc_type_conv_addr_loc_type [heap_independent, iff]: "h.addr_loc_type P h = addr_loc_type P" by(simp add: addr_loc_type_def h.addr_loc_type_def) lemmas addr_loc_type_cases [cases pred: addr_loc_type] = h.addr_loc_type.cases[unfolded heap_independent] lemmas addr_loc_type_intros = h.addr_loc_type.intros[unfolded heap_independent] definition typeof_addr_loc :: "'m prog \ 'addr \ addr_loc \ ty" where "typeof_addr_loc P = h.typeof_addr_loc P undefined" lemma typeof_addr_loc_conv_typeof_addr_loc [heap_independent, iff]: "h.typeof_addr_loc P h = typeof_addr_loc P" by(simp add: typeof_addr_loc_def h.typeof_addr_loc_def[abs_def]) definition conf :: "'a prog \ 'addr val \ ty \ bool" where "conf P \ h.conf P undefined" notation conf ("_ \ _ :\ _" [51,51,51] 50) lemma conf_conv_conf [heap_independent, iff]: "h.conf P h = conf P" by(simp add: conf_def heap_base.conf_def[abs_def]) lemmas defval_conf [simp] = h.defval_conf[unfolded heap_independent] definition lconf :: "'m prog \ (vname \ 'addr val) \ (vname \ ty) \ bool" where "lconf P = h.lconf P undefined" notation lconf ("_ \ _ '(:\') _" [51,51,51] 50) lemma lconf_conv_lconf [heap_independent, iff]: "h.lconf P h = lconf P" by(simp add: lconf_def h.lconf_def[abs_def]) definition confs :: "'m prog \ 'addr val list \ ty list \ bool" where "confs P = h.confs P undefined" notation confs ("_ \ _ [:\] _" [51,51,51] 50) lemma confs_conv_confs [heap_independent, iff]: "h.confs P h = confs P" by(simp add: confs_def) definition tconf :: "'m prog \ 'thread_id \ bool" where "tconf P = h.tconf P undefined" notation tconf ("_ \ _ \t" [51,51] 50) lemma tconf_conv_tconf [heap_independent, iff]: "h.tconf P h = tconf P" by(simp add: tconf_def h.tconf_def[abs_def]) definition vs_conf :: "'m prog \ ('addr \ addr_loc \ 'addr val set) \ bool" where "vs_conf P = h.vs_conf P undefined" lemma vs_conf_conv_vs_conf [heap_independent, iff]: "h.vs_conf P h = vs_conf P" by(simp add: vs_conf_def h.vs_conf_def[abs_def]) lemmas vs_confI = h.vs_confI[unfolded heap_independent] lemmas vs_confD = h.vs_confD[unfolded heap_independent] text \ use non-speculativity to express that only type-correct values are read \ primrec vs_type_all :: "'m prog \ 'addr \ addr_loc \ 'addr val set" where "vs_type_all P (ad, al) = {v. \T. P \ ad@al : T \ P \ v :\ T}" lemma vs_conf_vs_type_all [simp]: "vs_conf P (vs_type_all P)" by(rule h.vs_confI[unfolded heap_independent])(simp) lemma w_addrs_vs_type_all: "w_addrs (vs_type_all P) \ dom typeof_addr" by(auto simp add: w_addrs_def h.conf_def[unfolded heap_independent]) lemma w_addrs_vs_type_all_in_vs_type_all: "(\ad \ w_addrs (vs_type_all P). {(ad, al)|al. \T. P \ ad@al : T}) \ {adal. vs_type_all P adal \ {}}" by(auto simp add: w_addrs_def vs_type_all_def intro: defval_conf) declare vs_type_all.simps [simp del] lemmas vs_conf_insert_iff = h.vs_conf_insert_iff[unfolded heap_independent] end locale heap' = h: heap addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate "\_. typeof_addr" heap_read heap_write P for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and P :: "'m prog" sublocale heap' < heap_base' . context heap' begin lemma vs_conf_w_value_WriteMemD: "\ vs_conf P (w_value P vs ob); ob = NormalAction (WriteMem ad al v) \ \ \T. P \ ad@al : T \ P \ v :\ T" by(auto elim: vs_confD) lemma vs_conf_w_values_WriteMemD: "\ vs_conf P (w_values P vs obs); NormalAction (WriteMem ad al v) \ set obs \ \ \T. P \ ad@al : T \ P \ v :\ T" apply(induct obs arbitrary: vs) apply(auto 4 3 elim: vs_confD intro: w_values_mono[THEN subsetD]) done lemma w_values_vs_type_all_start_heap_obs: assumes wf: "wf_syscls P" shows "w_values P (vs_type_all P) (map snd (lift_start_obs h.start_tid h.start_heap_obs)) = vs_type_all P" (is "?lhs = ?rhs") proof(rule antisym, rule le_funI, rule subsetI) fix adal v assume v: "v \ ?lhs adal" obtain ad al where adal: "adal = (ad, al)" by(cases adal) show "v \ ?rhs adal" proof(rule ccontr) assume v': "\ ?thesis" from in_w_valuesD[OF v[unfolded adal] this[unfolded adal]] obtain obs' wa obs'' where eq: "map snd (lift_start_obs h.start_tid h.start_heap_obs) = obs' @ wa # obs''" and "write": "is_write_action wa" and loc: "(ad, al) \ action_loc_aux P wa" and vwa: "value_written_aux P wa al = v" by blast+ from "write" show False proof cases case (WriteMem ad' al' v') with vwa loc eq have "WriteMem ad al v \ set h.start_heap_obs" by(auto simp add: map_eq_append_conv Cons_eq_append_conv lift_start_obs_def) from h.start_heap_write_typeable[OF this] v' adal show ?thesis by(auto simp add: vs_type_all_def) next case (NewHeapElem ad' hT) with vwa loc eq have "NewHeapElem ad hT \ set h.start_heap_obs" by(auto simp add: map_eq_append_conv Cons_eq_append_conv lift_start_obs_def) hence "typeof_addr ad = \hT\" by(rule h.NewHeapElem_start_heap_obsD[OF wf]) with v' adal loc vwa NewHeapElem show ?thesis by(auto simp add: vs_type_all_def intro: addr_loc_type_intros h.addr_loc_default_conf[unfolded heap_independent]) qed qed qed(rule w_values_greater) end lemma lprefix_lappend2I: "lprefix xs ys \ lprefix xs (lappend ys zs)" by(auto simp add: lappend_assoc lprefix_conv_lappend) locale known_addrs_typing' = h: known_addrs_typing addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate "\_. typeof_addr" heap_read heap_write allocated known_addrs final r wfx P for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and allocated :: "'heap \ 'addr set" and known_addrs :: "'thread_id \ 'x \ 'addr set" and final :: "'x \ bool" and r :: "('addr, 'thread_id, 'x, 'heap, 'addr, ('addr, 'thread_id) obs_event) semantics" ("_ \ _ -_\ _" [50,0,0,50] 80) and wfx :: "'thread_id \ 'x \ 'heap \ bool" and P :: "'md prog" + assumes NewHeapElem_typed: \ \Should this be moved to known\_addrs\_typing?\ "\ t \ (x, h) -ta\ (x', h'); NewHeapElem ad CTn \ set \ta\\<^bsub>o\<^esub>; typeof_addr ad \ None \ \ typeof_addr ad = \CTn\" sublocale known_addrs_typing' < heap' by unfold_locales context known_addrs_typing' begin lemma known_addrs_typeable_in_vs_type_all: "h.if.known_addrs_state s \ dom typeof_addr \ (\a \ h.if.known_addrs_state s. {(a, al)|al. \T. P \ a@al : T}) \ {adal. vs_type_all P adal \ {}}" by(auto 4 4 dest: subsetD simp add: vs_type_all.simps intro: defval_conf) lemma if_NewHeapElem_typed: "\ t \ xh -ta\i x'h'; NormalAction (NewHeapElem ad CTn) \ set \ta\\<^bsub>o\<^esub>; typeof_addr ad \ None \ \ typeof_addr ad = \CTn\" by(cases rule: h.mthr.init_fin.cases)(auto dest: NewHeapElem_typed) lemma if_redT_NewHeapElem_typed: "\ h.mthr.if.redT s (t, ta) s'; NormalAction (NewHeapElem ad CTn) \ set \ta\\<^bsub>o\<^esub>; typeof_addr ad \ None \ \ typeof_addr ad = \CTn\" by(cases rule: h.mthr.if.redT.cases)(auto dest: if_NewHeapElem_typed) lemma non_speculative_written_value_typeable: assumes wfx_start: "ts_ok wfx (thr (h.start_state f P C M vs)) h.start_heap" and wfP: "wf_syscls P" and E: "E \ h.\_start f P C M vs status" and "write": "w \ write_actions E" and adal: "(ad, al) \ action_loc P E w" and ns: "non_speculative P (vs_type_all P) (lmap snd (ltake (enat w) E))" shows "\T. P \ ad@al : T \ P \ value_written P E w (ad, al) :\ T" proof - let ?start_state = "init_fin_lift_state status (h.start_state f P C M vs)" and ?start_obs = "lift_start_obs h.start_tid h.start_heap_obs" and ?v = "value_written P E w (ad, al)" from "write" have iwa: "is_write_action (action_obs E w)" by cases from E obtain E' where E': "E = lappend (llist_of ?start_obs) E'" and \: "E' \ h.mthr.if.\ ?start_state" by blast from \ obtain E'' where E'': "E' = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E'')" and Runs: "h.mthr.if.mthr.Runs ?start_state E''" by-(rule h.mthr.if.\.cases[OF \]) have wfx': "ts_ok (init_fin_lift wfx) (thr ?start_state) (shr ?start_state)" using wfx_start by(simp add: h.shr_start_state) from ns E' have ns: "non_speculative P (vs_type_all P) (lmap snd (ldropn (length (lift_start_obs h.start_tid h.start_heap_obs)) (ltake (enat w) E)))" by(subst (asm) lappend_ltake_ldrop[where n="enat (length (lift_start_obs h.start_tid h.start_heap_obs))", symmetric])(simp add: non_speculative_lappend min_def ltake_lappend1 w_values_vs_type_all_start_heap_obs[OF wfP] ldrop_enat split: if_split_asm) show ?thesis proof(cases "w < length ?start_obs") case True hence in_start: "action_obs E w \ set (map snd ?start_obs)" unfolding in_set_conv_nth E' by(simp add: lnth_lappend action_obs_def map_nth exI[where x="w"]) from iwa show ?thesis proof(cases) case (WriteMem ad' al' v') with adal have "ad' = ad" "al' = al" "?v = v'" by(simp_all add: value_written.simps) with WriteMem in_start have "WriteMem ad al ?v \ set h.start_heap_obs" by auto thus ?thesis by(rule h.start_heap_write_typeable[unfolded heap_independent]) next case (NewHeapElem ad' CTn) with adal have [simp]: "ad' = ad" by auto with NewHeapElem in_start have "NewHeapElem ad CTn \ set h.start_heap_obs" by auto with wfP have "typeof_addr ad = \CTn\" by(rule h.NewHeapElem_start_heap_obsD) with adal NewHeapElem show ?thesis by(cases al)(auto simp add: value_written.simps intro: addr_loc_type_intros h.addr_loc_default_conf[unfolded heap_independent]) qed next case False define w' where "w' = w - length ?start_obs" with "write" False have w'_len: "enat w' < llength E'" by(cases "llength E'")(auto simp add: actions_def E' elim: write_actions.cases) with Runs obtain m_w n_w t_w ta_w where E'_w: "lnth E' w' = (t_w, \ta_w\\<^bsub>o\<^esub> ! n_w)" and n_w: "n_w < length \ta_w\\<^bsub>o\<^esub>" and m_w: "enat m_w < llength E''" and w_sum: "w' = (\isnd (lnth E'' i)\\<^bsub>o\<^esub>) + n_w" and E''_m_w: "lnth E'' m_w = (t_w, ta_w)" unfolding E'' by(rule h.mthr.if.actions_\E_aux) from E'_w have obs_w: "action_obs E w = \ta_w\\<^bsub>o\<^esub> ! n_w" using False E' w'_def by(simp add: action_obs_def lnth_lappend) let ?E'' = "ldropn (Suc m_w) E''" let ?m_E'' = "ltake (enat m_w) E''" have E'_unfold: "E'' = lappend ?m_E'' (LCons (lnth E'' m_w) ?E'')" unfolding ldropn_Suc_conv_ldropn[OF m_w] by simp hence "h.mthr.if.mthr.Runs ?start_state (lappend ?m_E'' (LCons (lnth E'' m_w) ?E''))" using Runs by simp then obtain \' where \_\': "h.mthr.if.mthr.Trsys ?start_state (list_of ?m_E'') \'" and Runs': "h.mthr.if.mthr.Runs \' (LCons (lnth E'' m_w) ?E'')" by(rule h.mthr.if.mthr.Runs_lappendE) simp from Runs' obtain \''' where red_w: "h.mthr.if.redT \' (t_w, ta_w) \'''" and Runs'': "h.mthr.if.mthr.Runs \''' ?E''" unfolding E''_m_w by cases let ?EE'' = "lmap snd (lappend (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?m_E'')) (llist_of (map (Pair t_w) (take (n_w + 1) \ta_w\\<^bsub>o\<^esub>))))" have len_EE'': "llength ?EE'' = enat (w' + 1)" using n_w m_w apply(simp add: w_sum) apply(subst llength_lconcat_lfinite_conv_sum) apply(simp_all add: split_beta plus_enat_simps(1)[symmetric] add_Suc_right[symmetric] del: plus_enat_simps(1) add_Suc_right) - apply(subst sum_hom[symmetric, where f=enat]) + apply(subst sum_comp_morphism[symmetric, where h=enat]) apply(simp_all add: zero_enat_def min_def le_Suc_eq) apply(rule sum.cong) apply(auto simp add: lnth_ltake less_trans[where y="enat m_w"]) done have prefix: "lprefix ?EE'' (lmap snd E')" unfolding E'' by(subst (2) E'_unfold)(rule lmap_lprefix, clarsimp simp add: lmap_lappend_distrib E''_m_w lprefix_lappend2I[OF lprefix_llist_ofI[OF exI[where x="map (Pair t_w) (drop (n_w + 1) \ta_w\\<^bsub>o\<^esub>)"]]] map_append[symmetric]) from iwa False have iwa': "is_write_action (action_obs E' w')" by(simp add: E' action_obs_def lnth_lappend w'_def) from ns False have "non_speculative P (vs_type_all P) (lmap snd (ltake (enat w') E'))" by(simp add: E' ltake_lappend lmap_lappend_distrib non_speculative_lappend ldropn_lappend2 w'_def) with iwa' have "non_speculative P (vs_type_all P) (lappend (lmap snd (ltake (enat w') E')) (LCons (action_obs E' w') LNil))" by cases(simp_all add: non_speculative_lappend) also have "lappend (lmap snd (ltake (enat w') E')) (LCons (action_obs E' w') LNil) = lmap snd (ltake (enat (w' + 1)) E')" using w'_len by(simp add: ltake_Suc_conv_snoc_lnth lmap_lappend_distrib action_obs_def) also { have "lprefix (lmap snd (ltake (enat (w' + 1)) E')) (lmap snd E')" by(rule lmap_lprefix) simp with prefix have "lprefix ?EE'' (lmap snd (ltake (enat (w' + 1)) E')) \ lprefix (lmap snd (ltake (enat (w' + 1)) E')) ?EE''" by(rule lprefix_down_linear) moreover have "llength (lmap snd (ltake (enat (w' + 1)) E')) = enat (w' + 1)" using w'_len by(cases "llength E'") simp_all ultimately have "lmap snd (ltake (enat (w' + 1)) E') = ?EE''" using len_EE'' by(auto dest: lprefix_llength_eq_imp_eq) } finally have ns1: "non_speculative P (vs_type_all P) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of ?m_E''))))" and ns2: "non_speculative P (w_values P (vs_type_all P) (map snd (list_of (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?m_E''))))) (llist_of (take (Suc n_w) \ta_w\\<^bsub>o\<^esub>))" by(simp_all add: lmap_lappend_distrib non_speculative_lappend split_beta lconcat_llist_of[symmetric] lmap_lconcat llist.map_comp o_def split_def list_of_lmap[symmetric] del: list_of_lmap) have "vs_conf P (vs_type_all P)" by simp with \_\' wfx' ns1 have wfx': "ts_ok (init_fin_lift wfx) (thr \') (shr \')" and vs_conf: "vs_conf P (w_values P (vs_type_all P) (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of ?m_E''))))" by(rule h.if_RedT_non_speculative_invar[unfolded h.mthr.if.RedT_def heap_independent])+ have "concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of ?m_E'')) = map snd (list_of (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?m_E'')))" by(simp add: split_def lmap_lconcat llist.map_comp o_def list_of_lconcat map_concat) with vs_conf have "vs_conf P (w_values P (vs_type_all P) \)" by simp with red_w wfx' ns2 have vs_conf': "vs_conf P (w_values P (w_values P (vs_type_all P) (map snd (list_of (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?m_E''))))) (take (Suc n_w) \ta_w\\<^bsub>o\<^esub>))" (is "vs_conf _ ?vs'") by(rule h.if_redT_non_speculative_vs_conf[unfolded heap_independent]) from len_EE'' have "enat w' < llength ?EE''" by simp from w'_len have "lnth ?EE'' w' = action_obs E' w'" using lprefix_lnthD[OF prefix \enat w' < llength ?EE''\] by(simp add: action_obs_def) hence "\ \ lset ?EE''" using \enat w' < llength ?EE''\ unfolding lset_conv_lnth by(auto intro!: exI) also have "\ \ set (map snd (list_of (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?m_E''))) @ take (Suc n_w) \ta_w\\<^bsub>o\<^esub>)" by(auto 4 4 intro: rev_image_eqI rev_bexI simp add: split_beta lset_lconcat_lfinite dest: lset_lappend[THEN subsetD]) also have "action_obs E' w' = action_obs E w" using False by(simp add: E' w'_def lnth_lappend action_obs_def) also note obs_w_in_set = calculation and calculation = nothing from iwa have "?v \ w_values P (vs_type_all P) (map snd (list_of (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ?m_E''))) @ take (Suc n_w) \ta_w\\<^bsub>o\<^esub>) (ad, al)" proof(cases) case (WriteMem ad' al' v') with adal have "ad' = ad" "al' = al" "?v = v'" by(simp_all add: value_written.simps) with obs_w_in_set WriteMem show ?thesis by -(rule w_values_WriteMemD, simp) next case (NewHeapElem ad' CTn) with adal have [simp]: "ad' = ad" and v: "?v = addr_loc_default P CTn al" by(auto simp add: value_written.simps) with obs_w_in_set NewHeapElem adal show ?thesis by(unfold v)(rule w_values_new_actionD, simp_all) qed hence "?v \ ?vs' (ad, al)" by simp with vs_conf' show "\T. P \ ad@al : T \ P \ ?v :\ T" by(rule h.vs_confD[unfolded heap_independent]) qed qed lemma hb_read_value_typeable: assumes wfx_start: "ts_ok wfx (thr (h.start_state f P C M vs)) h.start_heap" (is "ts_ok wfx (thr ?start_state) _") and wfP: "wf_syscls P" and E: "E \ h.\_start f P C M vs status" and wf: "P \ (E, ws) \" and races: "\a ad al v. \ enat a < llength E; action_obs E a = NormalAction (ReadMem ad al v); \ P,E \ ws a \hb a \ \ \T. P \ ad@al : T \ P \ v :\ T" and r: "enat a < llength E" and read: "action_obs E a = NormalAction (ReadMem ad al v)" shows "\T. P \ ad@al : T \ P \ v :\ T" using r read proof(induction a arbitrary: ad al v rule: less_induct) case (less a) note r = \enat a < llength E\ and read = \action_obs E a = NormalAction (ReadMem ad al v)\ show ?case proof(cases "P,E \ ws a \hb a") case False with r read show ?thesis by(rule races) next case True note hb = this hence ao: "E \ ws a \a a" by(rule happens_before_into_action_order) from wf have ws: "is_write_seen P E ws" by(rule wf_exec_is_write_seenD) from r have "a \ actions E" by(simp add: actions_def) hence "a \ read_actions E" using read .. from is_write_seenD[OF ws this read] have "write": "ws a \ write_actions E" and adal_w: "(ad, al) \ action_loc P E (ws a)" and written: "value_written P E (ws a) (ad, al) = v" by simp_all from "write" have iwa: "is_write_action (action_obs E (ws a))" by cases let ?start_state = "init_fin_lift_state status (h.start_state f P C M vs)" and ?start_obs = "lift_start_obs h.start_tid h.start_heap_obs" show ?thesis proof(cases "ws a < a") case True let ?EE'' = "lmap snd (ltake (enat (ws a)) E)" have "non_speculative P (vs_type_all P) ?EE''" proof(rule non_speculative_nthI) fix i ad' al' v' assume i: "enat i < llength ?EE''" and nth_i: "lnth ?EE'' i = NormalAction (ReadMem ad' al' v')" from i have "i < ws a" by simp hence i': "i < a" using True by(simp) moreover with r have "enat i < llength E" by(metis enat_ord_code(2) order_less_trans) moreover with nth_i i \i < ws a\ have "action_obs E i = NormalAction (ReadMem ad' al' v')" by(simp add: action_obs_def lnth_ltake ac_simps) ultimately have "\T. P \ ad'@al' : T \ P \ v' :\ T" by(rule less.IH) hence "v' \ vs_type_all P (ad', al')" by(simp add: vs_type_all.simps) thus "v' \ w_values P (vs_type_all P) (list_of (ltake (enat i) ?EE'')) (ad', al')" by(rule w_values_mono[THEN subsetD]) qed with wfx_start wfP E "write" adal_w show ?thesis unfolding written[symmetric] by(rule non_speculative_written_value_typeable) next case False from E obtain E' where E': "E = lappend (llist_of ?start_obs) E'" and \: "E' \ h.mthr.if.\ ?start_state" by blast from \ obtain E'' where E'': "E' = lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E'')" and Runs: "h.mthr.if.mthr.Runs ?start_state E''" by-(rule h.mthr.if.\.cases[OF \]) have wfx': "ts_ok (init_fin_lift wfx) (thr ?start_state) (shr ?start_state)" using wfx_start by(simp add: h.shr_start_state) have a_start: "\ a < length ?start_obs" proof assume "a < length ?start_obs" with read have "NormalAction (ReadMem ad al v) \ snd ` set ?start_obs" unfolding set_map[symmetric] in_set_conv_nth by(auto simp add: E' lnth_lappend action_obs_def) hence "ReadMem ad al v \ set h.start_heap_obs" by auto thus False by(simp add: h.start_heap_obs_not_Read) qed hence ws_a_not_le: "\ ws a < length ?start_obs" using False by simp define w where "w = ws a - length ?start_obs" from "write" ws_a_not_le w_def have "enat w < llength (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E''))" by(cases "llength (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E''))")(auto simp add: actions_def E' E'' elim: write_actions.cases) with Runs obtain m_w n_w t_w ta_w where E'_w: "lnth E' w = (t_w, \ta_w\\<^bsub>o\<^esub> ! n_w)" and n_w: "n_w < length \ta_w\\<^bsub>o\<^esub>" and m_w: "enat m_w < llength E''" and w_sum: "w = (\isnd (lnth E'' i)\\<^bsub>o\<^esub>) + n_w" and E''_m_w: "lnth E'' m_w = (t_w, ta_w)" unfolding E'' by(rule h.mthr.if.actions_\E_aux) from E'_w have obs_w: "action_obs E (ws a) = \ta_w\\<^bsub>o\<^esub> ! n_w" using ws_a_not_le E' w_def by(simp add: action_obs_def lnth_lappend) let ?E'' = "ldropn (Suc m_w) E''" let ?m_E'' = "ltake (enat m_w) E''" have E'_unfold: "E'' = lappend ?m_E'' (LCons (lnth E'' m_w) ?E'')" unfolding ldropn_Suc_conv_ldropn[OF m_w] by simp hence "h.mthr.if.mthr.Runs ?start_state (lappend ?m_E'' (LCons (lnth E'' m_w) ?E''))" using Runs by simp then obtain \' where \_\': "h.mthr.if.mthr.Trsys ?start_state (list_of ?m_E'') \'" and Runs': "h.mthr.if.mthr.Runs \' (LCons (lnth E'' m_w) ?E'')" by(rule h.mthr.if.mthr.Runs_lappendE) simp from Runs' obtain \''' where red_w: "h.mthr.if.redT \' (t_w, ta_w) \'''" and Runs'': "h.mthr.if.mthr.Runs \''' ?E''" unfolding E''_m_w by cases from "write" \a \ read_actions E\ have "ws a \ a" by(auto dest: read_actions_not_write_actions) with False have "ws a > a" by simp with ao have new: "is_new_action (action_obs E (ws a))" by(simp add: action_order_def split: if_split_asm) then obtain CTn where obs_w': "action_obs E (ws a) = NormalAction (NewHeapElem ad CTn)" using adal_w by cases auto define a' where "a' = a - length ?start_obs" with False w_def have "enat a' < llength (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E''))" by(simp add: le_less_trans[OF _ \enat w < llength (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) E''))\]) with Runs obtain m_a n_a t_a ta_a where E'_a: "lnth E' a' = (t_a, \ta_a\\<^bsub>o\<^esub> ! n_a)" and n_a: "n_a < length \ta_a\\<^bsub>o\<^esub>" and m_a: "enat m_a < llength E''" and a_sum: "a' = (\isnd (lnth E'' i)\\<^bsub>o\<^esub>) + n_a" and E''_m_a: "lnth E'' m_a = (t_a, ta_a)" unfolding E'' by(rule h.mthr.if.actions_\E_aux) from a_start E'_a read have obs_a: "\ta_a\\<^bsub>o\<^esub> ! n_a = NormalAction (ReadMem ad al v)" using E' w_def by(simp add: action_obs_def lnth_lappend a'_def) let ?E'' = "ldropn (Suc m_a) E''" let ?m_E'' = "ltake (enat m_a) E''" have E'_unfold: "E'' = lappend ?m_E'' (LCons (lnth E'' m_a) ?E'')" unfolding ldropn_Suc_conv_ldropn[OF m_a] by simp hence "h.mthr.if.mthr.Runs ?start_state (lappend ?m_E'' (LCons (lnth E'' m_a) ?E''))" using Runs by simp then obtain \'' where \_\'': "h.mthr.if.mthr.Trsys ?start_state (list_of ?m_E'') \''" and Runs'': "h.mthr.if.mthr.Runs \'' (LCons (lnth E'' m_a) ?E'')" by(rule h.mthr.if.mthr.Runs_lappendE) simp from Runs'' obtain \''' where red_a: "h.mthr.if.redT \'' (t_a, ta_a) \'''" and Runs'': "h.mthr.if.mthr.Runs \''' ?E''" unfolding E''_m_a by cases let ?EE'' = "llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of ?m_E'')))" from m_a have "enat m_a \ llength E''" by simp hence len_EE'': "llength ?EE'' = enat (a' - n_a)" by(simp add: a_sum length_concat sum_list_sum_nth atLeast0LessThan length_list_of_conv_the_enat min_def split_beta lnth_ltake) have prefix: "lprefix ?EE'' (lmap snd E')" unfolding E'' by(subst (2) E'_unfold)(simp add: lmap_lappend_distrib lmap_lconcat llist.map_comp o_def split_def lconcat_llist_of[symmetric] lmap_llist_of[symmetric] lprefix_lappend2I del: lmap_llist_of) have ns: "non_speculative P (vs_type_all P) ?EE''" proof(rule non_speculative_nthI) fix i ad' al' v' assume i: "enat i < llength ?EE''" and lnth_i: "lnth ?EE'' i = NormalAction (ReadMem ad' al' v')" and "non_speculative P (vs_type_all P) (ltake (enat i) ?EE'')" let ?i = "i + length ?start_obs" from i len_EE'' have "i < a'" by simp hence i': "?i < a" by(simp add: a'_def) moreover hence "enat ?i < llength E" using \enat a < llength E\ by(simp add: less_trans[where y="enat a"]) moreover have "enat i < llength E'" using i by -(rule less_le_trans[OF _ lprefix_llength_le[OF prefix], simplified], simp) from lprefix_lnthD[OF prefix i] lnth_i have "lnth (lmap snd E') i = NormalAction (ReadMem ad' al' v')" by simp hence "action_obs E ?i = NormalAction (ReadMem ad' al' v')" using \enat i < llength E'\ by(simp add: E' action_obs_def lnth_lappend E'') ultimately have "\T. P \ ad'@al' : T \ P \ v' :\ T" by(rule less.IH) hence "v' \ vs_type_all P (ad', al')" by(simp add: vs_type_all.simps) thus "v' \ w_values P (vs_type_all P) (list_of (ltake (enat i) ?EE'')) (ad', al')" by(rule w_values_mono[THEN subsetD]) qed have "vs_conf P (vs_type_all P)" by simp with \_\'' wfx' ns have wfx'': "ts_ok (init_fin_lift wfx) (thr \'') (shr \'')" and vs'': "vs_conf P (w_values P (vs_type_all P) (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (list_of ?m_E''))))" by(rule h.if_RedT_non_speculative_invar[unfolded heap_independent h.mthr.if.RedT_def])+ note red_w moreover from n_w obs_w obs_w' have "NormalAction (NewHeapElem ad CTn) \ set \ta_w\\<^bsub>o\<^esub>" unfolding in_set_conv_nth by auto moreover have ta_a_read: "NormalAction (ReadMem ad al v) \ set \ta_a\\<^bsub>o\<^esub>" using n_a obs_a unfolding in_set_conv_nth by blast from red_a have "\T. P \ ad@al : T" proof(cases) case (redT_normal x x' h') from wfx'' \thr \'' t_a = \(x, no_wait_locks)\\ have "init_fin_lift wfx t_a x (shr \'')" by(rule ts_okD) with \t_a \ (x, shr \'') -ta_a\i (x', h')\ show ?thesis using ta_a_read by(rule h.init_fin_red_read_typeable[unfolded heap_independent]) next case redT_acquire thus ?thesis using n_a obs_a ta_a_read by auto qed hence "typeof_addr ad \ None" by(auto elim: addr_loc_type_cases) ultimately have "typeof_addr ad = \CTn\" by(rule if_redT_NewHeapElem_typed) with written adal_w obs_w' show ?thesis by(cases al)(auto simp add: value_written.simps intro: addr_loc_type_intros h.addr_loc_default_conf[unfolded heap_independent]) qed qed qed theorem assumes wfx_start: "ts_ok wfx (thr (h.start_state f P C M vs)) h.start_heap" and wfP: "wf_syscls P" and justified: "P \ (E, ws) weakly_justified_by J" and J: "range (justifying_exec \ J) \ h.\_start f P C M vs status" shows read_value_typeable_justifying: "\ 0 < n; enat a < llength (justifying_exec (J n)); action_obs (justifying_exec (J n)) a = NormalAction (ReadMem ad al v) \ \ \T. P \ ad@al : T \ P \ v :\ T" and read_value_typeable_justifed: "\ E \ h.\_start f P C M vs status; P \ (E, ws) \; enat a < llength E; action_obs E a = NormalAction (ReadMem ad al v) \ \ \T. P \ ad@al : T \ P \ v :\ T" proof - let ?E = "\n. justifying_exec (J n)" and ?\ = "\n. action_translation (J n)" and ?C = "\n. committed (J n)" and ?ws = "\n. justifying_ws (J n)" let ?\ = "h.\_start f P C M vs status" and ?start_obs = "lift_start_obs h.start_tid h.start_heap_obs" { fix a n assume "enat a < llength (justifying_exec (J n))" and "action_obs (justifying_exec (J n)) a = NormalAction (ReadMem ad al v)" and "n > 0" thus "\T. P \ ad@al : T \ P \ v :\ T" proof(induction n arbitrary: a ad al v) case 0 thus ?case by simp next case (Suc n') define n where "n = Suc n'" with Suc have n: "0 < n" and a: "enat a < llength (?E n)" and a_obs: "action_obs (?E n) a = NormalAction (ReadMem ad al v)" by simp_all have wf_n: "P \ (?E n, ?ws n) \" using justified by(simp add: justification_well_formed_def) from J have E: "?E n \ ?\" and E': "?E n' \ ?\" by auto from a a_obs wfx_start wfP E wf_n show ?case proof(rule hb_read_value_typeable[rotated -2]) fix a' ad' al' v' assume a': "enat a' < llength (?E n)" and a'_obs: "action_obs (?E n) a' = NormalAction (ReadMem ad' al' v')" and nhb: "\ P,?E n \ ?ws n a' \hb a'" from a' have "a' \ actions (?E n)" by(simp add: actions_def) hence read_a': "a' \ read_actions (?E n)" using a'_obs .. with justified nhb have committed': "?\ n a' \ ?\ n' ` ?C n'" unfolding is_weakly_justified_by.simps n_def uncommitted_reads_see_hb_def by blast from justified have wfa_n: "wf_action_translation E (J n)" and wfa_n': "wf_action_translation E (J n')" by(simp_all add: wf_action_translations_def) hence inj_n: "inj_on (?\ n) (actions (?E n))" and inj_n': "inj_on (?\ n') (actions (?E n'))" by(blast dest: wf_action_translation_on_inj_onD)+ from justified have C_n: "?C n \ actions (?E n)" and C_n': "?C n' \ actions (?E n')" and wf_n': "P \ (?E n', ?ws n') \" by(simp_all add: committed_subset_actions_def justification_well_formed_def) from justified have "?\ n' ` ?C n' \ ?\ n ` ?C n" unfolding n_def by(simp add: is_commit_sequence_def) with n_def committed' have "?\ n a' \ ?\ n ` ?C n" by auto with inj_n C_n have committed: "a' \ ?C n" using \a' \ actions (?E n)\ by(auto dest: inj_onD) with justified read_a' have ws_committed: "ws (?\ n a') \ ?\ n ` ?C n" by(rule weakly_justified_write_seen_hb_read_committed) from wf_n have ws_n: "is_write_seen P (?E n) (?ws n)" by(rule wf_exec_is_write_seenD) from is_write_seenD[OF this read_a' a'_obs] have ws_write: "?ws n a' \ write_actions (?E n)" and adal: "(ad', al') \ action_loc P (?E n) (?ws n a')" and written: "value_written P (?E n) (?ws n a') (ad', al') = v'" by simp_all define a'' where "a'' = inv_into (actions (?E n')) (?\ n') (?\ n a')" from C_n' n committed' have "?\ n a' \ ?\ n' ` actions (?E n')" by auto hence a'': "?\ n' a'' = ?\ n a'" and a''_action: "a'' \ actions (?E n')" using inj_n' committed' n by(simp_all add: a''_def f_inv_into_f inv_into_into) hence committed'': "a'' \ ?C n'" using committed' n inj_n' C_n' by(fastforce dest: inj_onD) from committed committed'' wfa_n wfa_n' a'' have "action_obs (?E n') a'' \ action_obs (?E n) a'" by(auto dest!: wf_action_translation_on_actionD intro: sim_action_trans sim_action_sym) with a'_obs committed'' C_n' have read_a'': "a'' \ read_actions (?E n')" by(auto intro: read_actions.intros) then obtain ad'' al'' v'' where a''_obs: "action_obs (?E n') a'' = NormalAction (ReadMem ad'' al'' v'')" by cases from committed'' have "n' > 0" using justified by(cases n')(simp_all add: is_commit_sequence_def) then obtain n'' where n'': "n' = Suc n''" by(cases n') simp_all from justified have wfa_n'': "wf_action_translation E (J n'')" by(simp add: wf_action_translations_def) hence inj_n'': "inj_on (?\ n'') (actions (?E n''))" by(blast dest: wf_action_translation_on_inj_onD)+ from justified have C_n'': "?C n'' \ actions (?E n'')" by(simp add: committed_subset_actions_def) from justified committed' committed'' n_def read_a' read_a'' n have "?\ n (?ws n (inv_into (actions (?E n)) (?\ n) (?\ n' a''))) = ws (?\ n' a'')" by(simp add: write_seen_committed_def) hence "?\ n (?ws n a') = ws (?\ n a')" using inj_n \a' \ actions (?E n)\ by(simp add: a'') from ws_committed obtain w where w: "ws (?\ n a') = ?\ n w" and committed_w: "w \ ?C n" by blast from committed_w C_n have "w \ actions (?E n)" by blast hence w_def: "w = ?ws n a'" using \?\ n (?ws n a') = ws (?\ n a')\ inj_n ws_write unfolding w by(auto dest: inj_onD) have committed_ws: "?ws n a' \ ?C n" using committed_w by(simp add: w_def) with wfa_n have sim_ws: "action_obs (?E n) (?ws n a') \ action_obs E (?\ n (?ws n a'))" by(blast dest: wf_action_translation_on_actionD) from wfa_n committed_ws have sim_ws: "action_obs (?E n) (?ws n a') \ action_obs E (?\ n (?ws n a'))" by(blast dest: wf_action_translation_on_actionD) with adal have adal_E: "(ad', al') \ action_loc P E (?\ n (?ws n a'))" by(simp add: action_loc_aux_sim_action) have "\w \ write_actions (?E n'). (ad', al') \ action_loc P (?E n') w \ value_written P (?E n') w (ad', al') = v'" proof(cases "?\ n' a'' \ ?\ n'' ` ?C n''") case True then obtain a''' where a''': "?\ n'' a''' = ?\ n' a''" and committed''': "a''' \ ?C n''" by auto from committed''' C_n'' have a'''_action: "a''' \ actions (?E n'')" by auto from committed'' committed''' wfa_n' wfa_n'' a''' have "action_obs (?E n'') a''' \ action_obs (?E n') a''" by(auto dest!: wf_action_translation_on_actionD intro: sim_action_trans sim_action_sym) with read_a'' committed''' C_n'' have read_a''': "a''' \ read_actions (?E n'')" by cases(auto intro: read_actions.intros) hence "?\ n' (?ws n' (inv_into (actions (?E n')) (?\ n') (?\ n'' a'''))) = ws (?\ n'' a''')" using justified committed''' unfolding is_weakly_justified_by.simps n'' Let_def write_seen_committed_def by blast also have "inv_into (actions (?E n')) (?\ n') (?\ n'' a''') = a''" using a''' inj_n' a''_action by(simp) also note a''' also note a'' finally have \_n': "?\ n' (?ws n' a'') = ws (?\ n a')" . then have "ws (?\ n a') = ?\ n' (?ws n' a'')" .. with \?\ n (?ws n a') = ws (?\ n a')\[symmetric] have eq_ws: "?\ n' (?ws n' a'') = ?\ n (?ws n a')" by simp from wf_n'[THEN wf_exec_is_write_seenD, THEN is_write_seenD, OF read_a'' a''_obs] have ws_write': "?ws n' a'' \ write_actions (?E n')" by simp from justified read_a'' committed'' have "ws (?\ n' a'') \ ?\ n' ` ?C n'" by(rule weakly_justified_write_seen_hb_read_committed) then obtain w' where w': "ws (?\ n' a'') = ?\ n' w'" and committed_w': "w' \ ?C n'" by blast from committed_w' C_n' have "w' \ actions (?E n')" by blast hence w'_def: "w' = ?ws n' a''" using \_n' inj_n' ws_write' unfolding w' a''[symmetric] by(auto dest: inj_onD) with committed_w' have committed_ws'': "?ws n' a'' \ committed (J n')" by simp with committed_ws wfa_n wfa_n' eq_ws have "action_obs (?E n') (?ws n' a'') \ action_obs (?E n) (?ws n a')" by(auto dest!: wf_action_translation_on_actionD intro: sim_action_trans sim_action_sym) hence adal_eq: "action_loc P (?E n') (?ws n' a'') = action_loc P (?E n) (?ws n a')" by(simp add: action_loc_aux_sim_action) with adal have adal': "(ad', al') \ action_loc P (?E n') (?ws n' a'')" by(simp add: action_loc_aux_sim_action) from committed_ws'' have "?ws n' a'' \ actions (?E n')" using C_n' by blast with ws_write \action_obs (?E n') (?ws n' a'') \ action_obs (?E n) (?ws n a')\ have ws_write'': "?ws n' a'' \ write_actions (?E n')" by(cases)(auto intro: write_actions.intros simp add: sim_action_is_write_action_eq) from wfa_n' committed_ws'' have sim_ws': "action_obs (?E n') (?ws n' a'') \ action_obs E (?\ n' (?ws n' a''))" by(blast dest: wf_action_translation_on_actionD) with adal' have adal'_E: "(ad', al') \ action_loc P E (?\ n' (?ws n' a''))" by(simp add: action_loc_aux_sim_action) from justified committed_ws ws_write adal_E have "value_written P (?E n) (?ws n a') (ad', al') = value_written P E (?\ n (?ws n a')) (ad', al')" unfolding is_weakly_justified_by.simps Let_def value_written_committed_def by blast also note eq_ws[symmetric] also from justified committed_ws'' ws_write'' adal'_E have "value_written P E (?\ n' (?ws n' a'')) (ad', al') = value_written P (?E n') (?ws n' a'') (ad', al')" unfolding is_weakly_justified_by.simps Let_def value_written_committed_def by(blast dest: sym) finally show ?thesis using written ws_write'' adal' by auto next case False with justified read_a'' committed'' have "ws (?\ n' a'') \ ?\ n'' ` ?C n''" unfolding is_weakly_justified_by.simps Let_def n'' committed_reads_see_committed_writes_weak_def by blast with a'' obtain w where w: "?\ n'' w = ws (?\ n a')" and committed_w: "w \ ?C n''" by auto from justified have "?\ n'' ` ?C n'' \ ?\ n' ` ?C n'" by(simp add: is_commit_sequence_def n'') with committed_w w[symmetric] have "ws (?\ n a') \ ?\ n' ` ?C n'" by(auto) then obtain w' where w': "ws (?\ n a') = ?\ n' w'" and committed_w': "w' \ ?C n'" by blast from wfa_n' committed_w' have "action_obs (?E n') w' \ action_obs E (?\ n' w')" by(blast dest: wf_action_translation_on_actionD) from this[folded w', folded \?\ n (?ws n a') = ws (?\ n a')\] sim_ws[symmetric] have sim_w': "action_obs (?E n') w' \ action_obs (?E n) (?ws n a')" by(rule sim_action_trans) with ws_write committed_w' C_n' have write_w': "w' \ write_actions (?E n')" by(cases)(auto intro!: write_actions.intros simp add: sim_action_is_write_action_eq) hence "value_written P (?E n') w' (ad', al') = value_written P E (?\ n' w') (ad', al')" using adal_E committed_w' justified unfolding \?\ n (?ws n a') = ws (?\ n a')\ w' is_weakly_justified_by.simps Let_def value_written_committed_def by blast also note w'[symmetric] also note \?\ n (?ws n a') = ws (?\ n a')\[symmetric] also have "value_written P E (?\ n (?ws n a')) (ad', al') = value_written P (?E n) (?ws n a') (ad', al')" using justified committed_ws ws_write adal_E unfolding is_weakly_justified_by.simps Let_def value_written_committed_def by(blast dest: sym) also have "(ad', al') \ action_loc P (?E n') w'" using sim_w' adal by(simp add: action_loc_aux_sim_action) ultimately show ?thesis using written write_w' by auto qed then obtain w where w: "w \ write_actions (?E n')" and adal: "(ad', al') \ action_loc P (?E n') w" and written: "value_written P (?E n') w (ad', al') = v'" by blast from w have w_len: "enat w < llength (?E n')" by(cases)(simp add: actions_def) let ?EE'' = "lmap snd (ltake (enat w) (?E n'))" have "non_speculative P (vs_type_all P) ?EE''" proof(rule non_speculative_nthI) fix i ad al v assume i: "enat i < llength ?EE''" and i_nth: "lnth ?EE'' i = NormalAction (ReadMem ad al v)" and ns: "non_speculative P (vs_type_all P) (ltake (enat i) ?EE'')" from i w_len have "i < w" by(simp add: min_def not_le split: if_split_asm) with w_len have "enat i < llength (?E n')" by(simp add: less_trans[where y="enat w"]) moreover from i_nth i \i < w\ w_len have "action_obs (?E n') i = NormalAction (ReadMem ad al v)" by(simp add: action_obs_def ac_simps less_trans[where y="enat w"] lnth_ltake) moreover from n'' have "0 < n'" by simp ultimately have "\T. P \ ad@al : T \ P \ v :\ T" by(rule Suc.IH) hence "v \ vs_type_all P (ad, al)" by(simp add: vs_type_all.simps) thus "v \ w_values P (vs_type_all P) (list_of (ltake (enat i) ?EE'')) (ad, al)" by(rule w_values_mono[THEN subsetD]) qed with wfx_start wfP E' w adal show "\T. P \ ad'@al' : T \ P \ v' :\ T" unfolding written[symmetric] by(rule non_speculative_written_value_typeable) qed qed } note justifying = this assume a: "enat a < llength E" and read: "action_obs E a = NormalAction (ReadMem ad al v)" and E: "E \ h.\_start f P C M vs status" and wf: "P \ (E, ws) \" from a have action: "a \ actions E" by(auto simp add: actions_def action_obs_def) with justified obtain n a' where a': "a = ?\ n a'" and committed': "a' \ ?C n" by(auto simp add: is_commit_sequence_def) from justified have C_n: "?C n \ actions (?E n)" and C_Sn: "?C (Suc n) \ actions (?E (Suc n))" and wf_tr: "wf_action_translation E (J n)" and wf_tr': "wf_action_translation E (J (Suc n))" by(auto simp add: committed_subset_actions_def wf_action_translations_def) from C_n committed' have action': "a' \ actions (?E n)" by blast from wf_tr committed' a' have "action_tid E a = action_tid (?E n) a'" "action_obs E a \ action_obs (?E n) a'" by(auto simp add: wf_action_translation_on_def intro: sim_action_sym) with read obtain v' where "action_obs (?E n) a' = NormalAction (ReadMem ad al v')" by(clarsimp simp add: action_obs_def) with action' have read': "a' \ read_actions (?E n)" .. from justified have "?\ n ` ?C n \ ?\ (Suc n) ` ?C (Suc n)" by(simp add: is_commit_sequence_def) with committed' a' have "a \ \" by auto then obtain a'' where a'': "a = ?\ (Suc n) a''" and committed'': "a'' \ ?C (Suc n)" by auto from committed'' C_Sn have action'': "a'' \ actions (?E (Suc n))" by blast with wf_tr' have "a'' = inv_into (actions (?E (Suc n))) (?\ (Suc n)) a" by(simp add: a'' wf_action_translation_on_def) with justified read' committed' a' have ws_a: "ws a = ?\ (Suc n) (?ws (Suc n) a'')" by(simp add: write_seen_committed_def) from wf_tr' committed'' a'' have "action_tid E a = action_tid (?E (Suc n)) a''" and "action_obs E a \ action_obs (?E (Suc n)) a''" by(auto simp add: wf_action_translation_on_def intro: sim_action_sym) with read obtain v'' where a_obs'': "action_obs (?E (Suc n)) a'' = NormalAction (ReadMem ad al v'')" by(clarsimp simp add: action_obs_def) with action'' have read'': "a'' \ read_actions (?E (Suc n))" by(auto intro: read_actions.intros simp add: action_obs_def) have "a \ read_actions E" "action_obs E a = NormalAction (ReadMem ad al v)" using action read by(auto intro: read_actions.intros simp add: action_obs_def read) from is_write_seenD[OF wf_exec_is_write_seenD[OF wf] this] have v_eq: "v = value_written P E (ws a) (ad, al)" and adal: "(ad, al) \ action_loc P E (ws a)" by simp_all from justified have "P \ (?E (Suc n), ?ws (Suc n)) \" by(simp add: justification_well_formed_def) from is_write_seenD[OF wf_exec_is_write_seenD[OF this] read'' a_obs''] have write'': "?ws (Suc n) a'' \ write_actions (?E (Suc n))" and written'': "value_written P (?E (Suc n)) (?ws (Suc n) a'') (ad, al) = v''" by simp_all from justified read'' committed'' have "ws (?\ (Suc n) a'') \ ?\ (Suc n) ` ?C (Suc n)" by(rule weakly_justified_write_seen_hb_read_committed) then obtain w where w: "ws (?\ (Suc n) a'') = ?\ (Suc n) w" and committed_w: "w \ ?C (Suc n)" by blast with C_Sn have "w \ actions (?E (Suc n))" by blast moreover have "ws (?\ (Suc n) a'') = ?\ (Suc n) (?ws (Suc n) a'')" using ws_a a'' by simp ultimately have w_def: "w = ?ws (Suc n) a''" using wf_action_translation_on_inj_onD[OF wf_tr'] write'' unfolding w by(auto dest: inj_onD) with committed_w have "?ws (Suc n) a'' \ ?C (Suc n)" by simp hence "value_written P E (ws a) (ad, al) = value_written P (?E (Suc n)) (?ws (Suc n) a'') (ad, al)" using adal justified write'' by(simp add: value_written_committed_def ws_a) with v_eq written'' have "v = v''" by simp from read'' have "enat a'' < llength (?E (Suc n))" by(cases)(simp add: actions_def) thus "\T. P \ ad@al : T \ P \ v :\ T" by(rule justifying)(simp_all add: a_obs'' \v = v''\) qed corollary weakly_legal_read_value_typeable: assumes wfx_start: "ts_ok wfx (thr (h.start_state f P C M vs)) h.start_heap" and wfP: "wf_syscls P" and legal: "weakly_legal_execution P (h.\_start f P C M vs status) (E, ws)" and a: "enat a < llength E" and read: "action_obs E a = NormalAction (ReadMem ad al v)" shows "\T. P \ ad@al : T \ P \ v :\ T" proof - from legal obtain J where "P \ (E, ws) weakly_justified_by J" and "range (justifying_exec \ J) \ h.\_start f P C M vs status" and "E \ h.\_start f P C M vs status" and "P \ (E, ws) \" by(rule legal_executionE) with wfx_start wfP show ?thesis using a read by(rule read_value_typeable_justifed) qed corollary legal_read_value_typeable: "\ ts_ok wfx (thr (h.start_state f P C M vs)) h.start_heap; wf_syscls P; legal_execution P (h.\_start f P C M vs status) (E, ws); enat a < llength E; action_obs E a = NormalAction (ReadMem ad al v) \ \ \T. P \ ad@al : T \ P \ v :\ T" by(erule (1) weakly_legal_read_value_typeable)(rule legal_imp_weakly_legal_execution) end end