diff --git a/thys/Collections/GenCF/Impl/Impl_Array_Hash_Map.thy b/thys/Collections/GenCF/Impl/Impl_Array_Hash_Map.thy --- a/thys/Collections/GenCF/Impl/Impl_Array_Hash_Map.thy +++ b/thys/Collections/GenCF/Impl/Impl_Array_Hash_Map.thy @@ -1,1590 +1,1590 @@ section \\isaheader{Array Based Hash-Maps}\ theory Impl_Array_Hash_Map imports Automatic_Refinement.Automatic_Refinement "../../Iterator/Array_Iterator" "../Gen/Gen_Map" "../Intf/Intf_Hash" "../Intf/Intf_Map" "../../Lib/HashCode" "../../Lib/Code_Target_ICF" "../../Lib/Diff_Array" Impl_List_Map begin subsection \Type definition and primitive operations\ definition load_factor :: nat \ \in percent\ where "load_factor = 75" datatype ('key, 'val) hashmap = HashMap "('key,'val) list_map array" "nat" subsection \Operations\ definition new_hashmap_with :: "nat \ ('k, 'v) hashmap" where "\size. new_hashmap_with size = HashMap (new_array [] size) 0" definition ahm_empty :: "nat \ ('k, 'v) hashmap" where "ahm_empty def_size \ new_hashmap_with def_size" definition bucket_ok :: "'k bhc \ nat \ nat \ ('k\'v) list \ bool" where "bucket_ok bhc len h kvs = (\k \ fst ` set kvs. bhc len k = h)" definition ahm_invar_aux :: "'k bhc \ nat \ ('k\'v) list array \ bool" where "ahm_invar_aux bhc n a \ (\h. h < array_length a \ bucket_ok bhc (array_length a) h (array_get a h) \ list_map_invar (array_get a h)) \ array_foldl (\_ n kvs. n + size kvs) 0 a = n \ array_length a > 1" primrec ahm_invar :: "'k bhc \ ('k, 'v) hashmap \ bool" where "ahm_invar bhc (HashMap a n) = ahm_invar_aux bhc n a" definition ahm_lookup_aux :: "'k eq \ 'k bhc \ 'k \ ('k, 'v) list_map array \ 'v option" where [simp]: "ahm_lookup_aux eq bhc k a = list_map_lookup eq k (array_get a (bhc (array_length a) k))" primrec ahm_lookup where "ahm_lookup eq bhc k (HashMap a _) = ahm_lookup_aux eq bhc k a" definition "ahm_\ bhc m \ \k. ahm_lookup (=) bhc k m" definition ahm_map_rel_def_internal: "ahm_map_rel Rk Rv \ {(HashMap a n, HashMap a' n)| a a' n n'. (a,a') \ \\\Rk,Rv\prod_rel\list_rel\array_rel \ (n,n') \ Id}" lemma ahm_map_rel_def: "\Rk,Rv\ ahm_map_rel \ {(HashMap a n, HashMap a' n)| a a' n n'. (a,a') \ \\\Rk,Rv\prod_rel\list_rel\array_rel \ (n,n') \ Id}" unfolding relAPP_def ahm_map_rel_def_internal . definition ahm_map_rel'_def: "ahm_map_rel' bhc \ br (ahm_\ bhc) (ahm_invar bhc)" definition ahm_rel_def_internal: "ahm_rel bhc Rk Rv = \Rk,Rv\ ahm_map_rel O ahm_map_rel' (abstract_bounded_hashcode Rk bhc)" lemma ahm_rel_def: "\Rk, Rv\ ahm_rel bhc \ \Rk,Rv\ ahm_map_rel O ahm_map_rel' (abstract_bounded_hashcode Rk bhc)" unfolding relAPP_def ahm_rel_def_internal . lemmas [autoref_rel_intf] = REL_INTFI[of "ahm_rel bhc" i_map] for bhc abbreviation "dflt_ahm_rel \ ahm_rel bounded_hashcode_nat" primrec ahm_iteratei_aux :: "(('k\'v) list array) \ ('k\'v, '\) set_iterator" where "ahm_iteratei_aux (Array xs) c f = foldli (concat xs) c f" primrec ahm_iteratei :: "(('k, 'v) hashmap) \ (('k\'v), '\) set_iterator" where "ahm_iteratei (HashMap a n) = ahm_iteratei_aux a" definition ahm_rehash_aux' :: "'k bhc \ nat \ 'k\'v \ ('k\'v) list array \ ('k\'v) list array" where "ahm_rehash_aux' bhc n kv a = (let h = bhc n (fst kv) in array_set a h (kv # array_get a h))" definition ahm_rehash_aux :: "'k bhc \ ('k\'v) list array \ nat \ ('k\'v) list array" where "ahm_rehash_aux bhc a sz = ahm_iteratei_aux a (\x. True) (ahm_rehash_aux' bhc sz) (new_array [] sz)" primrec ahm_rehash :: "'k bhc \ ('k,'v) hashmap \ nat \ ('k,'v) hashmap" where "ahm_rehash bhc (HashMap a n) sz = HashMap (ahm_rehash_aux bhc a sz) n" primrec hm_grow :: "('k,'v) hashmap \ nat" where "hm_grow (HashMap a n) = 2 * array_length a + 3" primrec ahm_filled :: "('k,'v) hashmap \ bool" where "ahm_filled (HashMap a n) = (array_length a * load_factor \ n * 100)" primrec ahm_update_aux :: "'k eq \ 'k bhc \ ('k,'v) hashmap \ 'k \ 'v \ ('k, 'v) hashmap" where "ahm_update_aux eq bhc (HashMap a n) k v = (let h = bhc (array_length a) k; m = array_get a h; insert = list_map_lookup eq k m = None in HashMap (array_set a h (list_map_update eq k v m)) (if insert then n + 1 else n))" definition ahm_update :: "'k eq \ 'k bhc \ 'k \ 'v \ ('k, 'v) hashmap \ ('k, 'v) hashmap" where "ahm_update eq bhc k v hm = (let hm' = ahm_update_aux eq bhc hm k v in (if ahm_filled hm' then ahm_rehash bhc hm' (hm_grow hm') else hm'))" primrec ahm_delete :: "'k eq \ 'k bhc \ 'k \ ('k,'v) hashmap \ ('k,'v) hashmap" where "ahm_delete eq bhc k (HashMap a n) = (let h = bhc (array_length a) k; m = array_get a h; deleted = (list_map_lookup eq k m \ None) in HashMap (array_set a h (list_map_delete eq k m)) (if deleted then n - 1 else n))" primrec ahm_isEmpty :: "('k,'v) hashmap \ bool" where "ahm_isEmpty (HashMap _ n) = (n = 0)" primrec ahm_isSng :: "('k,'v) hashmap \ bool" where "ahm_isSng (HashMap _ n) = (n = 1)" primrec ahm_size :: "('k,'v) hashmap \ nat" where "ahm_size (HashMap _ n) = n" lemma hm_grow_gt_1 [iff]: "Suc 0 < hm_grow hm" by(cases hm)(simp) lemma bucket_ok_Nil [simp]: "bucket_ok bhc len h [] = True" by(simp add: bucket_ok_def) lemma bucket_okD: "\ bucket_ok bhc len h xs; (k, v) \ set xs \ \ bhc len k = h" by(auto simp add: bucket_ok_def) lemma bucket_okI: "(\k. k \ fst ` set kvs \ bhc len k = h) \ bucket_ok bhc len h kvs" by(simp add: bucket_ok_def) subsection \Parametricity\ lemma param_HashMap[param]: "(HashMap, HashMap) \ \\\Rk,Rv\prod_rel\list_rel\array_rel \ nat_rel \ \Rk,Rv\ahm_map_rel" unfolding ahm_map_rel_def by force lemma param_case_hashmap[param]: "(case_hashmap, case_hashmap) \ (\\\Rk,Rv\prod_rel\list_rel\array_rel \ nat_rel \ R) \ \Rk,Rv\ahm_map_rel \ R" unfolding ahm_map_rel_def[abs_def] by (force split: hashmap.split dest: fun_relD) lemma rec_hashmap_is_case[simp]: "rec_hashmap = case_hashmap" by (intro ext, simp split: hashmap.split) subsection \@{term ahm_invar}\ lemma ahm_invar_auxD: assumes "ahm_invar_aux bhc n a" shows "\h. h < array_length a \ bucket_ok bhc (array_length a) h (array_get a h)" and "\h. h < array_length a \ list_map_invar (array_get a h)" and "n = array_foldl (\_ n kvs. n + length kvs) 0 a" and "array_length a > 1" using assms unfolding ahm_invar_aux_def by auto lemma ahm_invar_auxE: assumes "ahm_invar_aux bhc n a" obtains "\h. h < array_length a \ bucket_ok bhc (array_length a) h (array_get a h) \ list_map_invar (array_get a h)" and "n = array_foldl (\_ n kvs. n + length kvs) 0 a" and "array_length a > 1" using assms unfolding ahm_invar_aux_def by blast lemma ahm_invar_auxI: "\ \h. h < array_length a \ bucket_ok bhc (array_length a) h (array_get a h); \h. h < array_length a \ list_map_invar (array_get a h); n = array_foldl (\_ n kvs. n + length kvs) 0 a; array_length a > 1 \ \ ahm_invar_aux bhc n a" unfolding ahm_invar_aux_def by blast lemma ahm_invar_distinct_fst_concatD: assumes inv: "ahm_invar_aux bhc n (Array xs)" shows "distinct (map fst (concat xs))" proof - { fix h assume "h < length xs" with inv have "bucket_ok bhc (length xs) h (xs ! h)" and "list_map_invar (xs ! h)" by(simp_all add: ahm_invar_aux_def) } note no_junk = this show ?thesis unfolding map_concat proof(rule distinct_concat') have "distinct [x\xs . x \ []]" unfolding distinct_conv_nth proof(intro allI ballI impI) fix i j assume "i < length [x\xs . x \ []]" "j < length [x\xs . x \ []]" "i \ j" from filter_nth_ex_nth[OF \i < length [x\xs . x \ []]\] obtain i' where "i' \ i" "i' < length xs" and ith: "[x\xs . x \ []] ! i = xs ! i'" and eqi: "[x\take i' xs . x \ []] = take i [x\xs . x \ []]" by blast from filter_nth_ex_nth[OF \j < length [x\xs . x \ []]\] obtain j' where "j' \ j" "j' < length xs" and jth: "[x\xs . x \ []] ! j = xs ! j'" and eqj: "[x\take j' xs . x \ []] = take j [x\xs . x \ []]" by blast show "[x\xs . x \ []] ! i \ [x\xs . x \ []] ! j" proof assume "[x\xs . x \ []] ! i = [x\xs . x \ []] ! j" hence eq: "xs ! i' = xs ! j'" using ith jth by simp from \i < length [x\xs . x \ []]\ have "[x\xs . x \ []] ! i \ set [x\xs . x \ []]" by(rule nth_mem) with ith have "xs ! i' \ []" by simp then obtain kv where "kv \ set (xs ! i')" by(fastforce simp add: neq_Nil_conv) with no_junk[OF \i' < length xs\] have "bhc (length xs) (fst kv) = i'" by(simp add: bucket_ok_def) moreover from eq \kv \ set (xs ! i')\ have "kv \ set (xs ! j')" by simp with no_junk[OF \j' < length xs\] have "bhc (length xs) (fst kv) = j'" by(simp add: bucket_ok_def) ultimately have [simp]: "i' = j'" by simp from \i < length [x\xs . x \ []]\ have "i = length (take i [x\xs . x \ []])" by simp also from eqi eqj have "take i [x\xs . x \ []] = take j [x\xs . x \ []]" by simp finally show False using \i \ j\ \j < length [x\xs . x \ []]\ by simp qed qed moreover have "inj_on (map fst) {x \ set xs. x \ []}" proof(rule inj_onI) fix x y assume "x \ {x \ set xs. x \ []}" "y \ {x \ set xs. x \ []}" "map fst x = map fst y" hence "x \ set xs" "y \ set xs" "x \ []" "y \ []" by auto from \x \ set xs\ obtain i where "xs ! i = x" "i < length xs" unfolding set_conv_nth by fastforce from \y \ set xs\ obtain j where "xs ! j = y" "j < length xs" unfolding set_conv_nth by fastforce from \x \ []\ obtain k v x' where "x = (k, v) # x'" by(cases x) auto with no_junk[OF \i < length xs\] \xs ! i = x\ have "bhc (length xs) k = i" by(auto simp add: bucket_ok_def) moreover from \map fst x = map fst y\ \x = (k, v) # x'\ obtain v' where "(k, v') \ set y" by fastforce with no_junk[OF \j < length xs\] \xs ! j = y\ have "bhc (length xs) k = j" by(auto simp add: bucket_ok_def) ultimately have "i = j" by simp with \xs ! i = x\ \xs ! j = y\ show "x = y" by simp qed ultimately show "distinct [ys\map (map fst) xs . ys \ []]" by(simp add: filter_map o_def distinct_map) next fix ys have A: "\xs. distinct (map fst xs) = list_map_invar xs" by (simp add: list_map_invar_def) assume "ys \ set (map (map fst) xs)" thus "distinct ys" by(clarsimp simp add: set_conv_nth A) (erule no_junk(2)) next fix ys zs assume "ys \ set (map (map fst) xs)" "zs \ set (map (map fst) xs)" "ys \ zs" then obtain ys' zs' where [simp]: "ys = map fst ys'" "zs = map fst zs'" and "ys' \ set xs" "zs' \ set xs" by auto have "fst ` set ys' \ fst ` set zs' = {}" proof(rule equals0I) fix k assume "k \ fst ` set ys' \ fst ` set zs'" then obtain v v' where "(k, v) \ set ys'" "(k, v') \ set zs'" by(auto) from \ys' \ set xs\ obtain i where "xs ! i = ys'" "i < length xs" unfolding set_conv_nth by fastforce with \(k, v) \ set ys'\ have "bhc (length xs) k = i" by(auto dest: no_junk bucket_okD) moreover from \zs' \ set xs\ obtain j where "xs ! j = zs'" "j < length xs" unfolding set_conv_nth by fastforce with \(k, v') \ set zs'\ have "bhc (length xs) k = j" by(auto dest: no_junk bucket_okD) ultimately have "i = j" by simp with \xs ! i = ys'\ \xs ! j = zs'\ have "ys' = zs'" by simp with \ys \ zs\ show False by simp qed thus "set ys \ set zs = {}" by simp qed qed subsection \@{term "ahm_\"}\ (* TODO: Move this *) lemma list_map_lookup_is_map_of: "list_map_lookup (=) k l = map_of l k" using list_map_autoref_lookup_aux[where eq="(=)" and Rk=Id and Rv=Id] by force definition "ahm_\_aux bhc a \ (\k. ahm_lookup_aux (=) bhc k a)" lemma ahm_\_aux_def2: "ahm_\_aux bhc a = (\k. map_of (array_get a (bhc (array_length a) k)) k)" unfolding ahm_\_aux_def ahm_lookup_aux_def by (simp add: list_map_lookup_is_map_of) lemma ahm_\_def2: "ahm_\ bhc (HashMap a n) = ahm_\_aux bhc a" unfolding ahm_\_def ahm_\_aux_def by simp lemma finite_dom_ahm_\_aux: assumes "is_bounded_hashcode Id (=) bhc" "ahm_invar_aux bhc n a" shows "finite (dom (ahm_\_aux bhc a))" proof - have "dom (ahm_\_aux bhc a) \ (\h \ range (bhc (array_length a) :: 'a \ nat). dom (map_of (array_get a h)))" unfolding ahm_\_aux_def2 by(force simp add: dom_map_of_conv_image_fst dest: map_of_SomeD) moreover have "finite \" proof(rule finite_UN_I) from \ahm_invar_aux bhc n a\ have "array_length a > 1" by(simp add: ahm_invar_aux_def) hence "range (bhc (array_length a) :: 'a \ nat) \ {0.. nat))" by(rule finite_subset) simp qed(rule finite_dom_map_of) ultimately show ?thesis by(rule finite_subset) qed lemma ahm_\_aux_new_array[simp]: assumes bhc: "is_bounded_hashcode Id (=) bhc" "1 < sz" shows "ahm_\_aux bhc (new_array [] sz) k = None" using is_bounded_hashcodeD(3)[OF assms] unfolding ahm_\_aux_def ahm_lookup_aux_def by simp lemma ahm_\_aux_conv_map_of_concat: assumes bhc: "is_bounded_hashcode Id (=) bhc" assumes inv: "ahm_invar_aux bhc n (Array xs)" shows "ahm_\_aux bhc (Array xs) = map_of (concat xs)" proof fix k show "ahm_\_aux bhc (Array xs) k = map_of (concat xs) k" proof(cases "map_of (concat xs) k") case None hence "k \ fst ` set (concat xs)" by(simp add: map_of_eq_None_iff) hence "k \ fst ` set (xs ! bhc (length xs) k)" proof(rule contrapos_nn) assume "k \ fst ` set (xs ! bhc (length xs) k)" then obtain v where "(k, v) \ set (xs ! bhc (length xs) k)" by auto moreover from inv have "bhc (length xs) k < length xs" using bhc by (force simp: ahm_invar_aux_def) ultimately show "k \ fst ` set (concat xs)" by (force intro: rev_image_eqI) qed thus ?thesis unfolding None ahm_\_aux_def2 by (simp add: map_of_eq_None_iff) next case (Some v) hence "(k, v) \ set (concat xs)" by(rule map_of_SomeD) then obtain ys where "ys \ set xs" "(k, v) \ set ys" unfolding set_concat by blast from \ys \ set xs\ obtain i j where "i < length xs" "xs ! i = ys" unfolding set_conv_nth by auto with inv \(k, v) \ set ys\ show ?thesis unfolding Some unfolding ahm_\_aux_def2 by(force dest: bucket_okD simp add: ahm_invar_aux_def list_map_invar_def) qed qed lemma ahm_invar_aux_card_dom_ahm_\_auxD: assumes bhc: "is_bounded_hashcode Id (=) bhc" assumes inv: "ahm_invar_aux bhc n a" shows "card (dom (ahm_\_aux bhc a)) = n" proof(cases a) case [simp]: (Array xs) from inv have "card (dom (ahm_\_aux bhc (Array xs))) = card (dom (map_of (concat xs)))" by(simp add: ahm_\_aux_conv_map_of_concat[OF bhc]) also from inv have "distinct (map fst (concat xs))" by(simp add: ahm_invar_distinct_fst_concatD) hence "card (dom (map_of (concat xs))) = length (concat xs)" by(rule card_dom_map_of) also have "length (concat xs) = foldl (+) 0 (map length xs)" by (simp add: length_concat foldl_conv_fold add.commute fold_plus_sum_list_rev) also from inv have "\ = n" unfolding foldl_map by(simp add: ahm_invar_aux_def array_foldl_foldl) finally show ?thesis by(simp) qed lemma finite_dom_ahm_\: assumes "is_bounded_hashcode Id (=) bhc" "ahm_invar bhc hm" shows "finite (dom (ahm_\ bhc hm))" using assms by (cases hm, force intro: finite_dom_ahm_\_aux simp: ahm_\_def2) subsection \@{term ahm_empty}\ lemma ahm_invar_aux_new_array: assumes "n > 1" shows "ahm_invar_aux bhc 0 (new_array [] n)" proof - have "foldl (\b (k, v). b + length v) 0 (zip [0.. 1 \ ahm_invar bhc (new_hashmap_with n)" by(auto simp add: ahm_invar_def new_hashmap_with_def intro: ahm_invar_aux_new_array) lemma ahm_\_new_hashmap_with: assumes "is_bounded_hashcode Id (=) bhc" and "n > 1" shows "Map.empty = ahm_\ bhc (new_hashmap_with n)" unfolding new_hashmap_with_def ahm_\_def using is_bounded_hashcodeD(3)[OF assms] by force lemma ahm_empty_impl: assumes bhc: "is_bounded_hashcode Id (=) bhc" assumes def_size: "def_size > 1" shows "(ahm_empty def_size, Map.empty) \ ahm_map_rel' bhc" proof- from def_size and ahm_\_new_hashmap_with[OF bhc def_size] and ahm_invar_new_hashmap_with[OF def_size] show ?thesis unfolding ahm_empty_def ahm_map_rel'_def br_def by force qed lemma param_ahm_empty[param]: assumes def_size: "(def_size, def_size') \ nat_rel" shows "(ahm_empty def_size ,ahm_empty def_size') \ \Rk,Rv\ahm_map_rel" unfolding ahm_empty_def[abs_def] new_hashmap_with_def[abs_def] new_array_def[abs_def] using assms by parametricity lemma autoref_ahm_empty[autoref_rules]: fixes Rk :: "('kc\'ka) set" assumes bhc: "SIDE_GEN_ALGO (is_bounded_hashcode Rk eq bhc)" assumes def_size: "SIDE_GEN_ALGO (is_valid_def_hm_size TYPE('kc) def_size)" shows "(ahm_empty def_size, op_map_empty) \ \Rk, Rv\ahm_rel bhc" proof- from bhc have eq': "(eq,(=)) \ Rk \ Rk \ bool_rel" by (simp add: is_bounded_hashcodeD) with bhc have "is_bounded_hashcode Id (=) (abstract_bounded_hashcode Rk bhc)" unfolding autoref_tag_defs by blast thus ?thesis using assms unfolding op_map_empty_def unfolding ahm_rel_def is_valid_def_hm_size_def autoref_tag_defs apply (intro relcompI) apply (rule param_ahm_empty[of def_size def_size], simp) apply (blast intro: ahm_empty_impl) done qed subsection \@{term "ahm_lookup"}\ lemma param_ahm_lookup[param]: assumes bhc: "is_bounded_hashcode Rk eq bhc" defines bhc'_def: "bhc' \ abstract_bounded_hashcode Rk bhc" assumes inv: "ahm_invar bhc' m'" assumes K: "(k,k') \ Rk" assumes M: "(m,m') \ \Rk,Rv\ahm_map_rel" shows "(ahm_lookup eq bhc k m, ahm_lookup (=) bhc' k' m') \ \Rv\option_rel" proof- from bhc have eq': "(eq,(=)) \ Rk \ Rk \ bool_rel" by (simp add: is_bounded_hashcodeD) moreover from abstract_bhc_correct[OF bhc] have bhc': "(bhc,bhc') \ nat_rel \ Rk \ nat_rel" unfolding bhc'_def . moreover from M obtain a a' n n' where [simp]: "m = HashMap a n" and [simp]: "m' = HashMap a' n'" and A: "(a,a') \ \\\Rk,Rv\prod_rel\list_rel\array_rel" and N: "(n,n') \ Id" by (cases m, cases m', unfold ahm_map_rel_def, auto) moreover from inv and array_rel_imp_same_length[OF A] have "array_length a > 1" by (simp add: ahm_invar_aux_def) with abstract_bhc_is_bhc[OF bhc] have "bhc' (array_length a) k' < array_length a" unfolding bhc'_def by blast with bhc'[param_fo, OF _ K] have "bhc (array_length a) k < array_length a" by simp ultimately show ?thesis using K unfolding ahm_lookup_def[abs_def] rec_hashmap_is_case by (simp, parametricity) qed lemma ahm_lookup_impl: assumes bhc: "is_bounded_hashcode Id (=) bhc" shows "(ahm_lookup (=) bhc, op_map_lookup) \ Id \ ahm_map_rel' bhc \ Id" unfolding ahm_map_rel'_def br_def ahm_\_def by force lemma autoref_ahm_lookup[autoref_rules]: assumes bhc[unfolded autoref_tag_defs]: "SIDE_GEN_ALGO (is_bounded_hashcode Rk eq bhc)" shows "(ahm_lookup eq bhc, op_map_lookup) \ Rk \ \Rk,Rv\ahm_rel bhc \ \Rv\option_rel" proof (intro fun_relI) let ?bhc' = "abstract_bounded_hashcode Rk bhc" fix k k' a m' assume K: "(k,k') \ Rk" assume M: "(a,m') \ \Rk,Rv\ahm_rel bhc" from bhc have bhc': "is_bounded_hashcode Id (=) ?bhc'" by blast from M obtain a' where M1: "(a,a') \ \Rk,Rv\ahm_map_rel" and M2: "(a',m') \ ahm_map_rel' ?bhc'" unfolding ahm_rel_def by blast hence inv: "ahm_invar ?bhc' a'" unfolding ahm_map_rel'_def br_def by simp from relcompI[OF param_ahm_lookup[OF bhc inv K M1] ahm_lookup_impl[param_fo, OF bhc' _ M2]] show "(ahm_lookup eq bhc k a, op_map_lookup k' m') \ \Rv\option_rel" by simp qed subsection \@{term "ahm_iteratei"}\ abbreviation "ahm_to_list \ it_to_list ahm_iteratei" lemma param_ahm_iteratei_aux[param]: "(ahm_iteratei_aux,ahm_iteratei_aux) \ \\Ra\list_rel\array_rel \ (Rb \ bool_rel) \ (Ra \ Rb \ Rb) \ Rb \ Rb" unfolding ahm_iteratei_aux_def[abs_def] by parametricity lemma param_ahm_iteratei[param]: "(ahm_iteratei,ahm_iteratei) \ \Rk,Rv\ahm_map_rel \ (Rb \ bool_rel) \ (\Rk,Rv\prod_rel \ Rb \ Rb) \ Rb \ Rb" unfolding ahm_iteratei_def[abs_def] rec_hashmap_is_case by parametricity lemma param_ahm_to_list[param]: "(ahm_to_list,ahm_to_list) \ \Rk, Rv\ahm_map_rel \ \\Rk,Rv\prod_rel\list_rel" unfolding it_to_list_def[abs_def] by parametricity lemma ahm_to_list_distinct[simp,intro]: assumes bhc: "is_bounded_hashcode Id (=) bhc" assumes inv: "ahm_invar bhc m" shows "distinct (ahm_to_list m)" proof- obtain n a where [simp]: "m = HashMap a n" by (cases m) obtain l where [simp]: "a = Array l" by (cases a) from inv show "distinct (ahm_to_list m)" unfolding it_to_list_def by (force intro: distinct_mapI dest: ahm_invar_distinct_fst_concatD) qed lemma set_ahm_to_list: assumes bhc: "is_bounded_hashcode Id (=) bhc" assumes ref: "(m,m') \ ahm_map_rel' bhc" shows "map_to_set m' = set (ahm_to_list m)" proof- obtain n a where [simp]: "m = HashMap a n" by (cases m) obtain l where [simp]: "a = Array l" by (cases a) from ref have \[simp]: "m' = ahm_\ bhc m" and inv: "ahm_invar bhc m" unfolding ahm_map_rel'_def br_def by auto from inv have length: "length l > 1" unfolding ahm_invar_def ahm_invar_aux_def by force from inv have buckets_ok: "\h x. h < length l \ x \ set (l!h) \ bhc (length l) (fst x) = h" "\h. h < length l \ distinct (map fst (l!h))" by (simp_all add: ahm_invar_def ahm_invar_aux_def bucket_ok_def list_map_invar_def) show ?thesis unfolding it_to_list_def \ ahm_\_def ahm_iteratei_def apply (simp add: list_map_lookup_is_map_of) proof (intro equalityI subsetI, goal_cases) case prems: (1 x) let ?m = "\k. map_of (l ! bhc (length l) k) k" obtain k v where [simp]: "x = (k, v)" by (cases x) from prems have "set_to_map (map_to_set ?m) k = Some v" by (simp add: set_to_map_simp inj_on_fst_map_to_set) also note map_to_set_inverse finally have "map_of (l ! bhc (length l) k) k = Some v" . hence "(k,v) \ set (l ! bhc (length l) k)" by (simp add: map_of_SomeD) moreover have "bhc (length l) k < length l" using bhc length .. ultimately show ?case by force next case prems: (2 x) obtain k v where [simp]: "x = (k, v)" by (cases x) from prems obtain h where h_props: "(k,v) \ set (l!h)" "h < length l" by (force simp: set_conv_nth) moreover from h_props and buckets_ok have "bhc (length l) k = h" "distinct (map fst (l!h))" by auto ultimately have "map_of (l ! bhc (length l) k) k = Some v" by (force intro: map_of_is_SomeI) thus ?case by simp qed qed (* TODO: find out what the problem is here *) lemma ahm_iteratei_aux_impl: assumes inv: "ahm_invar_aux bhc n a" and bhc: "is_bounded_hashcode Id (=) bhc" shows "map_iterator (ahm_iteratei_aux a) (ahm_\_aux bhc a)" proof (cases a, rule) fix xs assume [simp]: "a = Array xs" show "ahm_iteratei_aux a = foldli (concat xs)" by (intro ext, simp) from ahm_invar_distinct_fst_concatD and inv show "distinct (map fst (concat xs))" by simp from ahm_\_aux_conv_map_of_concat and assms show "ahm_\_aux bhc a = map_of (concat xs)" by simp qed lemma ahm_iteratei_impl: assumes inv: "ahm_invar bhc m" and bhc: "is_bounded_hashcode Id (=) bhc" shows "map_iterator (ahm_iteratei m) (ahm_\ bhc m)" by (insert assms, cases m, simp add: ahm_\_def2, erule (1) ahm_iteratei_aux_impl) lemma autoref_ahm_is_iterator[autoref_ga_rules]: (*assumes eq: "GEN_OP_tag ((eq,OP (=) ::: (Rk \ Rk \ bool_rel)) \ (Rk \ Rk \ bool_rel))"*) assumes bhc: "GEN_ALGO_tag (is_bounded_hashcode Rk eq bhc)" shows "is_map_to_list Rk Rv (ahm_rel bhc) ahm_to_list" unfolding is_map_to_list_def is_map_to_sorted_list_def proof (intro allI impI) let ?bhc' = "abstract_bounded_hashcode Rk bhc" fix a m' assume M: "(a,m') \ \Rk,Rv\ahm_rel bhc" from bhc have bhc': "is_bounded_hashcode Id (=) ?bhc'" unfolding autoref_tag_defs apply (rule_tac abstract_bhc_is_bhc) by simp_all from M obtain a' where M1: "(a,a') \ \Rk,Rv\ahm_map_rel" and M2: "(a',m') \ ahm_map_rel' ?bhc'" unfolding ahm_rel_def by blast hence inv: "ahm_invar ?bhc' a'" unfolding ahm_map_rel'_def br_def by simp let ?l' = "ahm_to_list a'" from param_ahm_to_list[param_fo, OF M1] have "(ahm_to_list a, ?l') \ \\Rk,Rv\prod_rel\list_rel" . moreover from ahm_to_list_distinct[OF bhc' inv] have "distinct (ahm_to_list a')" . moreover from set_ahm_to_list[OF bhc' M2] have "map_to_set m' = set (ahm_to_list a')" . ultimately show "\l'. (ahm_to_list a, l') \ \\Rk,Rv\prod_rel\list_rel \ RETURN l' \ it_to_sorted_list (key_rel (\_ _. True)) (map_to_set m')" by (force simp: it_to_sorted_list_def key_rel_def[abs_def]) qed lemma ahm_iteratei_aux_code[code]: "ahm_iteratei_aux a c f \ = idx_iteratei array_get array_length a c (\x. foldli x c f) \" proof(cases a) case [simp]: (Array xs) have "ahm_iteratei_aux a c f \ = foldli (concat xs) c f \" by simp also have "\ = foldli xs c (\x. foldli x c f) \" by (simp add: foldli_concat) also have "\ = idx_iteratei (!) length xs c (\x. foldli x c f) \" by (simp add: idx_iteratei_nth_length_conv_foldli) also have "\ = idx_iteratei array_get array_length a c (\x. foldli x c f) \" by(simp add: idx_iteratei_array_get_Array_conv_nth) finally show ?thesis . qed subsection \@{term "ahm_rehash"}\ lemma array_length_ahm_rehash_aux': "array_length (ahm_rehash_aux' bhc n kv a) = array_length a" by(simp add: ahm_rehash_aux'_def Let_def) lemma ahm_rehash_aux'_preserves_ahm_invar_aux: assumes inv: "ahm_invar_aux bhc n a" and bhc: "is_bounded_hashcode Id (=) bhc" and fresh: "k \ fst ` set (array_get a (bhc (array_length a) k))" shows "ahm_invar_aux bhc (Suc n) (ahm_rehash_aux' bhc (array_length a) (k, v) a)" (is "ahm_invar_aux bhc _ ?a") proof(rule ahm_invar_auxI) note invD = ahm_invar_auxD[OF inv] let ?l = "array_length a" fix h assume "h < array_length ?a" hence hlen: "h < ?l" by(simp add: array_length_ahm_rehash_aux') from invD(1,2)[OF this] have bucket: "bucket_ok bhc ?l h (array_get a h)" and dist: "distinct (map fst (array_get a h))" by (simp_all add: list_map_invar_def) let ?h = "bhc (array_length a) k" from hlen bucket show "bucket_ok bhc (array_length ?a) h (array_get ?a h)" by(cases "h = ?h")(auto simp add: ahm_rehash_aux'_def Let_def array_length_ahm_rehash_aux' array_get_array_set_other dest: bucket_okD intro!: bucket_okI) from dist hlen fresh show "list_map_invar (array_get ?a h)" unfolding list_map_invar_def by(cases "h = ?h")(auto simp add: ahm_rehash_aux'_def Let_def array_get_array_set_other) next let ?f = "\n kvs. n + length kvs" { fix n :: nat and xs :: "('a \ 'b) list list" have "foldl ?f n xs = n + foldl ?f 0 xs" by(induct xs arbitrary: rule: rev_induct) simp_all } note fold = this let ?h = "bhc (array_length a) k" obtain xs where a [simp]: "a = Array xs" by(cases a) from inv and bhc have [simp]: "bhc (length xs) k < length xs" by (force simp add: ahm_invar_aux_def) have xs: "xs = take ?h xs @ (xs ! ?h) # drop (Suc ?h) xs" by(simp add: Cons_nth_drop_Suc) from inv have "n = array_foldl (\_ n kvs. n + length kvs) 0 a" by(auto elim: ahm_invar_auxE) hence "n = foldl ?f 0 (take ?h xs) + length (xs ! ?h) + foldl ?f 0 (drop (Suc ?h) xs)" by(simp add: array_foldl_foldl)(subst xs, simp, subst (1 2 3 4) fold, simp) thus "Suc n = array_foldl (\_ n kvs. n + length kvs) 0 ?a" by(simp add: ahm_rehash_aux'_def Let_def array_foldl_foldl foldl_list_update)(subst (1 2 3 4) fold, simp) next from inv have "1 < array_length a" by(auto elim: ahm_invar_auxE) thus "1 < array_length ?a" by(simp add: array_length_ahm_rehash_aux') qed (* TODO: Here be dragons *) lemma ahm_rehash_aux_correct: fixes a :: "('k\'v) list array" assumes bhc: "is_bounded_hashcode Id (=) bhc" and inv: "ahm_invar_aux bhc n a" and "sz > 1" shows "ahm_invar_aux bhc n (ahm_rehash_aux bhc a sz)" (is "?thesis1") and "ahm_\_aux bhc (ahm_rehash_aux bhc a sz) = ahm_\_aux bhc a" (is "?thesis2") proof - let ?a = "ahm_rehash_aux bhc a sz" define I where "I it a' \ ahm_invar_aux bhc (n - card it) a' \ array_length a' = sz \ (\k. if k \ it then ahm_\_aux bhc a' k = None else ahm_\_aux bhc a' k = ahm_\_aux bhc a k)" for it a' note iterator_rule = map_iterator_no_cond_rule_P[ OF ahm_iteratei_aux_impl[OF inv bhc], of I "new_array [] sz" "ahm_rehash_aux' bhc sz" "I {}"] from inv have "I {} ?a" unfolding ahm_rehash_aux_def proof(intro iterator_rule) from ahm_invar_aux_card_dom_ahm_\_auxD[OF bhc inv] have "card (dom (ahm_\_aux bhc a)) = n" . moreover from ahm_invar_aux_new_array[OF \1 < sz\] have "ahm_invar_aux bhc 0 (new_array ([]::('k\'v) list) sz)" . moreover { fix k assume "k \ dom (ahm_\_aux bhc a)" hence "ahm_\_aux bhc a k = None" by auto hence "ahm_\_aux bhc (new_array [] sz) k = ahm_\_aux bhc a k" using assms by simp } ultimately show "I (dom (ahm_\_aux bhc a)) (new_array [] sz)" using assms by (simp add: I_def) next fix k :: 'k and v :: 'v and it a' assume "k \ it" "ahm_\_aux bhc a k = Some v" and it_sub: "it \ dom (ahm_\_aux bhc a)" and I: "I it a'" from I have inv': "ahm_invar_aux bhc (n - card it) a'" and a'_eq_a: "\k. k \ it \ ahm_\_aux bhc a' k = ahm_\_aux bhc a k" and a'_None: "\k. k \ it \ ahm_\_aux bhc a' k = None" and [simp]: "sz = array_length a'" by (auto split: if_split_asm simp: I_def) from it_sub finite_dom_ahm_\_aux[OF bhc inv] have "finite it" by(rule finite_subset) moreover with \k \ it\ have "card it > 0" by (auto simp add: card_gt_0_iff) moreover from finite_dom_ahm_\_aux[OF bhc inv] it_sub have "card it \ card (dom (ahm_\_aux bhc a))" by (rule card_mono) moreover have "\ = n" using inv by(simp add: ahm_invar_aux_card_dom_ahm_\_auxD[OF bhc]) ultimately have "n - card (it - {k}) = (n - card it) + 1" using \k \ it\ by auto moreover from \k \ it\ have "ahm_\_aux bhc a' k = None" by (rule a'_None) hence "k \ fst ` set (array_get a' (bhc (array_length a') k))" by (simp add: ahm_\_aux_def2 map_of_eq_None_iff) ultimately have "ahm_invar_aux bhc (n - card (it - {k})) (ahm_rehash_aux' bhc sz (k, v) a')" using ahm_rehash_aux'_preserves_ahm_invar_aux[OF inv' bhc] by simp moreover have "array_length (ahm_rehash_aux' bhc sz (k, v) a') = sz" by (simp add: array_length_ahm_rehash_aux') moreover { fix k' assume "k' \ it - {k}" with is_bounded_hashcodeD(3)[OF bhc \1 < sz\, of k'] a'_None[of k'] have "ahm_\_aux bhc (ahm_rehash_aux' bhc sz (k, v) a') k' = None" unfolding ahm_\_aux_def2 by (cases "bhc sz k = bhc sz k'") (simp_all add: array_get_array_set_other ahm_rehash_aux'_def Let_def) } moreover { fix k' assume "k' \ it - {k}" with is_bounded_hashcodeD(3)[OF bhc \1 < sz\, of k] is_bounded_hashcodeD(3)[OF bhc \1 < sz\, of k'] a'_eq_a[of k'] \k \ it\ have "ahm_\_aux bhc (ahm_rehash_aux' bhc sz (k, v) a') k' = ahm_\_aux bhc a k'" unfolding ahm_rehash_aux'_def Let_def using \ahm_\_aux bhc a k = Some v\ unfolding ahm_\_aux_def2 by(cases "bhc sz k = bhc sz k'") (case_tac [!] "k' = k", simp_all add: array_get_array_set_other) } ultimately show "I (it - {k}) (ahm_rehash_aux' bhc sz (k, v) a')" unfolding I_def by simp qed simp_all thus ?thesis1 ?thesis2 unfolding ahm_rehash_aux_def I_def by auto qed lemma ahm_rehash_correct: fixes hm :: "('k, 'v) hashmap" assumes bhc: "is_bounded_hashcode Id (=) bhc" and inv: "ahm_invar bhc hm" and "sz > 1" shows "ahm_invar bhc (ahm_rehash bhc hm sz)" "ahm_\ bhc (ahm_rehash bhc hm sz) = ahm_\ bhc hm" proof- obtain a n where [simp]: "hm = HashMap a n" by (cases hm) from inv have "ahm_invar_aux bhc n a" by simp from ahm_rehash_aux_correct[OF bhc this \sz > 1\] show "ahm_invar bhc (ahm_rehash bhc hm sz)" and "ahm_\ bhc (ahm_rehash bhc hm sz) = ahm_\ bhc hm" by (simp_all add: ahm_\_def2) qed subsection \@{term ahm_update}\ lemma param_hm_grow[param]: "(hm_grow, hm_grow) \ \Rk,Rv\ahm_map_rel \ nat_rel" unfolding hm_grow_def[abs_def] rec_hashmap_is_case by parametricity lemma param_ahm_rehash_aux'[param]: assumes "is_bounded_hashcode Rk eq bhc" assumes "1 < n" assumes "(bhc,bhc') \ nat_rel \ Rk \ nat_rel" assumes "(n,n') \ nat_rel" and "n = array_length a" assumes "(kv,kv') \ \Rk,Rv\prod_rel" assumes "(a,a') \ \\\Rk,Rv\prod_rel\list_rel\array_rel" shows "(ahm_rehash_aux' bhc n kv a, ahm_rehash_aux' bhc' n' kv' a') \ \\\Rk,Rv\prod_rel\list_rel\array_rel" proof- from assms have "bhc n (fst kv) < array_length a" by force thus ?thesis unfolding ahm_rehash_aux'_def[abs_def] rec_hashmap_is_case Let_def using assms by parametricity qed (* TODO: Move this *) lemma param_new_array[param]: "(new_array, new_array) \ R \ nat_rel \ \R\array_rel" unfolding new_array_def[abs_def] by parametricity (* TODO: move *) lemma param_foldli_induct: assumes l: "(l,l') \ \Ra\list_rel" assumes c: "(c,c') \ Rb \ bool_rel" assumes \: "(\,\') \ Rb" assumes P\: "P \ \'" assumes f: "\a a' b b'. (a,a')\Ra \ (b,b')\Rb \ c b \ c' b' \ P b b' \ (f a b, f' a' b') \ Rb \ P (f a b) (f' a' b')" shows "(foldli l c f \, foldli l' c' f' \') \ Rb" using c \ P\ f by (induction arbitrary: \ \' rule: list_rel_induct[OF l], auto dest!: fun_relD) lemma param_foldli_induct_nocond: assumes l: "(l,l') \ \Ra\list_rel" assumes \: "(\,\') \ Rb" assumes P\: "P \ \'" assumes f: "\a a' b b'. (a,a')\Ra \ (b,b')\Rb \ P b b' \ (f a b, f' a' b') \ Rb \ P (f a b) (f' a' b')" shows "(foldli l (\_. True) f \, foldli l' (\_. True) f' \') \ Rb" using assms by (blast intro: param_foldli_induct) lemma param_ahm_rehash_aux[param]: assumes bhc: "is_bounded_hashcode Rk eq bhc" assumes bhc_rel: "(bhc,bhc') \ nat_rel \ Rk \ nat_rel" assumes A: "(a,a') \ \\\Rk,Rv\prod_rel\list_rel\array_rel" assumes N: "(n,n') \ nat_rel" "1 < n" shows "(ahm_rehash_aux bhc a n, ahm_rehash_aux bhc' a' n') \ \\\Rk,Rv\prod_rel\list_rel\array_rel" proof- obtain l l' where [simp]: "a = Array l" "a' = Array l'" by (cases a, cases a') from A have L: "(l,l') \ \\\Rk,Rv\prod_rel\list_rel\list_rel" unfolding array_rel_def by simp hence L': "(concat l, concat l') \ \\Rk,Rv\prod_rel\list_rel" by parametricity let ?P = "\a a'. n = array_length a" note induct_rule = param_foldli_induct_nocond[OF L', where P="?P"] show ?thesis unfolding ahm_rehash_aux_def by (simp, induction rule: induct_rule, insert N bhc bhc_rel, auto intro: param_new_array[param_fo] param_ahm_rehash_aux'[param_fo] simp: array_length_ahm_rehash_aux') qed (* TODO: Parametricity fails to prove this. Why? *) lemma param_ahm_rehash[param]: assumes bhc: "is_bounded_hashcode Rk eq bhc" assumes bhc_rel: "(bhc,bhc') \ nat_rel \ Rk \ nat_rel" assumes M: "(m,m') \ \Rk,Rv\ahm_map_rel" assumes N: "(n,n') \ nat_rel" "1 < n" shows "(ahm_rehash bhc m n, ahm_rehash bhc' m' n') \ \Rk,Rv\ahm_map_rel" proof- obtain a a' k k' where [simp]: "m = HashMap a k" "m' = HashMap a' k'" by (cases m, cases m') hence K: "(k,k') \ nat_rel" and A: "(a,a') \ \\\Rk,Rv\prod_rel\list_rel\array_rel" using M unfolding ahm_map_rel_def by simp_all show ?thesis unfolding ahm_rehash_def by (simp, insert K A assms, parametricity) qed lemma param_load_factor[param]: "(load_factor, load_factor) \ nat_rel" unfolding load_factor_def by simp lemma param_ahm_filled[param]: "(ahm_filled, ahm_filled) \ \Rk,Rv\ahm_map_rel \ bool_rel" unfolding ahm_filled_def[abs_def] rec_hashmap_is_case by parametricity lemma param_ahm_update_aux[param]: assumes bhc: "is_bounded_hashcode Rk eq bhc" assumes bhc_rel: "(bhc,bhc') \ nat_rel \ Rk \ nat_rel" assumes inv: "ahm_invar bhc' m'" assumes K: "(k,k') \ Rk" assumes V: "(v,v') \ Rv" assumes M: "(m,m') \ \Rk,Rv\ahm_map_rel" shows "(ahm_update_aux eq bhc m k v, ahm_update_aux (=) bhc' m' k' v' ) \ \Rk,Rv\ahm_map_rel" proof- from bhc have eq[param]: "(eq, (=))\Rk \ Rk \ bool_rel" by (simp add: is_bounded_hashcodeD) obtain a a' n n' where [simp]: "m = HashMap a n" and [simp]: "m' = HashMap a' n'" by (cases m, cases m') from M have A: "(a,a') \ \\\Rk,Rv\prod_rel\list_rel\array_rel" and N: "(n,n') \ nat_rel" unfolding ahm_map_rel_def by simp_all from inv have "1 < array_length a'" unfolding ahm_invar_def ahm_invar_aux_def by force hence "1 < array_length a" by (simp add: array_rel_imp_same_length[OF A]) with bhc have bhc_range: "bhc (array_length a) k < array_length a" by blast have option_compare: "\a a'. (a,a') \ \Rv\option_rel \ (a = None,a' = None) \ bool_rel" by force have "(array_get a (bhc (array_length a) k), array_get a' (bhc' (array_length a') k')) \ \\Rk,Rv\prod_rel\list_rel" using A K bhc_rel bhc_range by parametricity note cmp = option_compare[OF param_list_map_lookup[param_fo, OF eq K this]] show ?thesis apply simp unfolding ahm_update_aux_def Let_def rec_hashmap_is_case using assms A N bhc_range cmp by parametricity qed lemma param_ahm_update[param]: assumes bhc: "is_bounded_hashcode Rk eq bhc" assumes bhc_rel: "(bhc,bhc') \ nat_rel \ Rk \ nat_rel" assumes inv: "ahm_invar bhc' m'" assumes K: "(k,k') \ Rk" assumes V: "(v,v') \ Rv" assumes M: "(m,m') \ \Rk,Rv\ahm_map_rel" shows "(ahm_update eq bhc k v m, ahm_update (=) bhc' k' v' m') \ \Rk,Rv\ahm_map_rel" proof- have "1 < hm_grow (ahm_update_aux eq bhc m k v)" by simp with assms show ?thesis unfolding ahm_update_def[abs_def] Let_def by parametricity qed (* TODO: Move *) lemma length_list_map_update: "length (list_map_update (=) k v xs) = (if list_map_lookup (=) k xs = None then Suc (length xs) else length xs)" (is "?l_new = _") proof (cases "list_map_lookup (=) k xs", simp_all) case None hence "k \ dom (map_of xs)" by (force simp: list_map_lookup_is_map_of) hence "\a. list_map_update_aux (=) k v xs a = (k,v) # rev xs @ a" by (induction xs, auto) thus "?l_new = Suc (length xs)" unfolding list_map_update_def by simp next case (Some v') hence "(k,v') \ set xs" unfolding list_map_lookup_is_map_of by (rule map_of_SomeD) hence "\a. length (list_map_update_aux (=) k v xs a) = length xs + length a" by (induction xs, auto) thus "?l_new = length xs" unfolding list_map_update_def by simp qed lemma length_list_map_delete: "length (list_map_delete (=) k xs) = (if list_map_lookup (=) k xs = None then length xs else length xs - 1)" (is "?l_new = _") proof (cases "list_map_lookup (=) k xs", simp_all) case None hence "k \ dom (map_of xs)" by (force simp: list_map_lookup_is_map_of) hence "\a. list_map_delete_aux (=) k xs a = rev xs @ a" by (induction xs, auto) thus "?l_new = length xs" unfolding list_map_delete_def by simp next case (Some v') hence "(k,v') \ set xs" unfolding list_map_lookup_is_map_of by (rule map_of_SomeD) hence "\a. k \ fst`set a \ length (list_map_delete_aux (=) k xs a) = length xs + length a - 1" by (induction xs, auto) thus "?l_new = length xs - Suc 0" unfolding list_map_delete_def by simp qed lemma ahm_update_impl: assumes bhc: "is_bounded_hashcode Id (=) bhc" shows "(ahm_update (=) bhc, op_map_update) \ (Id::('k\'k) set) \ (Id::('v\'v) set) \ ahm_map_rel' bhc \ ahm_map_rel' bhc" proof (intro fun_relI, clarsimp) fix k::'k and v::'v and hm::"('k,'v) hashmap" and m::"'k\'v" assume "(hm,m) \ ahm_map_rel' bhc" hence \: "m = ahm_\ bhc hm" and inv: "ahm_invar bhc hm" unfolding ahm_map_rel'_def br_def by simp_all obtain a n where [simp]: "hm = HashMap a n" by (cases hm) have K: "(k,k) \ Id" and V: "(v,v) \ Id" by simp_all from inv have [simp]: "1 < array_length a" unfolding ahm_invar_def ahm_invar_aux_def by simp hence bhc_range[simp]: "\k. bhc (array_length a) k < array_length a" using bhc by blast let ?l = "array_length a" let ?h = "bhc (array_length a) k" let ?a' = "array_set a ?h (list_map_update (=) k v (array_get a ?h))" let ?n' = "if list_map_lookup (=) k (array_get a ?h) = None then n + 1 else n" let ?list = "array_get a (bhc ?l k)" let ?list' = "map_of ?list" have L: "(?list, ?list') \ br map_of list_map_invar" using inv unfolding ahm_invar_def ahm_invar_aux_def br_def by simp hence list: "list_map_invar ?list" by (simp_all add: br_def) let ?list_new = "list_map_update (=) k v ?list" let ?list_new' = "op_map_update k v (map_of (?list))" from list_map_autoref_update2[param_fo, OF K V L] have list_updated: "map_of ?list_new = ?list_new'" "list_map_invar ?list_new" unfolding br_def by simp_all have "ahm_invar bhc (HashMap ?a' ?n')" unfolding ahm_invar.simps proof(rule ahm_invar_auxI) fix h assume "h < array_length ?a'" hence h_in_range: "h < array_length a" by simp with inv have bucket_ok: "bucket_ok bhc ?l h (array_get a h)" by(auto elim: ahm_invar_auxD) thus "bucket_ok bhc (array_length ?a') h (array_get ?a' h)" proof (cases "h = bhc (array_length a) k") case False with bucket_ok show ?thesis by (intro bucket_okI, force simp add: array_get_array_set_other dest: bucket_okD) next case True show ?thesis proof (insert True, simp, intro bucket_okI, goal_cases) case prems: (1 k') show ?case proof (cases "k = k'") case False from prems have "k' \ dom ?list_new'" by (simp only: dom_map_of_conv_image_fst list_updated(1)[symmetric]) hence "k' \ fst`set ?list" using False by (simp add: dom_map_of_conv_image_fst) from imageE[OF this] obtain x where "fst x = k'" and "x \ set ?list" by force then obtain v' where "(k',v') \ set ?list" by (cases x, simp) with bucket_okD[OF bucket_ok] and \h = bhc (array_length a) k\ show ?thesis by simp qed simp qed qed from \h < array_length a\ inv have "list_map_invar (array_get a h)" by(auto dest: ahm_invar_auxD) with \h < array_length a\ show "list_map_invar (array_get ?a' h)" by (cases "h = ?h", simp_all add: list_updated array_get_array_set_other) next obtain xs where a [simp]: "a = Array xs" by(cases a) let ?f = "\n kvs. n + length kvs" { fix n :: nat and xs :: "('k \ 'v) list list" have "foldl ?f n xs = n + foldl ?f 0 xs" by(induct xs arbitrary: rule: rev_induct) simp_all } note fold = this from inv have [simp]: "bhc (length xs) k < length xs" using bhc_range by simp have xs: "xs = take ?h xs @ (xs ! ?h) # drop (Suc ?h) xs" by(simp add: Cons_nth_drop_Suc) from inv have "n = array_foldl (\_ n kvs. n + length kvs) 0 a" by (force dest: ahm_invar_auxD) hence "n = foldl ?f 0 (take ?h xs) + length (xs ! ?h) + foldl ?f 0 (drop (Suc ?h) xs)" apply (simp add: array_foldl_foldl) apply (subst xs) apply simp apply (metis fold) done thus "?n' = array_foldl (\_ n kvs. n + length kvs) 0 ?a'" apply(simp add: ahm_rehash_aux'_def Let_def array_foldl_foldl foldl_list_update map_of_eq_None_iff) apply(subst (1 2 3 4 5 6 7 8) fold) apply(simp add: length_list_map_update) done next from inv have "1 < array_length a" by(auto elim: ahm_invar_auxE) thus "1 < array_length ?a'" by simp next qed moreover have "ahm_\ bhc (ahm_update_aux (=) bhc hm k v) = - ahm_\ bhc hm(k \ v)" + (ahm_\ bhc hm)(k \ v)" proof fix k' - show "ahm_\ bhc (ahm_update_aux (=) bhc hm k v) k' = (ahm_\ bhc hm(k \ v)) k'" + show "ahm_\ bhc (ahm_update_aux (=) bhc hm k v) k' = ((ahm_\ bhc hm)(k \ v)) k'" proof (cases "bhc ?l k = bhc ?l k'") case False thus ?thesis by (force simp add: Let_def ahm_\_def array_get_array_set_other) next case True hence "bhc ?l k' = bhc ?l k" by simp thus ?thesis by (auto simp add: Let_def ahm_\_def list_map_lookup_is_map_of list_updated) qed qed ultimately have ref: "(ahm_update_aux (=) bhc hm k v, m(k \ v)) \ ahm_map_rel' bhc" (is "(?hm',_)\_") unfolding ahm_map_rel'_def br_def using \ by (auto simp: Let_def) show "(ahm_update (=) bhc k v hm, m(k \ v)) \ ahm_map_rel' bhc" proof (cases "ahm_filled ?hm'") case False with ref show ?thesis unfolding ahm_update_def by (simp del: ahm_update_aux.simps) next case True from ref have "(ahm_rehash bhc ?hm' (hm_grow ?hm'), m(k \ v)) \ ahm_map_rel' bhc" unfolding ahm_map_rel'_def br_def by (simp del: ahm_update_aux.simps add: ahm_rehash_correct[OF bhc]) thus ?thesis unfolding ahm_update_def using True by (simp del: ahm_update_aux.simps add: Let_def) qed qed lemma autoref_ahm_update[autoref_rules]: assumes bhc[unfolded autoref_tag_defs]: "SIDE_GEN_ALGO (is_bounded_hashcode Rk eq bhc)" shows "(ahm_update eq bhc, op_map_update) \ Rk \ Rv \ \Rk,Rv\ahm_rel bhc \ \Rk,Rv\ahm_rel bhc" proof (intro fun_relI) let ?bhc' = "abstract_bounded_hashcode Rk bhc" fix k k' v v' a m' assume K: "(k,k') \ Rk" and V: "(v,v') \ Rv" assume M: "(a,m') \ \Rk,Rv\ahm_rel bhc" (*from eq have eq': "(eq,(=)) \ Rk \ Rk \ bool_rel" by simp*) with bhc have bhc': "is_bounded_hashcode Id (=) ?bhc'" by blast from abstract_bhc_correct[OF bhc] have bhc_rel: "(bhc,?bhc') \ nat_rel \ Rk \ nat_rel" . from M obtain a' where M1: "(a,a') \ \Rk,Rv\ahm_map_rel" and M2: "(a',m') \ ahm_map_rel' ?bhc'" unfolding ahm_rel_def by blast hence inv: "ahm_invar ?bhc' a'" unfolding ahm_map_rel'_def br_def by simp from relcompI[OF param_ahm_update[OF bhc bhc_rel inv K V M1] ahm_update_impl[param_fo, OF bhc' _ _ M2]] show "(ahm_update eq bhc k v a, op_map_update k' v' m') \ \Rk,Rv\ahm_rel bhc" unfolding ahm_rel_def by simp qed subsection \@{term "ahm_delete"}\ lemma param_ahm_delete[param]: (*assumes eq: "(eq,(=)) \ Rk \ Rk \ bool_rel"*) assumes isbhc: "is_bounded_hashcode Rk eq bhc" assumes bhc: "(bhc,bhc') \ nat_rel \ Rk \ nat_rel" assumes inv: "ahm_invar bhc' m'" assumes K: "(k,k') \ Rk" assumes M: "(m,m') \ \Rk,Rv\ahm_map_rel" shows "(ahm_delete eq bhc k m, ahm_delete (=) bhc' k' m') \ \Rk,Rv\ahm_map_rel" proof- from isbhc have eq: "(eq,(=))\Rk\Rk\bool_rel" by (simp add: is_bounded_hashcodeD) obtain a a' n n' where [simp]: "m = HashMap a n" and [simp]: "m' = HashMap a' n'" by (cases m, cases m') from M have A: "(a,a') \ \\\Rk,Rv\prod_rel\list_rel\array_rel" and N: "(n,n') \ nat_rel" unfolding ahm_map_rel_def by simp_all from inv have "1 < array_length a'" unfolding ahm_invar_def ahm_invar_aux_def by force hence "1 < array_length a" by (simp add: array_rel_imp_same_length[OF A]) with isbhc have bhc_range: "bhc (array_length a) k < array_length a" by blast have option_compare: "\a a'. (a,a') \ \Rv\option_rel \ (a = None,a' = None) \ bool_rel" by force have "(array_get a (bhc (array_length a) k), array_get a' (bhc' (array_length a') k')) \ \\Rk,Rv\prod_rel\list_rel" using A K bhc bhc_range by parametricity note cmp = option_compare[OF param_list_map_lookup[param_fo, OF eq K this]] show ?thesis unfolding \m = HashMap a n\ \m' = HashMap a' n'\ by (simp only: ahm_delete.simps Let_def, insert eq isbhc bhc K A N bhc_range cmp, parametricity) qed lemma ahm_delete_impl: assumes bhc: "is_bounded_hashcode Id (=) bhc" shows "(ahm_delete (=) bhc, op_map_delete) \ (Id::('k\'k) set) \ ahm_map_rel' bhc \ ahm_map_rel' bhc" proof (intro fun_relI, clarsimp) fix k::'k and hm::"('k,'v) hashmap" and m::"'k\'v" assume "(hm,m) \ ahm_map_rel' bhc" hence \: "m = ahm_\ bhc hm" and inv: "ahm_invar bhc hm" unfolding ahm_map_rel'_def br_def by simp_all obtain a n where [simp]: "hm = HashMap a n" by (cases hm) have K: "(k,k) \ Id" by simp from inv have [simp]: "1 < array_length a" unfolding ahm_invar_def ahm_invar_aux_def by simp hence bhc_range[simp]: "\k. bhc (array_length a) k < array_length a" using bhc by blast let ?l = "array_length a" let ?h = "bhc ?l k" let ?a' = "array_set a ?h (list_map_delete (=) k (array_get a ?h))" let ?n' = "if list_map_lookup (=) k (array_get a ?h) = None then n else n - 1" let ?list = "array_get a ?h" let ?list' = "map_of ?list" let ?list_new = "list_map_delete (=) k ?list" let ?list_new' = "?list' |` (-{k})" from inv have "(?list, ?list') \ br map_of list_map_invar" unfolding br_def ahm_invar_def ahm_invar_aux_def by simp from list_map_autoref_delete2[param_fo, OF K this] have list_updated: "map_of ?list_new = ?list_new'" "list_map_invar ?list_new" by (simp_all add: br_def) have [simp]: "array_length ?a' = ?l" by simp have "ahm_invar_aux bhc ?n' ?a'" proof(rule ahm_invar_auxI) fix h assume "h < array_length ?a'" hence h_in_range[simp]: "h < array_length a" by simp with inv have inv': "bucket_ok bhc ?l h (array_get a h)" "1 < ?l" "list_map_invar (array_get a h)" by (auto elim: ahm_invar_auxE) show "bucket_ok bhc (array_length ?a') h (array_get ?a' h)" proof (cases "h = bhc ?l k") case False thus ?thesis using inv' by (simp add: array_get_array_set_other) next case True thus ?thesis proof (simp, intro bucket_okI, goal_cases) case prems: (1 k') show ?case proof (cases "k = k'") case False from prems have "k' \ dom ?list_new'" by (simp only: dom_map_of_conv_image_fst list_updated(1)[symmetric]) hence "k' \ fst`set ?list" using False by (simp add: dom_map_of_conv_image_fst) from imageE[OF this] obtain x where "fst x = k'" and "x \ set ?list" by force then obtain v' where "(k',v') \ set ?list" by (cases x, simp) with bucket_okD[OF inv'(1)] and \h = bhc (array_length a) k\ show ?thesis by blast qed simp qed qed from inv'(3) \h < array_length a\ show "list_map_invar (array_get ?a' h)" by (cases "h = ?h", simp_all add: list_updated array_get_array_set_other) next obtain xs where a [simp]: "a = Array xs" by(cases a) let ?f = "\n kvs. n + length (kvs::('k\'v) list)" { fix n :: nat and xs :: "('k\'v) list list" have "foldl ?f n xs = n + foldl ?f 0 xs" by(induct xs arbitrary: rule: rev_induct) simp_all } note fold = this from bhc_range have [simp]: "bhc (length xs) k < length xs" by simp moreover have xs: "xs = take ?h xs @ (xs ! ?h) # drop (Suc ?h) xs" by(simp add: Cons_nth_drop_Suc) from inv have "n = array_foldl (\_ n kvs. n + length kvs) 0 a" by(auto elim: ahm_invar_auxE) hence "n = foldl ?f 0 (take ?h xs) + length (xs ! ?h) + foldl ?f 0 (drop (Suc ?h) xs)" by(simp add: array_foldl_foldl)(subst xs, simp, subst (1 2 3 4) fold, simp) moreover have "\v a b. list_map_lookup (=) k (xs ! ?h) = Some v \ a + (length (xs ! ?h) - 1) + b = a + length (xs ! ?h) + b - 1" by (cases "xs ! ?h", simp_all) ultimately show "?n' = array_foldl (\_ n kvs. n + length kvs) 0 ?a'" apply(simp add: array_foldl_foldl foldl_list_update map_of_eq_None_iff) apply(subst (1 2 3 4 5 6 7 8) fold) apply (intro conjI impI) apply(auto simp add: length_list_map_delete) done next from inv show "1 < array_length ?a'" by(auto elim: ahm_invar_auxE) qed hence "ahm_invar bhc (HashMap ?a' ?n')" by simp moreover have "ahm_\_aux bhc ?a' = ahm_\_aux bhc a |` (- {k})" proof fix k' :: 'k show "ahm_\_aux bhc ?a' k' = (ahm_\_aux bhc a |` (- {k})) k'" proof (cases "bhc ?l k' = ?h") case False hence "k \ k'" by force thus ?thesis using False unfolding ahm_\_aux_def by (simp add: array_get_array_set_other list_map_lookup_is_map_of) next case True thus ?thesis unfolding ahm_\_aux_def by (simp add: list_map_lookup_is_map_of list_updated(1) restrict_map_def) qed qed hence "ahm_\ bhc (HashMap ?a' ?n') = op_map_delete k m" unfolding op_map_delete_def by (simp add: ahm_\_def2 \) ultimately have "(HashMap ?a' ?n', op_map_delete k m) \ ahm_map_rel' bhc" unfolding ahm_map_rel'_def br_def by simp thus "(ahm_delete (=) bhc k hm, m |` (-{k})) \ ahm_map_rel' bhc" by (simp only: \hm = HashMap a n\ ahm_delete.simps Let_def op_map_delete_def, force) qed lemma autoref_ahm_delete[autoref_rules]: assumes bhc[unfolded autoref_tag_defs]: "SIDE_GEN_ALGO (is_bounded_hashcode Rk eq bhc)" shows "(ahm_delete eq bhc, op_map_delete) \ Rk \ \Rk,Rv\ahm_rel bhc \ \Rk,Rv\ahm_rel bhc" proof (intro fun_relI) let ?bhc' = "abstract_bounded_hashcode Rk bhc" fix k k' a m' assume K: "(k,k') \ Rk" assume M: "(a,m') \ \Rk,Rv\ahm_rel bhc" (*from bhc have eq': "(eq,(=)) \ Rk \ Rk \ bool_rel" by (simp add: is_bounded_hashcodeD)*) with bhc have bhc': "is_bounded_hashcode Id (=) ?bhc'" by blast from abstract_bhc_correct[OF bhc] have bhc_rel: "(bhc,?bhc') \ nat_rel \ Rk \ nat_rel" . from M obtain a' where M1: "(a,a') \ \Rk,Rv\ahm_map_rel" and M2: "(a',m') \ ahm_map_rel' ?bhc'" unfolding ahm_rel_def by blast hence inv: "ahm_invar ?bhc' a'" unfolding ahm_map_rel'_def br_def by simp from relcompI[OF param_ahm_delete[OF bhc bhc_rel inv K M1] ahm_delete_impl[param_fo, OF bhc' _ M2]] show "(ahm_delete eq bhc k a, op_map_delete k' m') \ \Rk,Rv\ahm_rel bhc" unfolding ahm_rel_def by simp qed subsection \Various simple operations\ lemma param_ahm_isEmpty[param]: "(ahm_isEmpty, ahm_isEmpty) \ \Rk,Rv\ahm_map_rel \ bool_rel" unfolding ahm_isEmpty_def[abs_def] rec_hashmap_is_case by parametricity lemma param_ahm_isSng[param]: "(ahm_isSng, ahm_isSng) \ \Rk,Rv\ahm_map_rel \ bool_rel" unfolding ahm_isSng_def[abs_def] rec_hashmap_is_case by parametricity lemma param_ahm_size[param]: "(ahm_size, ahm_size) \ \Rk,Rv\ahm_map_rel \ nat_rel" unfolding ahm_size_def[abs_def] rec_hashmap_is_case by parametricity lemma ahm_isEmpty_impl: assumes "is_bounded_hashcode Id (=) bhc" shows "(ahm_isEmpty, op_map_isEmpty) \ ahm_map_rel' bhc \ bool_rel" proof (intro fun_relI) fix hm m assume rel: "(hm,m) \ ahm_map_rel' bhc" obtain a n where [simp]: "hm = HashMap a n" by (cases hm) from rel have \: "m = ahm_\_aux bhc a" and inv: "ahm_invar_aux bhc n a" unfolding ahm_map_rel'_def br_def by (simp_all add: ahm_\_def2) from ahm_invar_aux_card_dom_ahm_\_auxD[OF assms inv,symmetric] and finite_dom_ahm_\_aux[OF assms inv] show "(ahm_isEmpty hm, op_map_isEmpty m) \ bool_rel" unfolding ahm_isEmpty_def op_map_isEmpty_def by (simp add: \ card_eq_0_iff) qed lemma ahm_isSng_impl: assumes "is_bounded_hashcode Id (=) bhc" shows "(ahm_isSng, op_map_isSng) \ ahm_map_rel' bhc \ bool_rel" proof (intro fun_relI) fix hm m assume rel: "(hm,m) \ ahm_map_rel' bhc" obtain a n where [simp]: "hm = HashMap a n" by (cases hm) from rel have \: "m = ahm_\_aux bhc a" and inv: "ahm_invar_aux bhc n a" unfolding ahm_map_rel'_def br_def by (simp_all add: ahm_\_def2) note n_props[simp] = ahm_invar_aux_card_dom_ahm_\_auxD[OF assms inv,symmetric] note finite_dom[simp] = finite_dom_ahm_\_aux[OF assms inv] show "(ahm_isSng hm, op_map_isSng m) \ bool_rel" by (force simp add: \[symmetric] dom_eq_singleton_conv dest!: card_eq_SucD) qed lemma ahm_size_impl: assumes "is_bounded_hashcode Id (=) bhc" shows "(ahm_size, op_map_size) \ ahm_map_rel' bhc \ nat_rel" proof (intro fun_relI) fix hm m assume rel: "(hm,m) \ ahm_map_rel' bhc" obtain a n where [simp]: "hm = HashMap a n" by (cases hm) from rel have \: "m = ahm_\_aux bhc a" and inv: "ahm_invar_aux bhc n a" unfolding ahm_map_rel'_def br_def by (simp_all add: ahm_\_def2) from ahm_invar_aux_card_dom_ahm_\_auxD[OF assms inv,symmetric] show "(ahm_size hm, op_map_size m) \ nat_rel" unfolding ahm_isEmpty_def op_map_isEmpty_def by (simp add: \ card_eq_0_iff) qed lemma autoref_ahm_isEmpty[autoref_rules]: (*assumes eq: "GEN_OP eq (=) (Rk \ Rk \ bool_rel)"*) assumes bhc[unfolded autoref_tag_defs]: "SIDE_GEN_ALGO (is_bounded_hashcode Rk eq bhc)" shows "(ahm_isEmpty, op_map_isEmpty) \ \Rk,Rv\ahm_rel bhc \ bool_rel" proof (intro fun_relI) let ?bhc' = "abstract_bounded_hashcode Rk bhc" fix k k' a m' assume M: "(a,m') \ \Rk,Rv\ahm_rel bhc" (*from eq have "(eq,(=)) \ Rk \ Rk \ bool_rel" by simp*) from bhc have bhc': "is_bounded_hashcode Id (=) ?bhc'" by blast from M obtain a' where M1: "(a,a') \ \Rk,Rv\ahm_map_rel" and M2: "(a',m') \ ahm_map_rel' ?bhc'" unfolding ahm_rel_def by blast from relcompI[OF param_ahm_isEmpty[param_fo, OF M1] ahm_isEmpty_impl[param_fo, OF bhc' M2]] show "(ahm_isEmpty a, op_map_isEmpty m') \ bool_rel" by simp qed lemma autoref_ahm_isSng[autoref_rules]: assumes bhc[unfolded autoref_tag_defs]: "SIDE_GEN_ALGO (is_bounded_hashcode Rk eq bhc)" shows "(ahm_isSng, op_map_isSng) \ \Rk,Rv\ahm_rel bhc \ bool_rel" proof (intro fun_relI) let ?bhc' = "abstract_bounded_hashcode Rk bhc" fix k k' a m' assume M: "(a,m') \ \Rk,Rv\ahm_rel bhc" from bhc have bhc': "is_bounded_hashcode Id (=) ?bhc'" by blast from M obtain a' where M1: "(a,a') \ \Rk,Rv\ahm_map_rel" and M2: "(a',m') \ ahm_map_rel' ?bhc'" unfolding ahm_rel_def by blast from relcompI[OF param_ahm_isSng[param_fo, OF M1] ahm_isSng_impl[param_fo, OF bhc' M2]] show "(ahm_isSng a, op_map_isSng m') \ bool_rel" by simp qed lemma autoref_ahm_size[autoref_rules]: assumes bhc[unfolded autoref_tag_defs]: "SIDE_GEN_ALGO (is_bounded_hashcode Rk eq bhc)" shows "(ahm_size, op_map_size) \ \Rk,Rv\ahm_rel bhc \ nat_rel" proof (intro fun_relI) let ?bhc' = "abstract_bounded_hashcode Rk bhc" fix k k' a m' assume M: "(a,m') \ \Rk,Rv\ahm_rel bhc" from bhc have bhc': "is_bounded_hashcode Id (=) ?bhc'" by blast from M obtain a' where M1: "(a,a') \ \Rk,Rv\ahm_map_rel" and M2: "(a',m') \ ahm_map_rel' ?bhc'" unfolding ahm_rel_def by blast from relcompI[OF param_ahm_size[param_fo, OF M1] ahm_size_impl[param_fo, OF bhc' M2]] show "(ahm_size a, op_map_size m') \ nat_rel" by simp qed lemma ahm_map_rel_sv[relator_props]: assumes SK: "single_valued Rk" assumes SV: "single_valued Rv" shows "single_valued (\Rk, Rv\ahm_map_rel)" proof - from SK SV have 1: "single_valued (\\\Rk, Rv\prod_rel\list_rel\array_rel)" by (tagged_solver) thus ?thesis unfolding ahm_map_rel_def by (auto intro: single_valuedI dest: single_valuedD[OF 1]) qed lemma ahm_rel_sv[relator_props]: "\single_valued Rk; single_valued Rv\ \ single_valued (\Rk,Rv\ahm_rel bhc)" unfolding ahm_rel_def ahm_map_rel'_def by (tagged_solver (keep)) lemma rbt_map_rel_finite[relator_props]: assumes A[simplified]: "GEN_ALGO_tag (is_bounded_hashcode Rk eq bhc)" shows "finite_map_rel (\Rk,Rv\ahm_rel bhc)" unfolding ahm_rel_def finite_map_rel_def ahm_map_rel'_def br_def apply auto apply (case_tac y) apply (auto simp: ahm_\_def2) thm finite_dom_ahm_\_aux apply (rule finite_dom_ahm_\_aux) apply (rule abstract_bhc_is_bhc) apply (rule A) apply assumption done subsection \Proper iterator proofs\ lemma pi_ahm[icf_proper_iteratorI]: "proper_it (ahm_iteratei m) (ahm_iteratei m)" proof- obtain a n where [simp]: "m = HashMap a n" by (cases m) then obtain l where [simp]: "a = Array l" by (cases a) thus ?thesis unfolding proper_it_def list_map_iteratei_def by (simp add: ahm_iteratei_aux_def, blast) qed lemma pi'_ahm[icf_proper_iteratorI]: "proper_it' ahm_iteratei ahm_iteratei" by (rule proper_it'I, rule pi_ahm) (* hide_const (open) HashMap bucket_ok ahm_invar ahm_\ ahm_rehash hm_grow ahm_filled hide_type (open) hashmap *) lemmas autoref_ahm_rules = autoref_ahm_empty autoref_ahm_lookup autoref_ahm_update autoref_ahm_delete autoref_ahm_isEmpty autoref_ahm_isSng autoref_ahm_size lemmas autoref_ahm_rules_hashable[autoref_rules_raw] = autoref_ahm_rules[where Rk="Rk"] for Rk :: "(_\_::hashable) set" end diff --git a/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy b/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy --- a/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy +++ b/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy @@ -1,1114 +1,1114 @@ section \\isaheader{Implementing Unique Priority Queues by Annotated Lists}\ theory PrioUniqueByAnnotatedList imports "../spec/AnnotatedListSpec" "../spec/PrioUniqueSpec" begin text \ In this theory we use annotated lists to implement unique priority queues with totally ordered elements. This theory is written as a generic adapter from the AnnotatedList interface to the unique priority queue interface. The annotated list stores a sequence of elements annotated with priorities\footnote{Technically, the annotated list elements are of unit-type, and the annotations hold both, the priority queue elements and the priorities. This is required as we defined annotated lists to only sum up the elements annotations.} The monoids operations forms the maximum over the elements and the minimum over the priorities. The sequence of pairs is ordered by ascending elements' order. The insertion point for a new element, or the priority of an existing element can be found by splitting the sequence at the point where the maximum of the elements read so far gets bigger than the element to be inserted. The minimum priority can be read out as the sum over the whole sequence. Finding the element with minimum priority is done by splitting the sequence at the point where the minimum priority of the elements read so far becomes equal to the minimum priority of the whole sequence. \ subsection "Definitions" subsubsection "Monoid" datatype ('e, 'a) LP = Infty | LP 'e 'a fun p_unwrap :: "('e,'a) LP \ ('e \ 'a)" where "p_unwrap (LP e a) = (e , a)" fun p_min :: "('e::linorder, 'a::linorder) LP \ ('e, 'a) LP \ ('e, 'a) LP" where "p_min Infty Infty = Infty"| "p_min Infty (LP e a) = LP e a"| "p_min (LP e a) Infty = LP e a"| "p_min (LP e1 a) (LP e2 b) = (LP (max e1 e2) (min a b))" fun e_less_eq :: "'e \ ('e::linorder, 'a::linorder) LP \ bool" where "e_less_eq e Infty = False"| "e_less_eq e (LP e' _) = (e \ e')" text_raw\\paragraph{Instantiation of classes}\ \\\ lemma p_min_re_neut[simp]: "p_min a Infty = a" by (induct a) auto lemma p_min_le_neut[simp]: "p_min Infty a = a" by (induct a) auto lemma p_min_asso: "p_min (p_min a b) c = p_min a (p_min b c)" apply(induct a b rule: p_min.induct ) apply (auto) apply (induct c) apply (auto) apply (metis max.assoc) apply (metis min.assoc) done lemma lp_mono: "class.monoid_add p_min Infty" by unfold_locales (auto simp add: p_min_asso) instantiation LP :: (linorder,linorder) monoid_add begin definition zero_def: "0 == Infty" definition plus_def: "a+b == p_min a b" instance by intro_classes (auto simp add: p_min_asso zero_def plus_def) end fun p_less_eq :: "('e, 'a::linorder) LP \ ('e, 'a) LP \ bool" where "p_less_eq (LP e a) (LP f b) = (a \ b)"| "p_less_eq _ Infty = True"| "p_less_eq Infty (LP e a) = False" fun p_less :: "('e, 'a::linorder) LP \ ('e, 'a) LP \ bool" where "p_less (LP e a) (LP f b) = (a < b)"| "p_less (LP e a) Infty = True"| "p_less Infty _ = False" lemma p_less_le_not_le : "p_less x y \ p_less_eq x y \ \ (p_less_eq y x)" by (induct x y rule: p_less.induct) auto lemma p_order_refl : "p_less_eq x x" by (induct x) auto lemma p_le_inf : "p_less_eq Infty x \ x = Infty" by (induct x) auto lemma p_order_trans : "\p_less_eq x y; p_less_eq y z\ \ p_less_eq x z" apply (induct y z rule: p_less.induct) apply auto apply (induct x) apply auto apply (cases x) apply auto apply(induct x) apply (auto simp add: p_le_inf) apply (metis p_le_inf p_less_eq.simps(2)) done lemma p_linear2 : "p_less_eq x y \ p_less_eq y x" apply (induct x y rule: p_less_eq.induct) apply auto done instantiation LP :: (type, linorder) preorder begin definition plesseq_def: "less_eq = p_less_eq" definition pless_def: "less = p_less" instance apply (intro_classes) apply (simp only: p_less_le_not_le pless_def plesseq_def) apply (simp only: p_order_refl plesseq_def pless_def) apply (simp only: plesseq_def) apply (metis p_order_trans) done end subsubsection "Operations" definition aluprio_\ :: "('s \ (unit \ ('e::linorder,'a::linorder) LP) list) \ 's \ ('e::linorder \ 'a::linorder)" where "aluprio_\ \ ft == (map_of (map p_unwrap (map snd (\ ft))))" definition aluprio_invar :: "('s \ (unit \ ('c::linorder, 'd::linorder) LP) list) \ ('s \ bool) \ 's \ bool" where "aluprio_invar \ invar ft == invar ft \ (\ x\set (\ ft). snd x\Infty) \ sorted (map fst (map p_unwrap (map snd (\ ft)))) \ distinct (map fst (map p_unwrap (map snd (\ ft)))) " definition aluprio_empty where "aluprio_empty empt = empt" definition aluprio_isEmpty where "aluprio_isEmpty isEmpty = isEmpty" definition aluprio_insert :: "((('e::linorder,'a::linorder) LP \ bool) \ ('e,'a) LP \ 's \ ('s \ (unit \ ('e,'a) LP) \ 's)) \ ('s \ ('e,'a) LP) \ ('s \ bool) \ ('s \ 's \ 's) \ ('s \ unit \ ('e,'a) LP \ 's) \ 's \ 'e \ 'a \ 's" where " aluprio_insert splits annot isEmpty app consr s e a = (if e_less_eq e (annot s) \ \ isEmpty s then (let (l, (_,lp) , r) = splits (e_less_eq e) Infty s in (if e < fst (p_unwrap lp) then app (consr (consr l () (LP e a)) () lp) r else app (consr l () (LP e a)) r )) else consr s () (LP e a)) " definition aluprio_pop :: "((('e::linorder,'a::linorder) LP \ bool) \ ('e,'a) LP \ 's \ ('s \ (unit \ ('e,'a) LP) \ 's)) \ ('s \ ('e,'a) LP) \ ('s \ 's \ 's) \ 's \ 'e \'a \'s" where "aluprio_pop splits annot app s = (let (l, (_,lp) , r) = splits (\ x. x \ (annot s)) Infty s in (case lp of (LP e a) \ (e, a, app l r) ))" definition aluprio_prio :: "((('e::linorder,'a::linorder) LP \ bool) \ ('e,'a) LP \ 's \ ('s \ (unit \ ('e,'a) LP) \ 's)) \ ('s \ ('e,'a) LP) \ ('s \ bool) \ 's \ 'e \ 'a option" where " aluprio_prio splits annot isEmpty s e = (if e_less_eq e (annot s) \ \ isEmpty s then (let (l, (_,lp) , r) = splits (e_less_eq e) Infty s in (if e = fst (p_unwrap lp) then Some (snd (p_unwrap lp)) else None)) else None) " lemmas aluprio_defs = aluprio_invar_def aluprio_\_def aluprio_empty_def aluprio_isEmpty_def aluprio_insert_def aluprio_pop_def aluprio_prio_def subsection "Correctness" subsubsection "Auxiliary Lemmas" lemma p_linear: "(x::('e, 'a::linorder) LP) \ y \ y \ x" by (unfold plesseq_def) (simp only: p_linear2) lemma e_less_eq_mon1: "e_less_eq e x \ e_less_eq e (x + y)" apply (cases x) apply (auto simp add: plus_def) apply (cases y) apply (auto simp add: max.coboundedI1) done lemma e_less_eq_mon2: "e_less_eq e y \ e_less_eq e (x + y)" apply (cases x) apply (auto simp add: plus_def) apply (cases y) apply (auto simp add: max.coboundedI2) done lemmas e_less_eq_mon = e_less_eq_mon1 e_less_eq_mon2 lemma p_less_eq_mon: "(x::('e::linorder,'a::linorder) LP) \ z \ (x + y) \ z" apply(cases y) apply(auto simp add: plus_def) apply (cases x) apply (cases z) apply (auto simp add: plesseq_def) apply (cases z) apply (auto simp add: min.coboundedI1) done lemma p_less_eq_lem1: "\\ (x::('e::linorder,'a::linorder) LP) \ z; (x + y) \ z\ \ y \ z " apply (cases x,auto simp add: plus_def) apply (cases y, auto) apply (cases z, auto simp add: plesseq_def) apply (metis min_le_iff_disj) done lemma infadd: "x \ Infty \x + y \ Infty" apply (unfold plus_def) apply (induct x y rule: p_min.induct) apply auto done lemma e_less_eq_sum_list: "\\ e_less_eq e (sum_list xs)\ \ \x \ set xs. \ e_less_eq e x" proof (induct xs) case Nil thus ?case by simp next case (Cons a xs) hence "\ e_less_eq e (sum_list xs)" by (auto simp add: e_less_eq_mon) hence v1: "\x\set xs. \ e_less_eq e x" using Cons.hyps by simp from Cons.prems have "\ e_less_eq e a" by (auto simp add: e_less_eq_mon) with v1 show "\x\set (a#xs). \ e_less_eq e x" by simp qed lemma e_less_eq_p_unwrap: "\x \ Infty;\ e_less_eq e x\ \ fst (p_unwrap x) < e" by (cases x) auto lemma e_less_eq_refl : "b \ Infty \ e_less_eq (fst (p_unwrap b)) b" by (cases b) auto lemma e_less_eq_sum_list2: assumes "\x\set (\s). snd x \ Infty" "((), b) \ set (\s)" shows "e_less_eq (fst (p_unwrap b)) (sum_list (map snd (\s)))" apply(insert assms) apply (induct "\s") apply (auto simp add: zero_def e_less_eq_mon e_less_eq_refl) done lemma e_less_eq_lem1: "\\ e_less_eq e a;e_less_eq e (a + b)\ \ e_less_eq e b" apply (auto simp add: plus_def) apply (cases a) apply auto apply (cases b) apply auto apply (metis le_max_iff_disj) done lemma p_unwrap_less_sum: "snd (p_unwrap ((LP e aa) + b)) \ aa" apply (cases b) apply (auto simp add: plus_def) done lemma sum_list_less_elems: "\x\set xs. snd x \ Infty \ \y\set (map snd (map p_unwrap (map snd xs))). snd (p_unwrap (sum_list (map snd xs))) \ y" proof (induct xs) case Nil thus ?case by simp next case (Cons a as) thus ?case apply auto apply (cases "(snd a)" rule: p_unwrap.cases) apply auto apply (cases "sum_list (map snd as)") apply auto apply (metis linorder_linear p_min_re_neut p_unwrap.simps plus_def[abs_def] snd_eqD) apply (auto simp add: p_unwrap_less_sum) apply (unfold plus_def) apply (cases "(snd a, sum_list (map snd as))" rule: p_min.cases) apply auto apply (cases "map snd as") apply (auto simp add: infadd) apply (metis min.coboundedI2 snd_conv) done qed lemma distinct_sortet_list_app: "\sorted xs; distinct xs; xs = as @ b # cs\ \ \ x\ set cs. b < x" by (metis distinct.simps(2) distinct_append antisym_conv2 sorted_wrt.simps(2) sorted_append) lemma distinct_sorted_list_lem1: assumes "sorted xs" "sorted ys" "distinct xs" "distinct ys" " \ x \ set xs. x < e" " \ y \ set ys. e < y" shows "sorted (xs @ e # ys)" "distinct (xs @ e # ys)" proof - from assms (5,6) have "\x\set xs. \y\set ys. x \ y" by force thus "sorted (xs @ e # ys)" using assms by (auto simp add: sorted_append) have "set xs \ set ys = {}" using assms (5,6) by force thus "distinct (xs @ e # ys)" using assms by (auto) qed lemma distinct_sorted_list_lem2: assumes "sorted xs" "sorted ys" "distinct xs" "distinct ys" "e < e'" " \ x \ set xs. x < e" " \ y \ set ys. e' < y" shows "sorted (xs @ e # e' # ys)" "distinct (xs @ e # e' # ys)" proof - have "sorted (e' # ys)" "distinct (e' # ys)" "\ y \ set (e' # ys). e < y" using assms(2,4,5,7) by (auto) thus "sorted (xs @ e # e' # ys)" "distinct (xs @ e # e' # ys)" using assms(1,3,6) distinct_sorted_list_lem1[of xs "e' # ys" e] by auto qed lemma map_of_distinct_upd: "x \ set (map fst xs) \ [x \ y] ++ map_of xs = (map_of xs) (x \ y)" by (induct xs) (auto simp add: fun_upd_twist) lemma map_of_distinct_upd2: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd3: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ (x,y') # ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd4: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ ys) = (map_of (xs @ (x,y) # ys))(x := None)" apply(insert assms) apply(induct xs) apply clarsimp apply (metis dom_map_of_conv_image_fst fun_upd_None_restrict restrict_complement_singleton_eq restrict_map_self) apply (auto simp add: map_of_eq_None_iff) [] done lemma map_of_distinct_lookup: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) x = Some y" proof - have "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys)) (x \ y)" using assms map_of_distinct_upd2 by simp thus ?thesis by simp qed lemma ran_distinct: assumes dist: "distinct (map fst al)" shows "ran (map_of al) = snd ` set al" using assms proof (induct al) case Nil then show ?case by simp next case (Cons kv al) then have "ran (map_of al) = snd ` set al" by simp moreover from Cons.prems have "map_of al (fst kv) = None" by (simp add: map_of_eq_None_iff) ultimately show ?case by (simp only: map_of.simps ran_map_upd) simp qed subsubsection "Finite" lemma aluprio_finite_correct: "uprio_finite (aluprio_\ \) (aluprio_invar \ invar)" by(unfold_locales) (simp add: aluprio_defs finite_dom_map_of) subsubsection "Empty" lemma aluprio_empty_correct: assumes "al_empty \ invar empt" shows "uprio_empty (aluprio_\ \) (aluprio_invar \ invar) (aluprio_empty empt)" proof - interpret al_empty \ invar empt by fact show ?thesis apply (unfold_locales) apply (auto simp add: empty_correct aluprio_defs) done qed subsubsection "Is Empty" lemma aluprio_isEmpty_correct: assumes "al_isEmpty \ invar isEmpty" shows "uprio_isEmpty (aluprio_\ \) (aluprio_invar \ invar) (aluprio_isEmpty isEmpty)" proof - interpret al_isEmpty \ invar isEmpty by fact show ?thesis apply (unfold_locales) apply (auto simp add: aluprio_defs isEmpty_correct) done qed subsubsection "Insert" lemma annot_inf: assumes A: "invar s" "\x\set (\ s). snd x \ Infty" "al_annot \ invar annot" shows "annot s = Infty \ \ s = [] " proof - from A have invs: "invar s" by (simp add: aluprio_defs) interpret al_annot \ invar annot by fact show "annot s = Infty \ \ s = []" proof (cases "\ s = []") case True hence "map snd (\ s) = []" by simp hence "sum_list (map snd (\ s)) = Infty" by (auto simp add: zero_def) with invs have "annot s = Infty" by (auto simp add: annot_correct) with True show ?thesis by simp next case False hence " \x xs. (\ s) = x # xs" by (cases "\ s") auto from this obtain x xs where [simp]: "(\ s) = x # xs" by blast from this assms(2) have "snd x \ Infty" by (auto simp add: aluprio_defs) hence "sum_list (map snd (\ s)) \ Infty" by (auto simp add: infadd) thus ?thesis using annot_correct invs False by simp qed qed lemma e_less_eq_annot: assumes "al_annot \ invar annot" "invar s" "\x\set (\ s). snd x \ Infty" "\ e_less_eq e (annot s)" shows "\x \ set (map (fst \ (p_unwrap \ snd)) (\ s)). x < e" proof - interpret al_annot \ invar annot by fact from assms(2) have "annot s = sum_list (map snd (\ s))" by (auto simp add: annot_correct) with assms(4) have "\x \ set (map snd (\ s)). \ e_less_eq e x" by (metis e_less_eq_sum_list) with assms(3) show ?thesis by (auto simp add: e_less_eq_p_unwrap) qed lemma aluprio_insert_correct: assumes "al_splits \ invar splits" "al_annot \ invar annot" "al_isEmpty \ invar isEmpty" "al_app \ invar app" "al_consr \ invar consr" shows "uprio_insert (aluprio_\ \) (aluprio_invar \ invar) (aluprio_insert splits annot isEmpty app consr)" proof - interpret al_splits \ invar splits by fact interpret al_annot \ invar annot by fact interpret al_isEmpty \ invar isEmpty by fact interpret al_app \ invar app by fact interpret al_consr \ invar consr by fact show ?thesis proof (unfold_locales, unfold aluprio_defs, goal_cases) case g1asms: (1 s e a) thus ?case proof (cases "e_less_eq e (annot s) \ \ isEmpty s") case False with g1asms show ?thesis apply (auto simp add: consr_correct ) proof goal_cases case prems: 1 with assms(2) have "\x \ set (map (fst \ (p_unwrap \ snd)) (\ s)). x < e" by (simp add: e_less_eq_annot) with prems(3) show ?case by(auto simp add: sorted_append) next case prems: 2 hence "annot s = sum_list (map snd (\ s))" by (simp add: annot_correct) with prems show ?case by (auto simp add: e_less_eq_sum_list2) next case prems: 3 hence "\ s = []" by (auto simp add: isEmpty_correct) thus ?case by simp next case prems: 4 hence "\ s = []" by (auto simp add: isEmpty_correct) with prems show ?case by simp qed next case True note T1 = this obtain l uu lp r where l_lp_r: "(splits (e_less_eq e) Infty s) = (l, ((), lp), r) " by (cases "splits (e_less_eq e) Infty s", auto) note v2 = splits_correct[of s "e_less_eq e" Infty l "()" lp r] have v3: "invar s" "\ e_less_eq e Infty" "e_less_eq e (Infty + sum_list (map snd (\ s)))" using T1 g1asms annot_correct by (auto simp add: plus_def) have v4: "\ s = \ l @ ((), lp) # \ r" "\ e_less_eq e (Infty + sum_list (map snd (\ l)))" "e_less_eq e (Infty + sum_list (map snd (\ l)) + lp)" "invar l" "invar r" using v2[OF v3(1) _ v3(2) v3(3) l_lp_r] e_less_eq_mon(1) by auto hence v5: "e_less_eq e lp" by (metis e_less_eq_lem1) hence v6: "e \ (fst (p_unwrap lp))" by (cases lp) auto have "(Infty + sum_list (map snd (\ l))) = (annot l)" by (metis add_0_left annot_correct v4(4) zero_def) hence v7:"\ e_less_eq e (annot l)" using v4(2) by simp have "\x\set (\ l). snd x \ Infty" using g1asms v4(1) by simp hence v7: "\x \ set (map (fst \ (p_unwrap \ snd)) (\ l)). x < e" using v4(4) v7 assms(2) by(simp add: e_less_eq_annot) have v8:"map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ fst(p_unwrap lp) # map fst (map p_unwrap (map snd (\ r)))" using v4(1) by simp note distinct_sortet_list_app[of "map fst (map p_unwrap (map snd (\ s)))" "map fst (map p_unwrap (map snd (\ l)))" "fst(p_unwrap lp)" "map fst (map p_unwrap (map snd (\ r)))"] hence v9: "\ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). fst(p_unwrap lp) < x" using v4(1) g1asms v8 by auto have v10: "sorted (map fst (map p_unwrap (map snd (\ l))))" "distinct (map fst (map p_unwrap (map snd (\ l))))" "sorted (map fst (map p_unwrap (map snd (\ r))))" "distinct (map fst (map p_unwrap (map snd (\ l))))" using g1asms v8 by (auto simp add: sorted_append) from l_lp_r T1 g1asms show ?thesis proof (fold aluprio_insert_def, cases "e < fst (p_unwrap lp)") case True hence v11: "aluprio_insert splits annot isEmpty app consr s e a = app (consr (consr l () (LP e a)) () lp) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v12: "invar (app (consr (consr l () (LP e a)) () lp) r)" using v4(4,5) by (auto simp add: app_correct consr_correct) have v13: "\ (app (consr (consr l () (LP e a)) () lp) r) = \ l @ ((),(LP e a)) # ((), lp) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) hence v14: "(\x\set (\ (app (consr (consr l () (LP e a)) () lp) r)). snd x \ Infty)" using g1asms v4(1) by auto have v15: "e = fst(p_unwrap (LP e a))" by simp hence v16: "sorted (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # ((), lp) # \ r))))" "distinct (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # ((), lp) # \ r))))" using v10(1,3) v7 True v9 v4(1) g1asms distinct_sorted_list_lem2 by (auto simp add: sorted_append) thus "invar (aluprio_insert splits annot isEmpty app consr s e a) \ (\x\set (\ (aluprio_insert splits annot isEmpty app consr s e a)). snd x \ Infty) \ sorted (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a))))) \ distinct (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))))" using v11 v12 v13 v14 by simp next case False hence v11: "aluprio_insert splits annot isEmpty app consr s e a = app (consr l () (LP e a)) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v12: "invar (app (consr l () (LP e a)) r)" using v4(4,5) by (auto simp add: app_correct consr_correct) have v13: "\ (app (consr l () (LP e a)) r) = \ l @ ((),(LP e a)) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) hence v14: "(\x\set (\ (app (consr l () (LP e a)) r)). snd x \ Infty)" using g1asms v4(1) by auto have v15: "e = fst(p_unwrap (LP e a))" by simp have v16: "e = fst(p_unwrap lp)" using False v5 by (cases lp) auto hence v17: "sorted (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # \ r))))" "distinct (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # \ r))))" using v16 v15 v10(1,3) v7 True v9 v4(1) g1asms distinct_sorted_list_lem1 by (auto simp add: sorted_append) thus "invar (aluprio_insert splits annot isEmpty app consr s e a) \ (\x\set (\ (aluprio_insert splits annot isEmpty app consr s e a)). snd x \ Infty) \ sorted (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a))))) \ distinct (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))))" using v11 v12 v13 v14 by simp qed qed next case g1asms: (2 s e a) thus ?case proof (cases "e_less_eq e (annot s) \ \ isEmpty s") case False with g1asms show ?thesis apply (auto simp add: consr_correct) proof goal_cases case prems: 1 with assms(2) have "\x \ set (map (fst \ (p_unwrap \ snd)) (\ s)). x < e" by (simp add: e_less_eq_annot) hence "e \ set (map fst ((map (p_unwrap \ snd)) (\ s)))" by auto thus ?case by (auto simp add: map_of_distinct_upd) next case prems: 2 hence "\ s = []" by (auto simp add: isEmpty_correct) thus ?case by simp qed next case True note T1 = this obtain l lp r where l_lp_r: "(splits (e_less_eq e) Infty s) = (l, ((), lp), r) " by (cases "splits (e_less_eq e) Infty s", auto) note v2 = splits_correct[of s "e_less_eq e" Infty l "()" lp r] have v3: "invar s" "\ e_less_eq e Infty" "e_less_eq e (Infty + sum_list (map snd (\ s)))" using T1 g1asms annot_correct by (auto simp add: plus_def) have v4: "\ s = \ l @ ((), lp) # \ r" "\ e_less_eq e (Infty + sum_list (map snd (\ l)))" "e_less_eq e (Infty + sum_list (map snd (\ l)) + lp)" "invar l" "invar r" using v2[OF v3(1) _ v3(2) v3(3) l_lp_r] e_less_eq_mon(1) by auto hence v5: "e_less_eq e lp" by (metis e_less_eq_lem1) hence v6: "e \ (fst (p_unwrap lp))" by (cases lp) auto have "(Infty + sum_list (map snd (\ l))) = (annot l)" by (metis add_0_left annot_correct v4(4) zero_def) hence v7:"\ e_less_eq e (annot l)" using v4(2) by simp have "\x\set (\ l). snd x \ Infty" using g1asms v4(1) by simp hence v7: "\x \ set (map (fst \ (p_unwrap \ snd)) (\ l)). x < e" using v4(4) v7 assms(2) by(simp add: e_less_eq_annot) have v8:"map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ fst(p_unwrap lp) # map fst (map p_unwrap (map snd (\ r)))" using v4(1) by simp note distinct_sortet_list_app[of "map fst (map p_unwrap (map snd (\ s)))" "map fst (map p_unwrap (map snd (\ l)))" "fst(p_unwrap lp)" "map fst (map p_unwrap (map snd (\ r)))"] hence v9: " \ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). fst(p_unwrap lp) < x" using v4(1) g1asms v8 by auto hence v10: " \ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). e < x" using v6 by auto have v11: "e \ set (map fst (map p_unwrap (map snd (\ l))))" "e \ set (map fst (map p_unwrap (map snd (\ r))))" using v7 v10 v8 g1asms by auto from l_lp_r T1 g1asms show ?thesis proof (fold aluprio_insert_def, cases "e < fst (p_unwrap lp)") case True hence v12: "aluprio_insert splits annot isEmpty app consr s e a = app (consr (consr l () (LP e a)) () lp) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v13: "\ (app (consr (consr l () (LP e a)) () lp) r) = \ l @ ((),(LP e a)) # ((), lp) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) have v14: "e = fst(p_unwrap (LP e a))" by simp have v15: "e \ set (map fst (map p_unwrap (map snd(((),lp)#\ r))))" using v11(2) True by auto note map_of_distinct_upd2[OF v11(1) v15] thus "map_of (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))) - = map_of (map p_unwrap (map snd (\ s)))(e \ a)" + = (map_of (map p_unwrap (map snd (\ s))))(e \ a)" using v12 v13 v4(1) by simp next case False hence v12: "aluprio_insert splits annot isEmpty app consr s e a = app (consr l () (LP e a)) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v13: "\ (app (consr l () (LP e a)) r) = \ l @ ((),(LP e a)) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) have v14: "e = fst(p_unwrap lp)" using False v5 by (cases lp) auto note v15 = map_of_distinct_upd3[OF v11(1) v11(2)] have v16:"(map p_unwrap (map snd (\ s))) = (map p_unwrap (map snd (\ l))) @ (e,snd(p_unwrap lp)) # (map p_unwrap (map snd (\ r)))" using v4(1) v14 by simp note v15[of a "snd(p_unwrap lp)"] thus "map_of (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))) - = map_of (map p_unwrap (map snd (\ s)))(e \ a)" + = (map_of (map p_unwrap (map snd (\ s))))(e \ a)" using v12 v13 v16 by simp qed qed qed qed subsubsection "Prio" lemma aluprio_prio_correct: assumes "al_splits \ invar splits" "al_annot \ invar annot" "al_isEmpty \ invar isEmpty" shows "uprio_prio (aluprio_\ \) (aluprio_invar \ invar) (aluprio_prio splits annot isEmpty)" proof - interpret al_splits \ invar splits by fact interpret al_annot \ invar annot by fact interpret al_isEmpty \ invar isEmpty by fact show ?thesis proof (unfold_locales) fix s e assume inv1: "aluprio_invar \ invar s" hence sinv: "invar s" "(\ x\set (\ s). snd x\Infty)" "sorted (map fst (map p_unwrap (map snd (\ s))))" "distinct (map fst (map p_unwrap (map snd (\ s))))" by (auto simp add: aluprio_defs) show "aluprio_prio splits annot isEmpty s e = aluprio_\ \ s e" proof(cases "e_less_eq e (annot s) \ \ isEmpty s") case False note F1 = this thus ?thesis proof(cases "isEmpty s") case True hence "\ s = []" using sinv isEmpty_correct by simp hence "aluprio_\ \ s = Map.empty" by (simp add:aluprio_defs) hence "aluprio_\ \ s e = None" by simp thus "aluprio_prio splits annot isEmpty s e = aluprio_\ \ s e" using F1 by (auto simp add: aluprio_defs) next case False hence v3:"\ e_less_eq e (annot s)" using F1 by simp note v4=e_less_eq_annot[OF assms(2)] note v4[OF sinv(1) sinv(2) v3] hence v5:"e\set (map (fst \ (p_unwrap \ snd)) (\ s))" by auto hence "map_of (map (p_unwrap \ snd) (\ s)) e = None" using map_of_eq_None_iff by (metis map_map map_of_eq_None_iff set_map v5) thus "aluprio_prio splits annot isEmpty s e = aluprio_\ \ s e" using F1 by (auto simp add: aluprio_defs) qed next case True note T1 = this obtain l uu lp r where l_lp_r: "(splits (e_less_eq e) Infty s) = (l, ((), lp), r) " by (cases "splits (e_less_eq e) Infty s", auto) note v2 = splits_correct[of s "e_less_eq e" Infty l "()" lp r] have v3: "invar s" "\ e_less_eq e Infty" "e_less_eq e (Infty + sum_list (map snd (\ s)))" using T1 sinv annot_correct by (auto simp add: plus_def) have v4: "\ s = \ l @ ((), lp) # \ r" "\ e_less_eq e (Infty + sum_list (map snd (\ l)))" "e_less_eq e (Infty + sum_list (map snd (\ l)) + lp)" "invar l" "invar r" using v2[OF v3(1) _ v3(2) v3(3) l_lp_r] e_less_eq_mon(1) by auto hence v5: "e_less_eq e lp" by (metis e_less_eq_lem1) hence v6: "e \ (fst (p_unwrap lp))" by (cases lp) auto have "(Infty + sum_list (map snd (\ l))) = (annot l)" by (metis add_0_left annot_correct v4(4) zero_def) hence v7:"\ e_less_eq e (annot l)" using v4(2) by simp have "\x\set (\ l). snd x \ Infty" using sinv v4(1) by simp hence v7: "\x \ set (map (fst \ (p_unwrap \ snd)) (\ l)). x < e" using v4(4) v7 assms(2) by(simp add: e_less_eq_annot) have v8:"map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ fst(p_unwrap lp) # map fst (map p_unwrap (map snd (\ r)))" using v4(1) by simp note distinct_sortet_list_app[of "map fst (map p_unwrap (map snd (\ s)))" "map fst (map p_unwrap (map snd (\ l)))" "fst(p_unwrap lp)" "map fst (map p_unwrap (map snd (\ r)))"] hence v9: "\ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). fst(p_unwrap lp) < x" using v4(1) sinv v8 by auto hence v10: " \ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). e < x" using v6 by auto have v11: "e \ set (map fst (map p_unwrap (map snd (\ l))))" "e \ set (map fst (map p_unwrap (map snd (\ r))))" using v7 v10 v8 sinv by auto from l_lp_r T1 sinv show ?thesis proof (cases "e = fst (p_unwrap lp)") case False have v12: "e \ set (map fst (map p_unwrap (map snd(\ s))))" using v11 False v4(1) by auto hence "map_of (map (p_unwrap \ snd) (\ s)) e = None" using map_of_eq_None_iff by (metis map_map map_of_eq_None_iff set_map v12) thus ?thesis using T1 False l_lp_r by (auto simp add: aluprio_defs) next case True have v12: "map (p_unwrap \ snd) (\ s) = map p_unwrap (map snd (\ l)) @ (e,snd (p_unwrap lp)) # map p_unwrap (map snd (\ r))" using v4(1) True by simp note map_of_distinct_lookup[OF v11] hence "map_of (map (p_unwrap \ snd) (\ s)) e = Some (snd (p_unwrap lp))" using v12 by simp thus ?thesis using T1 True l_lp_r by (auto simp add: aluprio_defs) qed qed qed qed subsubsection "Pop" lemma aluprio_pop_correct: assumes "al_splits \ invar splits" "al_annot \ invar annot" "al_app \ invar app" shows "uprio_pop (aluprio_\ \) (aluprio_invar \ invar) (aluprio_pop splits annot app)" proof - interpret al_splits \ invar splits by fact interpret al_annot \ invar annot by fact interpret al_app \ invar app by fact show ?thesis proof (unfold_locales) fix s e a s' assume A: "aluprio_invar \ invar s" "aluprio_\ \ s \ Map.empty" "aluprio_pop splits annot app s = (e, a, s')" hence v1: "\ s \ []" by (auto simp add: aluprio_defs) obtain l lp r where l_lp_r: "splits (\ x. x\annot s) Infty s = (l,((),lp),r)" by (cases "splits (\ x. x\annot s) Infty s", auto) have invs: "invar s" "(\x\set (\ s). snd x \ Infty)" "sorted (map fst (map p_unwrap (map snd (\ s))))" "distinct (map fst (map p_unwrap (map snd (\ s))))" using A by (auto simp add:aluprio_defs) note a1 = annot_inf[of invar s \ annot] note a1[OF invs(1) invs(2) assms(2)] hence v2: "annot s \ Infty" using v1 by simp hence v3: "\ Infty \ annot s" by(cases "annot s") (auto simp add: plesseq_def) have v4: "annot s = sum_list (map snd (\ s))" by (auto simp add: annot_correct invs(1)) hence v5: "(Infty + sum_list (map snd (\ s))) \ annot s" by (auto simp add: plus_def) note p_mon = p_less_eq_mon[of _ "annot s"] note v6 = splits_correct[OF invs(1)] note v7 = v6[of "\ x. x \ annot s"] note v7[OF _ v3 v5 l_lp_r] p_mon hence v8: " \ s = \ l @ ((), lp) # \ r" "\ Infty + sum_list (map snd (\ l)) \ annot s" "Infty + sum_list (map snd (\ l)) + lp \ annot s" "invar l" "invar r" by auto hence v9: "lp \ Infty" using invs(2) by auto hence v10: "s' = app l r" "(e,a) = p_unwrap lp" using l_lp_r A(3) apply (auto simp add: aluprio_defs) apply (cases lp) apply auto apply (cases lp) apply auto done have "lp \ annot s" using v8(2,3) p_less_eq_lem1 by auto hence v11: "a \ snd (p_unwrap (annot s))" using v10(2) v2 v9 apply (cases "annot s") apply auto apply (cases lp) apply (auto simp add: plesseq_def) done note sum_list_less_elems[OF invs(2)] hence v12: "\y\set (map snd (map p_unwrap (map snd (\ s)))). a \ y" using v4 v11 by auto have "ran (aluprio_\ \ s) = set (map snd (map p_unwrap (map snd (\ s))))" using ran_distinct[OF invs(4)] apply (unfold aluprio_defs) apply (simp only: set_map) done hence ziel1: "\y\ran (aluprio_\ \ s). a \ y" using v12 by simp have v13: "map p_unwrap (map snd (\ s)) = map p_unwrap (map snd (\ l)) @ (e,a) # map p_unwrap (map snd (\ r))" using v8(1) v10 by auto hence v14: "map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ e # map fst (map p_unwrap (map snd (\ r)))" by auto hence v15: "e \ set (map fst (map p_unwrap (map snd (\ l))))" "e \ set (map fst (map p_unwrap (map snd (\ r))))" using invs(4) by auto note map_of_distinct_lookup[OF v15] note this[of a] hence ziel2: "aluprio_\ \ s e = Some a" using v13 by (unfold aluprio_defs, auto) have v16: "\ s' = \ l @ \ r" "invar s'" using v8(4,5) app_correct v10 by auto note map_of_distinct_upd4[OF v15] note this[of a] hence ziel3: "aluprio_\ \ s' = (aluprio_\ \ s)(e := None)" unfolding aluprio_defs using v16(1) v13 by auto have ziel4: "aluprio_invar \ invar s'" using v16 v8(1) invs(2,3,4) unfolding aluprio_defs by (auto simp add: sorted_append) show "aluprio_invar \ invar s' \ aluprio_\ \ s' = (aluprio_\ \ s)(e := None) \ aluprio_\ \ s e = Some a \ (\y\ran (aluprio_\ \ s). a \ y)" using ziel1 ziel2 ziel3 ziel4 by simp qed qed lemmas aluprio_correct = aluprio_finite_correct aluprio_empty_correct aluprio_isEmpty_correct aluprio_insert_correct aluprio_pop_correct aluprio_prio_correct locale aluprio_defs = StdALDefs ops for ops :: "(unit,('e::linorder,'a::linorder) LP,'s) alist_ops" begin definition [icf_rec_def]: "aluprio_ops \ \ upr_\ = aluprio_\ \, upr_invar = aluprio_invar \ invar, upr_empty = aluprio_empty empty, upr_isEmpty = aluprio_isEmpty isEmpty, upr_insert = aluprio_insert splits annot isEmpty app consr, upr_pop = aluprio_pop splits annot app, upr_prio = aluprio_prio splits annot isEmpty \" end locale aluprio = aluprio_defs ops + StdAL ops for ops :: "(unit,('e::linorder,'a::linorder) LP,'s) alist_ops" begin lemma aluprio_ops_impl: "StdUprio aluprio_ops" apply (rule StdUprio.intro) apply (simp_all add: icf_rec_unf) apply (rule aluprio_correct) apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] done end end diff --git a/thys/Collections/ICF/impl/ArrayHashMap_Impl.thy b/thys/Collections/ICF/impl/ArrayHashMap_Impl.thy --- a/thys/Collections/ICF/impl/ArrayHashMap_Impl.thy +++ b/thys/Collections/ICF/impl/ArrayHashMap_Impl.thy @@ -1,717 +1,717 @@ (* Title: Isabelle Collections Library Author: Andreas Lochbihler Maintainer: Andreas Lochbihler *) section \\isaheader{Array-based hash map implementation}\ theory ArrayHashMap_Impl imports "../../Lib/HashCode" "../../Lib/Code_Target_ICF" "../../Lib/Diff_Array" "../gen_algo/ListGA" ListMapImpl "../../Iterator/Array_Iterator" begin text \Misc.\ setup Locale_Code.open_block interpretation a_idx_it: idx_iteratei_loc list_of_array "\_. True" array_length array_get apply unfold_locales apply (case_tac [!] s) [2] apply auto done setup Locale_Code.close_block (* lemma idx_iteratei_aux_array_get_Array_conv_nth: "idx_iteratei_aux array_get sz i (Array xs) c f \ = idx_iteratei_aux (!) sz i xs c f \" apply(induct get\"(!) :: 'b list \ nat \ 'b" sz i xs c f \ rule: idx_iteratei_aux.induct) apply(subst (1 2) idx_iteratei_aux.simps) apply simp done lemma idx_iteratei_array_get_Array_conv_nth: "idx_iteratei array_get array_length (Array xs) = idx_iteratei nth length xs" by(simp add: idx_iteratei_def fun_eq_iff idx_iteratei_aux_array_get_Array_conv_nth) lemma idx_iteratei_aux_nth_conv_foldli_drop: fixes xs :: "'b list" assumes "i \ length xs" shows "idx_iteratei_aux (!) (length xs) i xs c f \ = foldli (drop (length xs - i) xs) c f \" using assms proof(induct get\"(!) :: 'b list \ nat \ 'b" sz\"length xs" i xs c f \ rule: idx_iteratei_aux.induct) case (1 i l c f \) show ?case proof(cases "i = 0 \ \ c \") case True thus ?thesis by(subst idx_iteratei_aux.simps)(auto) next case False hence i: "i > 0" and c: "c \" by auto hence "idx_iteratei_aux (!) (length l) i l c f \ = idx_iteratei_aux (!) (length l) (i - 1) l c f (f (l ! (length l - i)) \)" by(subst idx_iteratei_aux.simps) simp also have "\ = foldli (drop (length l - (i - 1)) l) c f (f (l ! (length l - i)) \)" using `i \ length l` i c by -(rule 1, auto) also from `i \ length l` i have "drop (length l - i) l = (l ! (length l - i)) # drop (length l - (i - 1)) l" by(subst Cons_nth_drop_Suc[symmetric])(simp_all, metis Suc_eq_plus1_left add_diff_assoc) hence "foldli (drop (length l - (i - 1)) l) c f (f (l ! (length l - i)) \) = foldli (drop (length l - i) l) c f \" using c by simp finally show ?thesis . qed qed lemma idx_iteratei_nth_length_conv_foldli: "idx_iteratei nth length = foldli" by(rule ext)+(simp add: idx_iteratei_def idx_iteratei_aux_nth_conv_foldli_drop) *) subsection \Type definition and primitive operations\ definition load_factor :: nat \ \in percent\ where "load_factor = 75" text \ We do not use @{typ "('k, 'v) assoc_list"} for the buckets but plain lists of key-value pairs. This speeds up rehashing because we then do not have to go through the abstract operations. \ datatype ('key, 'val) hashmap = HashMap "('key \ 'val) list array" "nat" subsection \Operations\ definition new_hashmap_with :: "nat \ ('key :: hashable, 'val) hashmap" where "\size. new_hashmap_with size = HashMap (new_array [] size) 0" definition ahm_empty :: "unit \ ('key :: hashable, 'val) hashmap" where "ahm_empty \ \_. new_hashmap_with (def_hashmap_size TYPE('key))" definition bucket_ok :: "nat \ nat \ (('key :: hashable) \ 'val) list \ bool" where "bucket_ok len h kvs = (\k \ fst ` set kvs. bounded_hashcode_nat len k = h)" definition ahm_invar_aux :: "nat \ (('key :: hashable) \ 'val) list array \ bool" where "ahm_invar_aux n a \ (\h. h < array_length a \ bucket_ok (array_length a) h (array_get a h) \ distinct (map fst (array_get a h))) \ array_foldl (\_ n kvs. n + size kvs) 0 a = n \ array_length a > 1" primrec ahm_invar :: "('key :: hashable, 'val) hashmap \ bool" where "ahm_invar (HashMap a n) = ahm_invar_aux n a" definition ahm_\_aux :: "(('key :: hashable) \ 'val) list array \ 'key \ 'val option" where [simp]: "ahm_\_aux a k = map_of (array_get a (bounded_hashcode_nat (array_length a) k)) k" primrec ahm_\ :: "('key :: hashable, 'val) hashmap \ 'key \ 'val option" where "ahm_\ (HashMap a _) = ahm_\_aux a" definition ahm_lookup :: "'key \ ('key :: hashable, 'val) hashmap \ 'val option" where "ahm_lookup k hm = ahm_\ hm k" primrec ahm_iteratei_aux :: "((('key :: hashable) \ 'val) list array) \ ('key \ 'val, '\) set_iterator" where "ahm_iteratei_aux (Array xs) c f = foldli (concat xs) c f" primrec ahm_iteratei :: "(('key :: hashable, 'val) hashmap) \ (('key \ 'val), '\) set_iterator" where "ahm_iteratei (HashMap a n) = ahm_iteratei_aux a" definition ahm_rehash_aux' :: "nat \ 'key \ 'val \ (('key :: hashable) \ 'val) list array \ ('key \ 'val) list array" where "ahm_rehash_aux' n kv a = (let h = bounded_hashcode_nat n (fst kv) in array_set a h (kv # array_get a h))" definition ahm_rehash_aux :: "(('key :: hashable) \ 'val) list array \ nat \ ('key \ 'val) list array" where "ahm_rehash_aux a sz = ahm_iteratei_aux a (\x. True) (ahm_rehash_aux' sz) (new_array [] sz)" primrec ahm_rehash :: "('key :: hashable, 'val) hashmap \ nat \ ('key, 'val) hashmap" where "ahm_rehash (HashMap a n) sz = HashMap (ahm_rehash_aux a sz) n" primrec hm_grow :: "('key :: hashable, 'val) hashmap \ nat" where "hm_grow (HashMap a n) = 2 * array_length a + 3" primrec ahm_filled :: "('key :: hashable, 'val) hashmap \ bool" where "ahm_filled (HashMap a n) = (array_length a * load_factor \ n * 100)" primrec ahm_update_aux :: "('key :: hashable, 'val) hashmap \ 'key \ 'val \ ('key, 'val) hashmap" where "ahm_update_aux (HashMap a n) k v = (let h = bounded_hashcode_nat (array_length a) k; m = array_get a h; insert = map_of m k = None in HashMap (array_set a h (AList.update k v m)) (if insert then n + 1 else n))" definition ahm_update :: "'key \ 'val \ ('key :: hashable, 'val) hashmap \ ('key, 'val) hashmap" where "ahm_update k v hm = (let hm' = ahm_update_aux hm k v in (if ahm_filled hm' then ahm_rehash hm' (hm_grow hm') else hm'))" primrec ahm_delete :: "'key \ ('key :: hashable, 'val) hashmap \ ('key, 'val) hashmap" where "ahm_delete k (HashMap a n) = (let h = bounded_hashcode_nat (array_length a) k; m = array_get a h; deleted = (map_of m k \ None) in HashMap (array_set a h (AList.delete k m)) (if deleted then n - 1 else n))" lemma hm_grow_gt_1 [iff]: "Suc 0 < hm_grow hm" by(cases hm)(simp) lemma bucket_ok_Nil [simp]: "bucket_ok len h [] = True" by(simp add: bucket_ok_def) lemma bucket_okD: "\ bucket_ok len h xs; (k, v) \ set xs \ \ bounded_hashcode_nat len k = h" by(auto simp add: bucket_ok_def) lemma bucket_okI: "(\k. k \ fst ` set kvs \ bounded_hashcode_nat len k = h) \ bucket_ok len h kvs" by(simp add: bucket_ok_def) subsection \@{term ahm_invar}\ lemma ahm_invar_auxE: assumes "ahm_invar_aux n a" obtains "\h. h < array_length a \ bucket_ok (array_length a) h (array_get a h) \ distinct (map fst (array_get a h))" and "n = array_foldl (\_ n kvs. n + length kvs) 0 a" and "array_length a > 1" using assms unfolding ahm_invar_aux_def by blast lemma ahm_invar_auxI: "\ \h. h < array_length a \ bucket_ok (array_length a) h (array_get a h); \h. h < array_length a \ distinct (map fst (array_get a h)); n = array_foldl (\_ n kvs. n + length kvs) 0 a; array_length a > 1 \ \ ahm_invar_aux n a" unfolding ahm_invar_aux_def by blast lemma ahm_invar_distinct_fst_concatD: assumes inv: "ahm_invar_aux n (Array xs)" shows "distinct (map fst (concat xs))" proof - { fix h assume "h < length xs" with inv have "bucket_ok (length xs) h (xs ! h)" "distinct (map fst (xs ! h))" by(simp_all add: ahm_invar_aux_def) } note no_junk = this show ?thesis unfolding map_concat proof(rule distinct_concat') have "distinct [x\xs . x \ []]" unfolding distinct_conv_nth proof(intro allI ballI impI) fix i j assume "i < length [x\xs . x \ []]" "j < length [x\xs . x \ []]" "i \ j" from filter_nth_ex_nth[OF \i < length [x\xs . x \ []]\] obtain i' where "i' \ i" "i' < length xs" and ith: "[x\xs . x \ []] ! i = xs ! i'" and eqi: "[x\take i' xs . x \ []] = take i [x\xs . x \ []]" by blast from filter_nth_ex_nth[OF \j < length [x\xs . x \ []]\] obtain j' where "j' \ j" "j' < length xs" and jth: "[x\xs . x \ []] ! j = xs ! j'" and eqj: "[x\take j' xs . x \ []] = take j [x\xs . x \ []]" by blast show "[x\xs . x \ []] ! i \ [x\xs . x \ []] ! j" proof assume "[x\xs . x \ []] ! i = [x\xs . x \ []] ! j" hence eq: "xs ! i' = xs ! j'" using ith jth by simp from \i < length [x\xs . x \ []]\ have "[x\xs . x \ []] ! i \ set [x\xs . x \ []]" by(rule nth_mem) with ith have "xs ! i' \ []" by simp then obtain kv where "kv \ set (xs ! i')" by(fastforce simp add: neq_Nil_conv) with no_junk[OF \i' < length xs\] have "bounded_hashcode_nat (length xs) (fst kv) = i'" by(simp add: bucket_ok_def) moreover from eq \kv \ set (xs ! i')\ have "kv \ set (xs ! j')" by simp with no_junk[OF \j' < length xs\] have "bounded_hashcode_nat (length xs) (fst kv) = j'" by(simp add: bucket_ok_def) ultimately have [simp]: "i' = j'" by simp from \i < length [x\xs . x \ []]\ have "i = length (take i [x\xs . x \ []])" by simp also from eqi eqj have "take i [x\xs . x \ []] = take j [x\xs . x \ []]" by simp finally show False using \i \ j\ \j < length [x\xs . x \ []]\ by simp qed qed moreover have "inj_on (map fst) {x \ set xs. x \ []}" proof(rule inj_onI) fix x y assume "x \ {x \ set xs. x \ []}" "y \ {x \ set xs. x \ []}" "map fst x = map fst y" hence "x \ set xs" "y \ set xs" "x \ []" "y \ []" by auto from \x \ set xs\ obtain i where "xs ! i = x" "i < length xs" unfolding set_conv_nth by fastforce from \y \ set xs\ obtain j where "xs ! j = y" "j < length xs" unfolding set_conv_nth by fastforce from \x \ []\ obtain k v x' where "x = (k, v) # x'" by(cases x) auto with no_junk[OF \i < length xs\] \xs ! i = x\ have "bounded_hashcode_nat (length xs) k = i" by(auto simp add: bucket_ok_def) moreover from \map fst x = map fst y\ \x = (k, v) # x'\ obtain v' where "(k, v') \ set y" by fastforce with no_junk[OF \j < length xs\] \xs ! j = y\ have "bounded_hashcode_nat (length xs) k = j" by(auto simp add: bucket_ok_def) ultimately have "i = j" by simp with \xs ! i = x\ \xs ! j = y\ show "x = y" by simp qed ultimately show "distinct [ys\map (map fst) xs . ys \ []]" by(simp add: filter_map o_def distinct_map) next fix ys assume "ys \ set (map (map fst) xs)" thus "distinct ys" by(clarsimp simp add: set_conv_nth)(rule no_junk) next fix ys zs assume "ys \ set (map (map fst) xs)" "zs \ set (map (map fst) xs)" "ys \ zs" then obtain ys' zs' where [simp]: "ys = map fst ys'" "zs = map fst zs'" and "ys' \ set xs" "zs' \ set xs" by auto have "fst ` set ys' \ fst ` set zs' = {}" proof(rule equals0I) fix k assume "k \ fst ` set ys' \ fst ` set zs'" then obtain v v' where "(k, v) \ set ys'" "(k, v') \ set zs'" by(auto) from \ys' \ set xs\ obtain i where "xs ! i = ys'" "i < length xs" unfolding set_conv_nth by fastforce with \(k, v) \ set ys'\ have "bounded_hashcode_nat (length xs) k = i" by(auto dest: no_junk bucket_okD) moreover from \zs' \ set xs\ obtain j where "xs ! j = zs'" "j < length xs" unfolding set_conv_nth by fastforce with \(k, v') \ set zs'\ have "bounded_hashcode_nat (length xs) k = j" by(auto dest: no_junk bucket_okD) ultimately have "i = j" by simp with \xs ! i = ys'\ \xs ! j = zs'\ have "ys' = zs'" by simp with \ys \ zs\ show False by simp qed thus "set ys \ set zs = {}" by simp qed qed subsection \@{term "ahm_\"}\ lemma finite_dom_ahm_\_aux: assumes "ahm_invar_aux n a" shows "finite (dom (ahm_\_aux a))" proof - have "dom (ahm_\_aux a) \ (\h \ range (bounded_hashcode_nat (array_length a) :: 'a \ nat). dom (map_of (array_get a h)))" by(force simp add: dom_map_of_conv_image_fst ahm_\_aux_def dest: map_of_SomeD) moreover have "finite \" proof(rule finite_UN_I) from \ahm_invar_aux n a\ have "array_length a > 1" by(simp add: ahm_invar_aux_def) hence "range (bounded_hashcode_nat (array_length a) :: 'a \ nat) \ {0.. nat))" by(rule finite_subset) simp qed(rule finite_dom_map_of) ultimately show ?thesis by(rule finite_subset) qed lemma ahm_\_aux_conv_map_of_concat: assumes inv: "ahm_invar_aux n (Array xs)" shows "ahm_\_aux (Array xs) = map_of (concat xs)" proof fix k show "ahm_\_aux (Array xs) k = map_of (concat xs) k" proof(cases "map_of (concat xs) k") case None hence "k \ fst ` set (concat xs)" by(simp add: map_of_eq_None_iff) hence "k \ fst ` set (xs ! bounded_hashcode_nat (length xs) k)" proof(rule contrapos_nn) assume "k \ fst ` set (xs ! bounded_hashcode_nat (length xs) k)" then obtain v where "(k, v) \ set (xs ! bounded_hashcode_nat (length xs) k)" by auto moreover from inv have "bounded_hashcode_nat (length xs) k < length xs" by(simp add: bounded_hashcode_nat_bounds ahm_invar_aux_def) ultimately show "k \ fst ` set (concat xs)" by(force intro: rev_image_eqI) qed thus ?thesis unfolding None by(simp add: map_of_eq_None_iff) next case (Some v) hence "(k, v) \ set (concat xs)" by(rule map_of_SomeD) then obtain ys where "ys \ set xs" "(k, v) \ set ys" unfolding set_concat by blast from \ys \ set xs\ obtain i j where "i < length xs" "xs ! i = ys" unfolding set_conv_nth by auto with inv \(k, v) \ set ys\ show ?thesis unfolding Some by(force dest: bucket_okD simp add: ahm_invar_aux_def) qed qed lemma ahm_invar_aux_card_dom_ahm_\_auxD: assumes inv: "ahm_invar_aux n a" shows "card (dom (ahm_\_aux a)) = n" proof(cases a) case [simp]: (Array xs) from inv have "card (dom (ahm_\_aux (Array xs))) = card (dom (map_of (concat xs)))" by(simp add: ahm_\_aux_conv_map_of_concat) also from inv have "distinct (map fst (concat xs))" by(simp add: ahm_invar_distinct_fst_concatD) hence "card (dom (map_of (concat xs))) = length (concat xs)" by(rule card_dom_map_of) also have "length (concat xs) = foldl (+) 0 (map length xs)" by (simp add: length_concat foldl_conv_fold add.commute fold_plus_sum_list_rev) also from inv have "\ = n" unfolding foldl_map by(simp add: ahm_invar_aux_def array_foldl_foldl) finally show ?thesis by(simp) qed lemma finite_dom_ahm_\: "ahm_invar hm \ finite (dom (ahm_\ hm))" by(cases hm)(auto intro: finite_dom_ahm_\_aux) lemma finite_map_ahm_\_aux: "finite_map ahm_\_aux (ahm_invar_aux n)" by(unfold_locales)(rule finite_dom_ahm_\_aux) lemma finite_map_ahm_\: "finite_map ahm_\ ahm_invar" by(unfold_locales)(rule finite_dom_ahm_\) subsection \@{term ahm_empty}\ lemma ahm_invar_aux_new_array: assumes "n > 1" shows "ahm_invar_aux 0 (new_array [] n)" proof - have "foldl (\b (k, v). b + length v) 0 (zip [0.. 1 \ ahm_invar (new_hashmap_with n)" by(auto simp add: ahm_invar_def new_hashmap_with_def intro: ahm_invar_aux_new_array) lemma ahm_\_new_hashmap_with: "n > 1 \ ahm_\ (new_hashmap_with n) = Map.empty" by(simp add: new_hashmap_with_def bounded_hashcode_nat_bounds fun_eq_iff) lemma ahm_invar_ahm_empty [simp]: "ahm_invar (ahm_empty ())" using def_hashmap_size[where ?'a = 'a] by(auto intro: ahm_invar_new_hashmap_with simp add: ahm_empty_def) lemma ahm_empty_correct [simp]: "ahm_\ (ahm_empty ()) = Map.empty" using def_hashmap_size[where ?'a = 'a] by(auto intro: ahm_\_new_hashmap_with simp add: ahm_empty_def) lemma ahm_empty_impl: "map_empty ahm_\ ahm_invar ahm_empty" by(unfold_locales)(auto) subsection \@{term "ahm_lookup"}\ lemma ahm_lookup_impl: "map_lookup ahm_\ ahm_invar ahm_lookup" by(unfold_locales)(simp add: ahm_lookup_def) subsection \@{term "ahm_iteratei"}\ lemma ahm_iteratei_aux_impl: assumes invar_m: "ahm_invar_aux n m" shows "map_iterator (ahm_iteratei_aux m) (ahm_\_aux m)" proof - obtain ms where m_eq[simp]: "m = Array ms" by (cases m) from ahm_invar_distinct_fst_concatD[of n ms] invar_m have dist: "distinct (map fst (concat ms))" by simp show "map_iterator (ahm_iteratei_aux m) (ahm_\_aux m)" using set_iterator_foldli_correct[of "concat ms"] dist by (simp add: ahm_\_aux_conv_map_of_concat[OF invar_m[unfolded m_eq]] ahm_iteratei_aux_def map_to_set_map_of[OF dist] distinct_map) qed lemma ahm_iteratei_correct: assumes invar_hm: "ahm_invar hm" shows "map_iterator (ahm_iteratei hm) (ahm_\ hm)" proof - obtain A n where hm_eq [simp]: "hm = HashMap A n" by(cases hm) from ahm_iteratei_aux_impl[of n A] invar_hm show map_it: "map_iterator (ahm_iteratei hm) (ahm_\ hm)" by simp qed lemma ahm_iteratei_aux_code [code]: "ahm_iteratei_aux a c f \ = a_idx_it.idx_iteratei a c (\x. foldli x c f) \" proof(cases a) case [simp]: (Array xs) have "ahm_iteratei_aux a c f \ = foldli (concat xs) c f \" by simp also have "\ = foldli xs c (\x. foldli x c f) \" by (simp add: foldli_concat) thm a_idx_it.idx_iteratei_correct also have "\ = a_idx_it.idx_iteratei a c (\x. foldli x c f) \" by (simp add: a_idx_it.idx_iteratei_correct) finally show ?thesis . qed subsection \@{term "ahm_rehash"}\ lemma array_length_ahm_rehash_aux': "array_length (ahm_rehash_aux' n kv a) = array_length a" by(simp add: ahm_rehash_aux'_def Let_def) lemma ahm_rehash_aux'_preserves_ahm_invar_aux: assumes inv: "ahm_invar_aux n a" and fresh: "k \ fst ` set (array_get a (bounded_hashcode_nat (array_length a) k))" shows "ahm_invar_aux (Suc n) (ahm_rehash_aux' (array_length a) (k, v) a)" (is "ahm_invar_aux _ ?a") proof(rule ahm_invar_auxI) fix h assume "h < array_length ?a" hence hlen: "h < array_length a" by(simp add: array_length_ahm_rehash_aux') with inv have bucket: "bucket_ok (array_length a) h (array_get a h)" and dist: "distinct (map fst (array_get a h))" by(auto elim: ahm_invar_auxE) let ?h = "bounded_hashcode_nat (array_length a) k" from hlen bucket show "bucket_ok (array_length ?a) h (array_get ?a h)" by(cases "h = ?h")(auto simp add: ahm_rehash_aux'_def Let_def array_length_ahm_rehash_aux' array_get_array_set_other dest: bucket_okD intro!: bucket_okI) from dist hlen fresh show "distinct (map fst (array_get ?a h))" by(cases "h = ?h")(auto simp add: ahm_rehash_aux'_def Let_def array_get_array_set_other) next let ?f = "\n kvs. n + length kvs" { fix n :: nat and xs :: "('a \ 'b) list list" have "foldl ?f n xs = n + foldl ?f 0 xs" by(induct xs arbitrary: rule: rev_induct) simp_all } note fold = this let ?h = "bounded_hashcode_nat (array_length a) k" obtain xs where a [simp]: "a = Array xs" by(cases a) from inv have [simp]: "bounded_hashcode_nat (length xs) k < length xs" by(simp add: ahm_invar_aux_def bounded_hashcode_nat_bounds) have xs: "xs = take ?h xs @ (xs ! ?h) # drop (Suc ?h) xs" by(simp add: Cons_nth_drop_Suc) from inv have "n = array_foldl (\_ n kvs. n + length kvs) 0 a" by(auto elim: ahm_invar_auxE) hence "n = foldl ?f 0 (take ?h xs) + length (xs ! ?h) + foldl ?f 0 (drop (Suc ?h) xs)" by(simp add: array_foldl_foldl)(subst xs, simp, subst (1 2 3 4) fold, simp) thus "Suc n = array_foldl (\_ n kvs. n + length kvs) 0 ?a" by(simp add: ahm_rehash_aux'_def Let_def array_foldl_foldl foldl_list_update)(subst (1 2 3 4) fold, simp) next from inv have "1 < array_length a" by(auto elim: ahm_invar_auxE) thus "1 < array_length ?a" by(simp add: array_length_ahm_rehash_aux') qed declare [[coercion_enabled = false]] lemma ahm_rehash_aux_correct: fixes a :: "(('key :: hashable) \ 'val) list array" assumes inv: "ahm_invar_aux n a" and "sz > 1" shows "ahm_invar_aux n (ahm_rehash_aux a sz)" (is "?thesis1") and "ahm_\_aux (ahm_rehash_aux a sz) = ahm_\_aux a" (is "?thesis2") proof - (*interpret ahm: map_iterator "ahm_\_aux" "ahm_invar_aux n" "ahm_iteratei_aux" by(rule ahm_iteratei_aux_impl)*) let ?a = "ahm_rehash_aux a sz" let ?I = "\it a'. ahm_invar_aux (n - card it) a' \ array_length a' = sz \ (\k. if k \ it then ahm_\_aux a' k = None else ahm_\_aux a' k = ahm_\_aux a k)" have "?I {} ?a \ (\it\dom (ahm_\_aux a). it \ {} \ \ True \ ?I it ?a)" unfolding ahm_rehash_aux_def proof (rule map_iterator_rule_P[OF ahm_iteratei_aux_impl[OF inv], where c = "\_. True" and f="ahm_rehash_aux' sz" and ?\0.0 = "new_array [] sz" and I="?I"] ) from inv have "card (dom (ahm_\_aux a)) = n" by(rule ahm_invar_aux_card_dom_ahm_\_auxD) moreover from \1 < sz\ have "ahm_invar_aux 0 (new_array ([] :: ('key \ 'val) list) sz)" by(rule ahm_invar_aux_new_array) moreover { fix k assume "k \ dom (ahm_\_aux a)" hence "ahm_\_aux a k = None" by auto moreover have "bounded_hashcode_nat sz k < sz" using \1 < sz\ by(simp add: bounded_hashcode_nat_bounds) ultimately have "ahm_\_aux (new_array [] sz) k = ahm_\_aux a k" by simp } ultimately show "?I (dom (ahm_\_aux a)) (new_array [] sz)" by(auto simp add: bounded_hashcode_nat_bounds[OF \1 < sz\]) next fix k :: 'key and v :: 'val and it a' assume "k \ it" "ahm_\_aux a k = Some v" and it_sub: "it \ dom (ahm_\_aux a)" and I: "?I it a'" from I have inv': "ahm_invar_aux (n - card it) a'" and a'_eq_a: "\k. k \ it \ ahm_\_aux a' k = ahm_\_aux a k" and a'_None: "\k. k \ it \ ahm_\_aux a' k = None" and [simp]: "sz = array_length a'" by(auto split: if_split_asm) from it_sub finite_dom_ahm_\_aux[OF inv] have "finite it" by(rule finite_subset) moreover with \k \ it\ have "card it > 0" by(auto simp add: card_gt_0_iff) moreover from finite_dom_ahm_\_aux[OF inv] it_sub have "card it \ card (dom (ahm_\_aux a))" by(rule card_mono) moreover have "\ = n" using inv by(simp add: ahm_invar_aux_card_dom_ahm_\_auxD) ultimately have "n - card (it - {k}) = (n - card it) + 1" using \k \ it\ by auto moreover from \k \ it\ have "ahm_\_aux a' k = None" by(rule a'_None) hence "k \ fst ` set (array_get a' (bounded_hashcode_nat (array_length a') k))" by(simp add: map_of_eq_None_iff) ultimately have "ahm_invar_aux (n - card (it - {k})) (ahm_rehash_aux' sz (k, v) a')" using inv' by(auto intro: ahm_rehash_aux'_preserves_ahm_invar_aux) moreover have "array_length (ahm_rehash_aux' sz (k, v) a') = sz" by(simp add: array_length_ahm_rehash_aux') moreover { fix k' assume "k' \ it - {k}" with bounded_hashcode_nat_bounds[OF \1 < sz\, of k'] a'_None[of k'] have "ahm_\_aux (ahm_rehash_aux' sz (k, v) a') k' = None" by(cases "bounded_hashcode_nat sz k = bounded_hashcode_nat sz k'")(auto simp add: array_get_array_set_other ahm_rehash_aux'_def Let_def) } moreover { fix k' assume "k' \ it - {k}" with bounded_hashcode_nat_bounds[OF \1 < sz\, of k'] bounded_hashcode_nat_bounds[OF \1 < sz\, of k] a'_eq_a[of k'] \k \ it\ have "ahm_\_aux (ahm_rehash_aux' sz (k, v) a') k' = ahm_\_aux a k'" unfolding ahm_rehash_aux'_def Let_def using \ahm_\_aux a k = Some v\ by(cases "bounded_hashcode_nat sz k = bounded_hashcode_nat sz k'")(case_tac [!] "k' = k", simp_all add: array_get_array_set_other) } ultimately show "?I (it - {k}) (ahm_rehash_aux' sz (k, v) a')" by simp qed auto thus ?thesis1 ?thesis2 unfolding ahm_rehash_aux_def by(auto intro: ext) qed lemma ahm_rehash_correct: fixes hm :: "('key :: hashable, 'val) hashmap" assumes inv: "ahm_invar hm" and "sz > 1" shows "ahm_invar (ahm_rehash hm sz)" "ahm_\ (ahm_rehash hm sz) = ahm_\ hm" using assms by -(case_tac [!] hm, auto intro: ahm_rehash_aux_correct) subsection \@{term ahm_update}\ lemma ahm_update_aux_correct: assumes inv: "ahm_invar hm" shows "ahm_invar (ahm_update_aux hm k v)" (is ?thesis1) and "ahm_\ (ahm_update_aux hm k v) = (ahm_\ hm)(k \ v)" (is ?thesis2) proof - obtain a n where [simp]: "hm = HashMap a n" by(cases hm) let ?h = "bounded_hashcode_nat (array_length a) k" let ?a' = "array_set a ?h (AList.update k v (array_get a ?h))" let ?n' = "if map_of (array_get a ?h) k = None then n + 1 else n" have "ahm_invar (HashMap ?a' ?n')" unfolding ahm_invar.simps proof(rule ahm_invar_auxI) fix h assume "h < array_length ?a'" hence "h < array_length a" by simp with inv have "bucket_ok (array_length a) h (array_get a h)" by(auto elim: ahm_invar_auxE) thus "bucket_ok (array_length ?a') h (array_get ?a' h)" using \h < array_length a\ apply(cases "h = bounded_hashcode_nat (array_length a) k") apply(fastforce intro!: bucket_okI simp add: dom_update array_get_array_set_other dest: bucket_okD del: imageE elim: imageE)+ done from \h < array_length a\ inv have "distinct (map fst (array_get a h))" by(auto elim: ahm_invar_auxE) with \h < array_length a\ show "distinct (map fst (array_get ?a' h))" by(cases "h = bounded_hashcode_nat (array_length a) k")(auto simp add: array_get_array_set_other intro: distinct_update) next obtain xs where a [simp]: "a = Array xs" by(cases a) let ?f = "\n kvs. n + length kvs" { fix n :: nat and xs :: "('a \ 'b) list list" have "foldl ?f n xs = n + foldl ?f 0 xs" by(induct xs arbitrary: rule: rev_induct) simp_all } note fold = this from inv have [simp]: "bounded_hashcode_nat (length xs) k < length xs" by(simp add: ahm_invar_aux_def bounded_hashcode_nat_bounds) have xs: "xs = take ?h xs @ (xs ! ?h) # drop (Suc ?h) xs" by(simp add: Cons_nth_drop_Suc) from inv have "n = array_foldl (\_ n kvs. n + length kvs) 0 a" by(auto elim: ahm_invar_auxE) hence "n = foldl ?f 0 (take ?h xs) + length (xs ! ?h) + foldl ?f 0 (drop (Suc ?h) xs)" by(simp add: array_foldl_foldl)(subst xs, simp, subst (1 2 3 4) fold, simp) thus "?n' = array_foldl (\_ n kvs. n + length kvs) 0 ?a'" apply(simp add: ahm_rehash_aux'_def Let_def array_foldl_foldl foldl_list_update map_of_eq_None_iff) apply(subst (1 2 3 4 5 6 7 8) fold) apply(simp add: length_update) done next from inv have "1 < array_length a" by(auto elim: ahm_invar_auxE) thus "1 < array_length ?a'" by simp qed - moreover have "ahm_\ (ahm_update_aux hm k v) = ahm_\ hm(k \ v)" + moreover have "ahm_\ (ahm_update_aux hm k v) = (ahm_\ hm)(k \ v)" proof fix k' from inv have "1 < array_length a" by(auto elim: ahm_invar_auxE) with bounded_hashcode_nat_bounds[OF this, of k] - show "ahm_\ (ahm_update_aux hm k v) k' = (ahm_\ hm(k \ v)) k'" + show "ahm_\ (ahm_update_aux hm k v) k' = ((ahm_\ hm)(k \ v)) k'" by(cases "bounded_hashcode_nat (array_length a) k = bounded_hashcode_nat (array_length a) k'")(auto simp add: Let_def update_conv array_get_array_set_other) qed ultimately show ?thesis1 ?thesis2 by(simp_all add: Let_def) qed lemma ahm_update_correct: assumes inv: "ahm_invar hm" shows "ahm_invar (ahm_update k v hm)" and "ahm_\ (ahm_update k v hm) = (ahm_\ hm)(k \ v)" using assms by(simp_all add: ahm_update_def Let_def ahm_rehash_correct ahm_update_aux_correct) lemma ahm_update_impl: "map_update ahm_\ ahm_invar ahm_update" by(unfold_locales)(simp_all add: ahm_update_correct) subsection \@{term "ahm_delete"}\ lemma ahm_delete_correct: assumes inv: "ahm_invar hm" shows "ahm_invar (ahm_delete k hm)" (is "?thesis1") and "ahm_\ (ahm_delete k hm) = (ahm_\ hm) |` (- {k})" (is "?thesis2") proof - obtain a n where hm [simp]: "hm = HashMap a n" by(cases hm) let ?h = "bounded_hashcode_nat (array_length a) k" let ?a' = "array_set a ?h (AList.delete k (array_get a ?h))" let ?n' = "if map_of (array_get a (bounded_hashcode_nat (array_length a) k)) k = None then n else n - 1" have "ahm_invar_aux ?n' ?a'" proof(rule ahm_invar_auxI) fix h assume "h < array_length ?a'" hence "h < array_length a" by simp with inv have "bucket_ok (array_length a) h (array_get a h)" and "1 < array_length a" and "distinct (map fst (array_get a h))" by(auto elim: ahm_invar_auxE) thus "bucket_ok (array_length ?a') h (array_get ?a' h)" and "distinct (map fst (array_get ?a' h))" using bounded_hashcode_nat_bounds[of "array_length a" k] by-(case_tac [!] "h = bounded_hashcode_nat (array_length a) k", auto simp add: array_get_array_set_other set_delete_conv intro!: bucket_okI dest: bucket_okD intro: distinct_delete) next obtain xs where a [simp]: "a = Array xs" by(cases a) let ?f = "\n kvs. n + length kvs" { fix n :: nat and xs :: "('a \ 'b) list list" have "foldl ?f n xs = n + foldl ?f 0 xs" by(induct xs arbitrary: rule: rev_induct) simp_all } note fold = this from inv have [simp]: "bounded_hashcode_nat (length xs) k < length xs" by(simp add: ahm_invar_aux_def bounded_hashcode_nat_bounds) from inv have "distinct (map fst (array_get a ?h))" by(auto elim: ahm_invar_auxE) moreover have xs: "xs = take ?h xs @ (xs ! ?h) # drop (Suc ?h) xs" by(simp add: Cons_nth_drop_Suc) from inv have "n = array_foldl (\_ n kvs. n + length kvs) 0 a" by(auto elim: ahm_invar_auxE) hence "n = foldl ?f 0 (take ?h xs) + length (xs ! ?h) + foldl ?f 0 (drop (Suc ?h) xs)" by(simp add: array_foldl_foldl)(subst xs, simp, subst (1 2 3 4) fold, simp) ultimately show "?n' = array_foldl (\_ n kvs. n + length kvs) 0 ?a'" apply(simp add: array_foldl_foldl foldl_list_update map_of_eq_None_iff) apply(subst (1 2 3 4 5 6 7 8) fold) apply(auto simp add: length_distinct in_set_conv_nth) done next from inv show "1 < array_length ?a'" by(auto elim: ahm_invar_auxE) qed thus "?thesis1" by(auto simp add: Let_def) have "ahm_\_aux ?a' = ahm_\_aux a |` (- {k})" proof fix k' :: 'a from inv have "bounded_hashcode_nat (array_length a) k < array_length a" by(auto elim: ahm_invar_auxE simp add: bounded_hashcode_nat_bounds) thus "ahm_\_aux ?a' k' = (ahm_\_aux a |` (- {k})) k'" by(cases "?h = bounded_hashcode_nat (array_length a) k'")(auto simp add: restrict_map_def array_get_array_set_other delete_conv) qed thus ?thesis2 by(simp add: Let_def) qed lemma ahm_delete_impl: "map_delete ahm_\ ahm_invar ahm_delete" by(unfold_locales)(blast intro: ahm_delete_correct)+ hide_const (open) HashMap ahm_empty bucket_ok ahm_invar ahm_\ ahm_lookup ahm_iteratei ahm_rehash hm_grow ahm_filled ahm_update ahm_delete hide_type (open) hashmap end diff --git a/thys/Collections/ICF/impl/HashMap_Impl.thy b/thys/Collections/ICF/impl/HashMap_Impl.thy --- a/thys/Collections/ICF/impl/HashMap_Impl.thy +++ b/thys/Collections/ICF/impl/HashMap_Impl.thy @@ -1,455 +1,455 @@ (* Title: Isabelle Collections Library Author: Peter Lammich Maintainer: Peter Lammich *) section \\isaheader{Hash maps implementation}\ theory HashMap_Impl imports RBTMapImpl ListMapImpl "../../Lib/HashCode" "../../Lib/Code_Target_ICF" begin text \ We use a red-black tree instead of an indexed array. This has the disadvantage of being more complex, however we need not bother about a fixed-size array and rehashing if the array becomes too full. The entries of the red-black tree are lists of (key,value) pairs. \ subsection \Abstract Hashmap\ text \ We first specify the behavior of our hashmap on the level of maps. We will then show that our implementation based on hashcode-map and bucket-map is a correct implementation of this specification. \ type_synonym ('k,'v) abs_hashmap = "hashcode \ ('k \ 'v)" \ \Map entry of map by function\ abbreviation map_entry where "map_entry k f m == m(k := f (m k))" \ \Invariant: Buckets only contain entries with the right hashcode and there are no empty buckets\ definition ahm_invar:: "('k::hashable,'v) abs_hashmap \ bool" where "ahm_invar m == (\hc cm k. m hc = Some cm \ k\dom cm \ hashcode k = hc) \ (\hc cm. m hc = Some cm \ cm \ Map.empty)" \ \Abstract a hashmap to the corresponding map\ definition ahm_\ where "ahm_\ m k == case m (hashcode k) of None \ None | Some cm \ cm k" \ \Lookup an entry\ definition ahm_lookup :: "'k::hashable \ ('k,'v) abs_hashmap \ 'v option" where "ahm_lookup k m == (ahm_\ m) k" \ \The empty hashmap\ definition ahm_empty :: "('k::hashable,'v) abs_hashmap" where "ahm_empty = Map.empty" \ \Update/insert an entry\ definition ahm_update where "ahm_update k v m == case m (hashcode k) of None \ m (hashcode k \ [k \ v]) | Some cm \ m (hashcode k \ cm (k \ v)) " \ \Delete an entry\ definition ahm_delete where "ahm_delete k m == map_entry (hashcode k) (\v. case v of None \ None | Some bm \ ( if bm |` (- {k}) = Map.empty then None else Some ( bm |` (- {k})) ) ) m " definition ahm_isEmpty where "ahm_isEmpty m == m=Map.empty" text \ Now follow correctness lemmas, that relate the hashmap operations to operations on the corresponding map. Those lemmas are named op\_correct, where (is) the operation. \ lemma ahm_invarI: "\ !!hc cm k. \m hc = Some cm; k\dom cm\ \ hashcode k = hc; !!hc cm. \ m hc = Some cm \ \ cm \ Map.empty \ \ ahm_invar m" by (unfold ahm_invar_def) blast lemma ahm_invarD: "\ ahm_invar m; m hc = Some cm; k\dom cm \ \ hashcode k = hc" by (unfold ahm_invar_def) blast lemma ahm_invarDne: "\ ahm_invar m; m hc = Some cm \ \ cm \ Map.empty" by (unfold ahm_invar_def) blast lemma ahm_invar_bucket_not_empty[simp]: "ahm_invar m \ m hc \ Some Map.empty" by (auto dest: ahm_invarDne) lemmas ahm_lookup_correct = ahm_lookup_def lemma ahm_empty_correct: "ahm_\ ahm_empty = Map.empty" "ahm_invar ahm_empty" apply (rule ext) apply (unfold ahm_empty_def) apply (auto simp add: ahm_\_def intro: ahm_invarI split: option.split) done lemma ahm_update_correct: - "ahm_\ (ahm_update k v m) = ahm_\ m (k \ v)" + "ahm_\ (ahm_update k v m) = (ahm_\ m)(k \ v)" "ahm_invar m \ ahm_invar (ahm_update k v m)" apply (rule ext) apply (unfold ahm_update_def) apply (auto simp add: ahm_\_def split: option.split) apply (rule ahm_invarI) apply (auto dest: ahm_invarD ahm_invarDne split: if_split_asm) apply (rule ahm_invarI) apply (auto dest: ahm_invarD split: if_split_asm) apply (drule (1) ahm_invarD) apply auto done lemma fun_upd_apply_ne: "x\y \ (f(x:=v)) y = f y" by simp lemma cancel_one_empty_simp: "m |` (-{k}) = Map.empty \ dom m \ {k}" proof assume "m |` (- {k}) = Map.empty" hence "{} = dom (m |` (-{k}))" by auto also have "\ = dom m - {k}" by auto finally show "dom m \ {k}" by blast next assume "dom m \ {k}" hence "dom m - {k} = {}" by auto hence "dom (m |` (-{k})) = {}" by auto thus "m |` (-{k}) = Map.empty" by blast qed lemma ahm_delete_correct: "ahm_\ (ahm_delete k m) = (ahm_\ m) |` (-{k})" "ahm_invar m \ ahm_invar (ahm_delete k m)" apply (rule ext) apply (unfold ahm_delete_def) apply (auto simp add: ahm_\_def Let_def Map.restrict_map_def split: option.split)[1] apply (drule_tac x=x in fun_cong) apply (auto)[1] apply (rule ahm_invarI) apply (auto split: if_split_asm option.split_asm dest: ahm_invarD) apply (drule (1) ahm_invarD) apply (auto simp add: restrict_map_def split: if_split_asm option.split_asm) done lemma ahm_isEmpty_correct: "ahm_invar m \ ahm_isEmpty m \ ahm_\ m = Map.empty" proof assume "ahm_invar m" "ahm_isEmpty m" thus "ahm_\ m = Map.empty" by (auto simp add: ahm_isEmpty_def ahm_\_def intro: ext) next assume I: "ahm_invar m" and E: "ahm_\ m = Map.empty" show "ahm_isEmpty m" proof (rule ccontr) assume "\ahm_isEmpty m" then obtain hc bm where MHC: "m hc = Some bm" by (unfold ahm_isEmpty_def) (blast elim: nempty_dom dest: domD) from ahm_invarDne[OF I, OF MHC] obtain k v where BMK: "bm k = Some v" by (blast elim: nempty_dom dest: domD) from ahm_invarD[OF I, OF MHC] BMK have [simp]: "hashcode k = hc" by auto hence "ahm_\ m k = Some v" by (simp add: ahm_\_def MHC BMK) with E show False by simp qed qed lemmas ahm_correct = ahm_empty_correct ahm_lookup_correct ahm_update_correct ahm_delete_correct ahm_isEmpty_correct \ \Bucket entries correspond to map entries\ lemma ahm_be_is_e: assumes I: "ahm_invar m" assumes A: "m hc = Some bm" "bm k = Some v" shows "ahm_\ m k = Some v" using A apply (auto simp add: ahm_\_def split: option.split dest: ahm_invarD[OF I]) apply (frule ahm_invarD[OF I, where k=k]) apply auto done \ \Map entries correspond to bucket entries\ lemma ahm_e_is_be: "\ ahm_\ m k = Some v; !!bm. \m (hashcode k) = Some bm; bm k = Some v \ \ P \ \ P" by (unfold ahm_\_def) (auto split: option.split_asm) subsection \Concrete Hashmap\ text \ In this section, we define the concrete hashmap that is made from the hashcode map and the bucket map. We then show the correctness of the operations w.r.t. the abstract hashmap, and thus, indirectly, w.r.t. the corresponding map. \ type_synonym ('k,'v) hm_impl = "(hashcode, ('k,'v) lm) rm" subsubsection "Operations" \ \Auxiliary function: Apply function to value of an entry\ definition rm_map_entry :: "hashcode \ ('v option \ 'v option) \ (hashcode, 'v) rm \ (hashcode,'v) rm" where "rm_map_entry k f m == case rm.lookup k m of None \ ( case f None of None \ m | Some v \ rm.update k v m ) | Some v \ ( case f (Some v) of None \ rm.delete k m | Some v' \ rm.update k v' m ) " \ \Empty hashmap\ definition empty :: "unit \ ('k :: hashable, 'v) hm_impl" where "empty == rm.empty" \ \Update/insert entry\ definition update :: "'k::hashable \ 'v \ ('k,'v) hm_impl \ ('k,'v) hm_impl" where "update k v m == let hc = hashcode k in case rm.lookup hc m of None \ rm.update hc (lm.update k v (lm.empty ())) m | Some bm \ rm.update hc (lm.update k v bm) m" \ \Lookup value by key\ definition lookup :: "'k::hashable \ ('k,'v) hm_impl \ 'v option" where "lookup k m == case rm.lookup (hashcode k) m of None \ None | Some lm \ lm.lookup k lm" \ \Delete entry by key\ definition delete :: "'k::hashable \ ('k,'v) hm_impl \ ('k,'v) hm_impl" where "delete k m == rm_map_entry (hashcode k) (\v. case v of None \ None | Some lm \ ( let lm' = lm.delete k lm in if lm.isEmpty lm' then None else Some lm' ) ) m" \ \Emptiness check\ definition "isEmpty == rm.isEmpty" \ \Interruptible iterator\ definition "iteratei m c f \0 == rm.iteratei m c (\(hc, lm) \. lm.iteratei lm c f \ ) \0" lemma iteratei_alt_def : "iteratei m = set_iterator_image snd ( set_iterator_product (rm.iteratei m) (\hclm. lm.iteratei (snd hclm)))" proof - have aux: "\c f. (\(hc, lm). lm.iteratei lm c f) = (\m. lm.iteratei (snd m) c f)" by auto show ?thesis unfolding set_iterator_product_def set_iterator_image_alt_def iteratei_def[abs_def] aux by (simp add: split_beta) qed subsubsection "Correctness w.r.t. Abstract HashMap" text \ The following lemmas establish the correctness of the operations w.r.t. the abstract hashmap. They have the naming scheme op\_correct', where (is) the name of the operation. \ \ \Abstract concrete hashmap to abstract hashmap\ definition hm_\' where "hm_\' m == \hc. case rm.\ m hc of None \ None | Some lm \ Some (lm.\ lm)" \ \Invariant for concrete hashmap: The hashcode-map and bucket-maps satisfy their invariants and the invariant of the corresponding abstract hashmap is satisfied.\ definition "invar m == ahm_invar (hm_\' m)" lemma rm_map_entry_correct: "rm.\ (rm_map_entry k f m) = (rm.\ m)(k := f (rm.\ m k))" apply (auto simp add: rm_map_entry_def rm.delete_correct rm.lookup_correct rm.update_correct split: option.split) done lemma empty_correct': "hm_\' (empty ()) = ahm_empty" "invar (empty ())" by (simp_all add: hm_\'_def empty_def ahm_empty_def rm.correct invar_def ahm_invar_def) lemma lookup_correct': "invar m \ lookup k m = ahm_lookup k (hm_\' m)" apply (unfold lookup_def invar_def) apply (auto split: option.split simp add: ahm_lookup_def ahm_\_def hm_\'_def rm.correct lm.correct) done lemma update_correct': "invar m \ hm_\' (update k v m) = ahm_update k v (hm_\' m)" "invar m \ invar (update k v m)" proof - assume "invar m" thus "hm_\' (update k v m) = ahm_update k v (hm_\' m)" apply (unfold invar_def) apply (rule ext) apply (auto simp add: update_def ahm_update_def hm_\'_def Let_def rm.correct lm.correct split: option.split) done thus "invar m \ invar (update k v m)" by (simp add: invar_def ahm_update_correct) qed lemma delete_correct': "invar m \ hm_\' (delete k m) = ahm_delete k (hm_\' m)" "invar m \ invar (delete k m)" proof - assume "invar m" thus "hm_\' (delete k m) = ahm_delete k (hm_\' m)" apply (unfold invar_def) apply (rule ext) apply (auto simp add: delete_def ahm_delete_def hm_\'_def rm_map_entry_correct rm.correct lm.correct Let_def split: option.split option.split_asm) done thus "invar (delete k m)" using \invar m\ by (simp add: ahm_delete_correct invar_def) qed lemma isEmpty_correct': "invar hm \ isEmpty hm \ ahm_\ (hm_\' hm) = Map.empty" apply (simp add: isEmpty_def rm.isEmpty_correct invar_def ahm_isEmpty_correct[unfolded ahm_isEmpty_def, symmetric]) by (auto simp add: hm_\'_def ahm_\_def fun_eq_iff split: option.split_asm) (* lemma sel_correct': assumes "invar hm" shows "\ sel hm f = Some r; \u v. \ ahm_\ (hm_\' hm) u = Some v; f (u, v) = Some r \ \ P \ \ P" and "\ sel hm f = None; ahm_\ (hm_\' hm) u = Some v \ \ f (u, v) = None" proof - assume sel: "sel hm f = Some r" and P: "\u v. \ahm_\ (hm_\' hm) u = Some v; f (u, v) = Some r\ \ P" from `invar hm` have IA: "ahm_invar (hm_\' hm)" by(simp add: invar_def) from TrueI sel obtain hc lm where "rm_\ hm hc = Some lm" and "lm_sel lm f = Some r" unfolding sel_def by (rule rm.sel_someE) simp from TrueI `lm_sel lm f = Some r` obtain k v where "lm_\ lm k = Some v" "f (k, v) = Some r" by(rule lm.sel_someE) from `rm_\ hm hc = Some lm` have "hm_\' hm hc = Some (lm_\ lm)" by(simp add: hm_\'_def) with IA have "ahm_\ (hm_\' hm) k = Some v" using `lm_\ lm k = Some v` by(rule ahm_be_is_e) thus P using `f (k, v) = Some r` by(rule P) next assume sel: "sel hm f = None" and \: "ahm_\ (hm_\' hm) u = Some v" from `invar hm` have IA: "ahm_invar (hm_\' hm)" by(simp add: invar_def) from \ obtain lm where \_u: "hm_\' hm (hashcode u) = Some lm" and "lm u = Some v" by (rule ahm_e_is_be) from \_u obtain lmc where "rm_\ hm (hashcode u) = Some lmc" "lm = lm_\ lmc" by(auto simp add: hm_\'_def split: option.split_asm) with `lm u = Some v` have "lm_\ lmc u = Some v" by simp from sel rm.sel_noneD [OF TrueI _ `rm_\ hm (hashcode u) = Some lmc`, of "(\(hc, lm). lm_sel lm f)"] have "lm_sel lmc f = None" unfolding sel_def by simp with TrueI show "f (u, v) = None" using `lm_\ lmc u = Some v` by(rule lm.sel_noneD) qed *) lemma iteratei_correct': assumes invar: "invar hm" shows "map_iterator (iteratei hm) (ahm_\ (hm_\' hm))" proof - from rm.iteratei_correct have it_rm: "map_iterator (rm.iteratei hm) (rm.\ hm)" by simp from lm.iteratei_correct have it_lm: "\lm. map_iterator (lm.iteratei lm) (lm.\ lm)" by simp from set_iterator_product_correct [OF it_rm, of "\hclm. lm.iteratei (snd hclm)" "\hclm. map_to_set (lm.\ (snd hclm))", OF it_lm] have it_prod: "set_iterator (set_iterator_product (rm.iteratei hm) (\hclm. lm.iteratei (snd hclm))) (SIGMA hclm:map_to_set (rm.\ hm). map_to_set (lm.\ (snd hclm)))" by simp show ?thesis unfolding iteratei_alt_def proof (rule set_iterator_image_correct[OF it_prod]) from invar show "inj_on snd (SIGMA hclm:map_to_set (rm.\ hm). map_to_set (lm.\ (snd hclm)))" apply (simp add: inj_on_def invar_def ahm_invar_def hm_\'_def Ball_def map_to_set_def split: option.splits) apply (metis domI option.inject) done next from invar show "map_to_set (ahm_\ (hm_\' hm)) = snd ` (SIGMA hclm:map_to_set (rm.\ hm). map_to_set (lm.\ (snd hclm)))" apply (simp add: inj_on_def invar_def ahm_invar_def hm_\'_def Ball_def map_to_set_def set_eq_iff image_iff split: option.splits) apply (auto simp add: dom_def ahm_\_def split: option.splits) done qed qed lemmas hm_correct' = empty_correct' lookup_correct' update_correct' delete_correct' isEmpty_correct' iteratei_correct' lemmas hm_invars = empty_correct'(2) update_correct'(2) delete_correct'(2) hide_const (open) empty invar lookup update delete isEmpty iteratei end diff --git a/thys/Containers/Examples/Map_To_Mapping_Ex.thy b/thys/Containers/Examples/Map_To_Mapping_Ex.thy --- a/thys/Containers/Examples/Map_To_Mapping_Ex.thy +++ b/thys/Containers/Examples/Map_To_Mapping_Ex.thy @@ -1,44 +1,44 @@ (* Title: Containers/Map_To_Mapping_Ex.thy Author: Andreas Lochbihler, ETH Zurich *) theory Map_To_Mapping_Ex imports "../Map_To_Mapping" begin subsection \Test cases for replacing @{typ "'a \ 'b"} with @{typ "('a, 'b) mapping"}\ lift_definition mapping_filter :: "('a, 'b) mapping \ 'a list \ 'b list" is "List.map_filter" . lemmas mapping_filter_code [code] = map_filter_simps[containers_identify] definition test :: "(nat \ int option) \ nat list \ int list" where "test f xs = - (if f = Map.empty then [] else List.map_filter (f(2 := None)(1 \ -1)) xs)" + (if f = Map.empty then [] else List.map_filter ((f(2 := None))(1 \ -1)) xs)" lift_definition test' :: "(nat, int) mapping \ nat list \ int list" is test . lemmas [code] = test_def[containers_identify] export_code test' checking SML fun iter :: "('a \ 'a option) \ nat \ 'a \ 'a option" where "iter m 0 x = Some x" | "iter m (Suc n) x = Option.bind (m x) (iter m n)" lift_definition iter' :: "('a, 'a) mapping \ nat \ 'a \ 'a option" is iter . lemmas [code] = iter.simps[containers_identify] export_code iter' in SML definition dom_test :: "bool" where "dom_test \ (dom [(1 :: int) \ ([()] :: unit list)] = {1})" lemmas [code] = dom_test_def[containers_identify] ML \val true = @{code dom_test}\ end diff --git a/thys/Jinja/Common/Exceptions.thy b/thys/Jinja/Common/Exceptions.thy --- a/thys/Jinja/Common/Exceptions.thy +++ b/thys/Jinja/Common/Exceptions.thy @@ -1,109 +1,109 @@ (* Title: HOL/MicroJava/J/Exceptions.thy Author: Gerwin Klein, Martin Strecker Copyright 2002 Technische Universitaet Muenchen *) section \Exceptions\ theory Exceptions imports Objects begin definition NullPointer :: cname where "NullPointer \ ''NullPointer''" definition ClassCast :: cname where "ClassCast \ ''ClassCast''" definition OutOfMemory :: cname where "OutOfMemory \ ''OutOfMemory''" definition sys_xcpts :: "cname set" where "sys_xcpts \ {NullPointer, ClassCast, OutOfMemory}" definition addr_of_sys_xcpt :: "cname \ addr" where "addr_of_sys_xcpt s \ if s = NullPointer then 0 else if s = ClassCast then 1 else if s = OutOfMemory then 2 else undefined" definition start_heap :: "'c prog \ heap" where - "start_heap G \ Map.empty (addr_of_sys_xcpt NullPointer \ blank G NullPointer) - (addr_of_sys_xcpt ClassCast \ blank G ClassCast) - (addr_of_sys_xcpt OutOfMemory \ blank G OutOfMemory)" + "start_heap G \ Map.empty (addr_of_sys_xcpt NullPointer \ blank G NullPointer, + addr_of_sys_xcpt ClassCast \ blank G ClassCast, + addr_of_sys_xcpt OutOfMemory \ blank G OutOfMemory)" definition preallocated :: "heap \ bool" where "preallocated h \ \C \ sys_xcpts. \fs. h(addr_of_sys_xcpt C) = Some (C,fs)" subsection "System exceptions" lemma [simp]: "NullPointer \ sys_xcpts \ OutOfMemory \ sys_xcpts \ ClassCast \ sys_xcpts" (*<*)by(simp add: sys_xcpts_def)(*>*) lemma sys_xcpts_cases [consumes 1, cases set]: "\ C \ sys_xcpts; P NullPointer; P OutOfMemory; P ClassCast\ \ P C" (*<*)by (auto simp add: sys_xcpts_def)(*>*) subsection "@{term preallocated}" lemma preallocated_dom [simp]: "\ preallocated h; C \ sys_xcpts \ \ addr_of_sys_xcpt C \ dom h" (*<*)by (fastforce simp:preallocated_def dom_def)(*>*) lemma preallocatedD: "\ preallocated h; C \ sys_xcpts \ \ \fs. h(addr_of_sys_xcpt C) = Some (C, fs)" (*<*)by(auto simp add: preallocated_def sys_xcpts_def)(*>*) lemma preallocatedE [elim?]: "\ preallocated h; C \ sys_xcpts; \fs. h(addr_of_sys_xcpt C) = Some(C,fs) \ P h C\ \ P h C" (*<*)by (fast dest: preallocatedD)(*>*) lemma cname_of_xcp [simp]: "\ preallocated h; C \ sys_xcpts \ \ cname_of h (addr_of_sys_xcpt C) = C" (*<*)by (auto elim: preallocatedE)(*>*) lemma typeof_ClassCast [simp]: "preallocated h \ typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt ClassCast)) = Some(Class ClassCast)" (*<*)by (auto elim: preallocatedE)(*>*) lemma typeof_OutOfMemory [simp]: "preallocated h \ typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt OutOfMemory)) = Some(Class OutOfMemory)" (*<*)by (auto elim: preallocatedE)(*>*) lemma typeof_NullPointer [simp]: "preallocated h \ typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NullPointer)) = Some(Class NullPointer)" (*<*)by (auto elim: preallocatedE)(*>*) lemma preallocated_hext: "\ preallocated h; h \ h' \ \ preallocated h'" (*<*)by (simp add: preallocated_def hext_def)(*>*) (*<*) lemmas preallocated_upd_obj = preallocated_hext [OF _ hext_upd_obj] lemmas preallocated_new = preallocated_hext [OF _ hext_new] (*>*) lemma preallocated_start: "preallocated (start_heap P)" (*<*)by (auto simp add: start_heap_def blank_def sys_xcpts_def fun_upd_apply addr_of_sys_xcpt_def preallocated_def)(*>*) end diff --git a/thys/Jinja/Compiler/Hidden.thy b/thys/Jinja/Compiler/Hidden.thy --- a/thys/Jinja/Compiler/Hidden.thy +++ b/thys/Jinja/Compiler/Hidden.thy @@ -1,51 +1,51 @@ theory Hidden imports "List-Index.List_Index" begin definition hidden :: "'a list \ nat \ bool" where "hidden xs i \ i < size xs \ xs!i \ set(drop (i+1) xs)" lemma hidden_last_index: "x \ set xs \ hidden (xs @ [x]) (last_index xs x)" by(auto simp add: hidden_def nth_append rev_nth[symmetric] dest: last_index_less[OF _ le_refl]) lemma hidden_inacc: "hidden xs i \ last_index xs x \ i" by(auto simp add: hidden_def last_index_drop last_index_less_size_conv) lemma [simp]: "hidden xs i \ hidden (xs@[x]) i" by(auto simp add:hidden_def nth_append) lemma fun_upds_apply: "(m(xs[\]ys)) x = (let xs' = take (size ys) xs in if x \ set xs' then Some(ys ! last_index xs' x) else m x)" proof(induct xs arbitrary: m ys) case Nil then show ?case by(simp add: Let_def) next case Cons show ?case proof(cases ys) case Nil then show ?thesis by(simp add:Let_def) next case Cons': Cons then show ?thesis using Cons by(simp add: Let_def last_index_Cons) qed qed lemma map_upds_apply_eq_Some: "((m(xs[\]ys)) x = Some y) = (let xs' = take (size ys) xs in if x \ set xs' then ys ! last_index xs' x = y else m x = Some y)" by(simp add:fun_upds_apply Let_def) lemma map_upds_upd_conv_last_index: "\x \ set xs; size xs \ size ys \ - \ m(xs[\]ys)(x\y) = m(xs[\]ys[last_index xs x := y])" + \ m(xs[\]ys, x\y) = m(xs[\]ys[last_index xs x := y])" by(rule ext) (simp add:fun_upds_apply eq_sym_conv Let_def) end diff --git a/thys/Jinja/J/Equivalence.thy b/thys/Jinja/J/Equivalence.thy --- a/thys/Jinja/J/Equivalence.thy +++ b/thys/Jinja/J/Equivalence.thy @@ -1,1606 +1,1606 @@ (* Title: Jinja/J/Equivalence.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Equivalence of Big Step and Small Step Semantics\ theory Equivalence imports BigStep SmallStep WWellForm begin subsection\Small steps simulate big step\ subsubsection "Cast" lemma CastReds: "P \ \e,s\ \* \e',s'\ \ P \ \Cast C e,s\ \* \Cast C e',s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) CastRed[OF step(2)]]) qed (*>*) lemma CastRedsNull: "P \ \e,s\ \* \null,s'\ \ P \ \Cast C e,s\ \* \null,s'\" (*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastNull])(*>*) lemma CastRedsAddr: "\ P \ \e,s\ \* \addr a,s'\; hp s' a = Some(D,fs); P \ D \\<^sup>* C \ \ P \ \Cast C e,s\ \* \addr a,s'\" (*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCast])(*>*) lemma CastRedsFail: "\ P \ \e,s\ \* \addr a,s'\; hp s' a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \Cast C e,s\ \* \THROW ClassCast,s'\" (*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastFail])(*>*) lemma CastRedsThrow: "\ P \ \e,s\ \* \throw a,s'\ \ \ P \ \Cast C e,s\ \* \throw a,s'\" (*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CastThrow])(*>*) subsubsection "LAss" lemma LAssReds: "P \ \e,s\ \* \e',s'\ \ P \ \ V:=e,s\ \* \ V:=e',s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) LAssRed[OF step(2)]]) qed (*>*) lemma LAssRedsVal: "\ P \ \e,s\ \* \Val v,(h',l')\ \ \ P \ \ V:=e,s\ \* \unit,(h',l'(V\v))\" (*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ RedLAss])(*>*) lemma LAssRedsThrow: "\ P \ \e,s\ \* \throw a,s'\ \ \ P \ \ V:=e,s\ \* \throw a,s'\" (*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ red_reds.LAssThrow])(*>*) subsubsection "BinOp" lemma BinOp1Reds: "P \ \e,s\ \* \e',s'\ \ P \ \ e \bop\ e\<^sub>2, s\ \* \e' \bop\ e\<^sub>2, s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed1[OF step(2)]]) qed (*>*) lemma BinOp2Reds: "P \ \e,s\ \* \e',s'\ \ P \ \(Val v) \bop\ e, s\ \* \(Val v) \bop\ e', s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed2[OF step(2)]]) qed (*>*) lemma BinOpRedsVal: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \Val v\<^sub>2,s\<^sub>2\" and op: "binop(bop,v\<^sub>1,v\<^sub>2) = Some v" shows "P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \* \Val v,s\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(Val v\<^sub>1 \bop\ e\<^sub>2, s\<^sub>1)" let ?y' = "(Val v\<^sub>1 \bop\ Val v\<^sub>2, s\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" by(rule RedBinOp[OF op]) ultimately show ?thesis by simp qed (*>*) lemma BinOpRedsThrow1: "P \ \e,s\ \* \throw e',s'\ \ P \ \e \bop\ e\<^sub>2, s\ \* \throw e', s'\" (*<*)by(rule BinOp1Reds[THEN rtrancl_into_rtrancl, OF _ red_reds.BinOpThrow1])(*>*) lemma BinOpRedsThrow2: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \throw e,s\<^sub>2\" shows "P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \* \throw e,s\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(Val v\<^sub>1 \bop\ e\<^sub>2, s\<^sub>1)" let ?y' = "(Val v\<^sub>1 \bop\ throw e, s\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" by(rule red_reds.BinOpThrow2) ultimately show ?thesis by simp qed (*>*) subsubsection "FAcc" lemma FAccReds: "P \ \e,s\ \* \e',s'\ \ P \ \e\F{D}, s\ \* \e'\F{D}, s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) FAccRed[OF step(2)]]) qed (*>*) lemma FAccRedsVal: "\P \ \e,s\ \* \addr a,s'\; hp s' a = Some(C,fs); fs(F,D) = Some v \ \ P \ \e\F{D},s\ \* \Val v,s'\" (*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAcc])(*>*) lemma FAccRedsNull: "P \ \e,s\ \* \null,s'\ \ P \ \e\F{D},s\ \* \THROW NullPointer,s'\" (*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccNull])(*>*) lemma FAccRedsThrow: "P \ \e,s\ \* \throw a,s'\ \ P \ \e\F{D},s\ \* \throw a,s'\" (*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ red_reds.FAccThrow])(*>*) subsubsection "FAss" lemma FAssReds1: "P \ \e,s\ \* \e',s'\ \ P \ \e\F{D}:=e\<^sub>2, s\ \* \e'\F{D}:=e\<^sub>2, s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) FAssRed1[OF step(2)]]) qed (*>*) lemma FAssReds2: "P \ \e,s\ \* \e',s'\ \ P \ \Val v\F{D}:=e, s\ \* \Val v\F{D}:=e', s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) FAssRed2[OF step(2)]]) qed (*>*) lemma FAssRedsVal: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \addr a,s\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2)\" and hC: "Some(C,fs) = h\<^sub>2 a" shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0\ \* \unit, (h\<^sub>2(a\(C,fs((F,D)\v))),l\<^sub>2)\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(addr a\F{D}:=e\<^sub>2, s\<^sub>1)" let ?y' = "(addr a\F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2))" have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" using RedFAss hC by simp ultimately show ?thesis by simp qed (*>*) lemma FAssRedsNull: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \null,s\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \Val v,s\<^sub>2\" shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0\ \* \THROW NullPointer, s\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(null\F{D}:=e\<^sub>2, s\<^sub>1)" let ?y' = "(null\F{D}:=Val v::expr,s\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" by(rule RedFAssNull) ultimately show ?thesis by simp qed (*>*) lemma FAssRedsThrow1: "P \ \e,s\ \* \throw e',s'\ \ P \ \e\F{D}:=e\<^sub>2, s\ \* \throw e', s'\" (*<*)by(rule FAssReds1[THEN rtrancl_into_rtrancl, OF _ red_reds.FAssThrow1])(*>*) lemma FAssRedsThrow2: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \Val v,s\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \throw e,s\<^sub>2\" shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \* \throw e,s\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(Val v\F{D}:=e\<^sub>2,s\<^sub>1)" let ?y' = "(Val v\F{D}:=throw e,s\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" by(rule red_reds.FAssThrow2) ultimately show ?thesis by simp qed (*>*) subsubsection";;" lemma SeqReds: "P \ \e,s\ \* \e',s'\ \ P \ \e;;e\<^sub>2, s\ \* \e';;e\<^sub>2, s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) SeqRed[OF step(2)]]) qed (*>*) lemma SeqRedsThrow: "P \ \e,s\ \* \throw e',s'\ \ P \ \e;;e\<^sub>2, s\ \* \throw e', s'\" (*<*)by(rule SeqReds[THEN rtrancl_into_rtrancl, OF _ red_reds.SeqThrow])(*>*) lemma SeqReds2: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \e\<^sub>2',s\<^sub>2\" shows "P \ \e\<^sub>1;;e\<^sub>2, s\<^sub>0\ \* \e\<^sub>2',s\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(Val v\<^sub>1;; e\<^sub>2,s\<^sub>1)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule SeqReds[OF e\<^sub>1_steps]) also have "(?y, ?z) \ (red P)\<^sup>*" by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps]) ultimately show ?thesis by simp qed (*>*) subsubsection"If" lemma CondReds: "P \ \e,s\ \* \e',s'\ \ P \ \if (e) e\<^sub>1 else e\<^sub>2,s\ \* \if (e') e\<^sub>1 else e\<^sub>2,s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) CondRed[OF step(2)]]) qed (*>*) lemma CondRedsThrow: "P \ \e,s\ \* \throw a,s'\ \ P \ \if (e) e\<^sub>1 else e\<^sub>2, s\ \* \throw a,s'\" (*<*)by(rule CondReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CondThrow])(*>*) lemma CondReds2T: assumes e_steps: "P \ \e,s\<^sub>0\ \* \true,s\<^sub>1\" and e\<^sub>1_steps: "P \ \e\<^sub>1, s\<^sub>1\ \* \e',s\<^sub>2\" shows "P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \* \e',s\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(if (true) e\<^sub>1 else e\<^sub>2,s\<^sub>1)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF e_steps]) also have "(?y, ?z) \ (red P)\<^sup>*" by(rule RedCondT[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>1_steps]) ultimately show ?thesis by simp qed (*>*) lemma CondReds2F: assumes e_steps: "P \ \e,s\<^sub>0\ \* \false,s\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2, s\<^sub>1\ \* \e',s\<^sub>2\" shows "P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \* \e',s\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(if (false) e\<^sub>1 else e\<^sub>2,s\<^sub>1)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF e_steps]) also have "(?y, ?z) \ (red P)\<^sup>*" by(rule RedCondF[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps]) ultimately show ?thesis by simp qed (*>*) subsubsection "While" lemma WhileFReds: assumes b_steps: "P \ \b,s\ \* \false,s'\" shows "P \ \while (b) c,s\ \* \unit,s'\" (*<*) by(rule RedWhile[THEN converse_rtrancl_into_rtrancl, OF CondReds[THEN rtrancl_into_rtrancl, OF b_steps RedCondF]]) (*>*) lemma WhileRedsThrow: assumes b_steps: "P \ \b,s\ \* \throw e,s'\" shows "P \ \while (b) c,s\ \* \throw e,s'\" (*<*) by(rule RedWhile[THEN converse_rtrancl_into_rtrancl, OF CondReds[THEN rtrancl_into_rtrancl, OF b_steps red_reds.CondThrow]]) (*>*) lemma WhileTReds: assumes b_steps: "P \ \b,s\<^sub>0\ \* \true,s\<^sub>1\" and c_steps: "P \ \c,s\<^sub>1\ \* \Val v\<^sub>1,s\<^sub>2\" and while_steps: "P \ \while (b) c,s\<^sub>2\ \* \e,s\<^sub>3\" shows "P \ \while (b) c,s\<^sub>0\ \* \e,s\<^sub>3\" (*<*)(is "(?a, ?c) \ (red P)\<^sup>*") proof - let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0)" let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1)" let ?b' = "(c;; while (b) c,s\<^sub>1)" let ?y' = "(Val v\<^sub>1;; while (b) c,s\<^sub>2)" have "(?a, ?b) \ (red P)\<^sup>*" using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp also have "(?b, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF b_steps]) also have "(?y, ?b') \ (red P)\<^sup>*" using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp also have "(?b', ?y') \ (red P)\<^sup>*" by(rule SeqReds[OF c_steps]) also have "(?y', ?c) \ (red P)\<^sup>*" by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF while_steps]) ultimately show ?thesis by simp qed (*>*) lemma WhileTRedsThrow: assumes b_steps: "P \ \b,s\<^sub>0\ \* \true,s\<^sub>1\" and c_steps: "P \ \c,s\<^sub>1\ \* \throw e,s\<^sub>2\" shows "P \ \while (b) c,s\<^sub>0\ \* \throw e,s\<^sub>2\" (*<*)(is "(?a, ?c) \ (red P)\<^sup>*") proof - let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0)" let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1)" let ?b' = "(c;; while (b) c,s\<^sub>1)" let ?y' = "(throw e;; while (b) c,s\<^sub>2)" have "(?a, ?b) \ (red P)\<^sup>*" using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp also have "(?b, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF b_steps]) also have "(?y, ?b') \ (red P)\<^sup>*" using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp also have "(?b', ?y') \ (red P)\<^sup>*" by(rule SeqReds[OF c_steps]) also have "(?y', ?c) \ (red P)" by(rule red_reds.SeqThrow) ultimately show ?thesis by simp qed (*>*) subsubsection"Throw" lemma ThrowReds: "P \ \e,s\ \* \e',s'\ \ P \ \throw e,s\ \* \throw e',s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) ThrowRed[OF step(2)]]) qed (*>*) lemma ThrowRedsNull: "P \ \e,s\ \* \null,s'\ \ P \ \throw e,s\ \* \THROW NullPointer,s'\" (*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ RedThrowNull])(*>*) lemma ThrowRedsThrow: "P \ \e,s\ \* \throw a,s'\ \ P \ \throw e,s\ \* \throw a,s'\" (*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ red_reds.ThrowThrow])(*>*) subsubsection "InitBlock" lemma InitBlockReds_aux: "P \ \e,s\ \* \e',s'\ \ \h l h' l' v. s = (h,l(V\v)) \ s' = (h',l') \ P \ \{V:T := Val v; e},(h,l)\ \* \{V:T := Val(the(l' V)); e'},(h',l'(V:=(l V)))\" (*<*) proof(induct rule: converse_rtrancl_induct2) case refl then show ?case by(fastforce simp: fun_upd_same simp del:fun_upd_apply) next case (step e0 s0 e1 s1) obtain h1 l1 where s1[simp]: "s1 = (h1, l1)" by(cases s1) simp { fix h0 l0 h2 l2 v0 assume [simp]: "s0 = (h0, l0(V \ v0))" and s'[simp]: "s' = (h2, l2)" then have "V \ dom l1" using step(1) by(auto dest!: red_lcl_incr) then obtain v1 where l1V[simp]: "l1 V = \v1\" by blast let ?a = "({V:T; V:=Val v0;; e0},(h0, l0))" let ?b = "({V:T; V:=Val v1;; e1},(h1, l1(V := l0 V)))" let ?c = "({V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l0 V)))" let ?l = "l1(V := l0 V)" and ?v = v1 have e0_steps: "P \ \e0,(h0, l0(V \ v0))\ \ \e1,(h1, l1)\" using step(1) by simp have lv: "\l v. l1 = l(V \ v) \ P \ \{V:T; V:=Val v;; e1},(h1, l)\ \* \{V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l V))\" using step(3) s' s1 by blast moreover have "l1 = ?l(V \ ?v)" by(rule ext) (simp add:fun_upd_def) ultimately have "(?b, ?c) \ (red P)\<^sup>*" using lv[of ?l ?v] by simp then have "(?a, ?c) \ (red P)\<^sup>*" by(rule InitBlockRed[THEN converse_rtrancl_into_rtrancl, OF e0_steps l1V]) } then show ?case by blast qed (*>*) lemma InitBlockReds: "P \ \e, (h,l(V\v))\ \* \e', (h',l')\ \ P \ \{V:T := Val v; e}, (h,l)\ \* \{V:T := Val(the(l' V)); e'}, (h',l'(V:=(l V)))\" (*<*)by(blast dest:InitBlockReds_aux)(*>*) lemma InitBlockRedsFinal: "\ P \ \e,(h,l(V\v))\ \* \e',(h',l')\; final e' \ \ P \ \{V:T := Val v; e},(h,l)\ \* \e',(h', l'(V := l V))\" (*<*) by(fast elim!:InitBlockReds[THEN rtrancl_into_rtrancl] finalE intro:RedInitBlock InitBlockThrow) (*>*) subsubsection "Block" lemma BlockRedsFinal: assumes reds: "P \ \e\<^sub>0,s\<^sub>0\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2)\" and fin: "final e\<^sub>2" shows "\h\<^sub>0 l\<^sub>0. s\<^sub>0 = (h\<^sub>0,l\<^sub>0(V:=None)) \ P \ \{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0)\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V))\" (*<*) using reds proof (induct rule:converse_rtrancl_induct2) case refl thus ?case by(fastforce intro:finalE[OF fin] RedBlock BlockThrow simp del:fun_upd_apply) next case (step e\<^sub>0 s\<^sub>0 e\<^sub>1 s\<^sub>1) have red: "P \ \e\<^sub>0,s\<^sub>0\ \ \e\<^sub>1,s\<^sub>1\" and reds: "P \ \e\<^sub>1,s\<^sub>1\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2)\" and IH: "\h l. s\<^sub>1 = (h,l(V := None)) \ P \ \{V:T; e\<^sub>1},(h,l)\ \* \e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l V))\" and s\<^sub>0: "s\<^sub>0 = (h\<^sub>0, l\<^sub>0(V := None))" by fact+ obtain h\<^sub>1 l\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1)" by fastforce show ?case proof cases assume "assigned V e\<^sub>0" then obtain v e where e\<^sub>0: "e\<^sub>0 = V := Val v;; e" by (unfold assigned_def)blast from red e\<^sub>0 s\<^sub>0 have e\<^sub>1: "e\<^sub>1 = unit;;e" and s\<^sub>1: "s\<^sub>1 = (h\<^sub>0, l\<^sub>0(V \ v))" by auto from e\<^sub>1 fin have "e\<^sub>1 \ e\<^sub>2" by (auto simp:final_def) then obtain e' s' where red1: "P \ \e\<^sub>1,s\<^sub>1\ \ \e',s'\" and reds': "P \ \e',s'\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2)\" using converse_rtranclE2[OF reds] by blast from red1 e\<^sub>1 have es': "e' = e" "s' = s\<^sub>1" by auto show ?case using e\<^sub>0 s\<^sub>1 es' reds' by(fastforce intro!: InitBlockRedsFinal[OF _ fin] simp del:fun_upd_apply) next assume unass: "\ assigned V e\<^sub>0" show ?thesis proof (cases "l\<^sub>1 V") assume None: "l\<^sub>1 V = None" hence "P \ \{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0)\ \ \{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V))\" using s\<^sub>0 s\<^sub>1 red by(simp add: BlockRedNone[OF _ _ unass]) moreover have "P \ \{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V))\ \* \e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l\<^sub>0 V))\" using IH[of _ "l\<^sub>1(V := l\<^sub>0 V)"] s\<^sub>1 None by(simp add:fun_upd_idem) ultimately show ?case by(rule converse_rtrancl_into_rtrancl) next fix v assume Some: "l\<^sub>1 V = Some v" hence "P \ \{V:T;e\<^sub>0},(h\<^sub>0,l\<^sub>0)\ \ \{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V := l\<^sub>0 V))\" using s\<^sub>0 s\<^sub>1 red by(simp add: BlockRedSome[OF _ _ unass]) moreover have "P \ \{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V:= l\<^sub>0 V))\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V))\" using InitBlockRedsFinal[OF _ fin,of _ _ "l\<^sub>1(V:=l\<^sub>0 V)" V] Some reds s\<^sub>1 by(simp add:fun_upd_idem) ultimately show ?case by(rule converse_rtrancl_into_rtrancl) qed qed qed (*>*) subsubsection "try-catch" lemma TryReds: "P \ \e,s\ \* \e',s'\ \ P \ \try e catch(C V) e\<^sub>2,s\ \* \try e' catch(C V) e\<^sub>2,s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) TryRed[OF step(2)]]) qed (*>*) lemma TryRedsVal: "P \ \e,s\ \* \Val v,s'\ \ P \ \try e catch(C V) e\<^sub>2,s\ \* \Val v,s'\" (*<*)by(rule TryReds[THEN rtrancl_into_rtrancl, OF _ RedTry])(*>*) lemma TryCatchRedsFinal: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \Throw a,(h\<^sub>1,l\<^sub>1)\" and h\<^sub>1a: "h\<^sub>1 a = Some(D,fs)" and sub: "P \ D \\<^sup>* C" and e\<^sub>2_steps: "P \ \e\<^sub>2, (h\<^sub>1, l\<^sub>1(V \ Addr a))\ \* \e\<^sub>2', (h\<^sub>2,l\<^sub>2)\" and final: "final e\<^sub>2'" shows "P \ \try e\<^sub>1 catch(C V) e\<^sub>2, s\<^sub>0\ \* \e\<^sub>2', (h\<^sub>2, (l\<^sub>2::locals)(V := l\<^sub>1 V))\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(try Throw a catch(C V) e\<^sub>2,(h\<^sub>1, l\<^sub>1))" let ?b = "({V:Class C; V:=addr a;; e\<^sub>2},(h\<^sub>1, l\<^sub>1))" have bz: "(?b, ?z) \ (red P)\<^sup>*" by(rule InitBlockRedsFinal[OF e\<^sub>2_steps final]) have hp: "hp (h\<^sub>1, l\<^sub>1) a = \(D, fs)\" using h\<^sub>1a by simp have "(?x, ?y) \ (red P)\<^sup>*" by(rule TryReds[OF e\<^sub>1_steps]) also have "(?y, ?z) \ (red P)\<^sup>*" by(rule RedTryCatch[THEN converse_rtrancl_into_rtrancl, OF hp sub bz]) ultimately show ?thesis by simp qed (*>*) lemma TryRedsFail: "\ P \ \e\<^sub>1,s\ \* \Throw a,(h,l)\; h a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \try e\<^sub>1 catch(C V) e\<^sub>2,s\ \* \Throw a,(h,l)\" (*<*)by(fastforce intro!: TryReds[THEN rtrancl_into_rtrancl, OF _ RedTryFail])(*>*) subsubsection "List" lemma ListReds1: "P \ \e,s\ \* \e',s'\ \ P \ \e#es,s\ [\]* \e' # es,s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) ListRed1[OF step(2)]]) qed (*>*) lemma ListReds2: "P \ \es,s\ [\]* \es',s'\ \ P \ \Val v # es,s\ [\]* \Val v # es',s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) ListRed2[OF step(2)]]) qed (*>*) lemma ListRedsVal: "\ P \ \e,s\<^sub>0\ \* \Val v,s\<^sub>1\; P \ \es,s\<^sub>1\ [\]* \es',s\<^sub>2\ \ \ P \ \e#es,s\<^sub>0\ [\]* \Val v # es',s\<^sub>2\" (*<*)by(rule rtrancl_trans[OF ListReds1 ListReds2])(*>*) subsubsection"Call" text\First a few lemmas on what happens to free variables during redction.\ lemma assumes wf: "wwf_J_prog P" shows Red_fv: "P \ \e,(h,l)\ \ \e',(h',l')\ \ fv e' \ fv e" and "P \ \es,(h,l)\ [\] \es',(h',l')\ \ fvs es' \ fvs es" (*<*) proof (induct rule:red_reds_inducts) case (RedCall h l a C fs M Ts T pns body D vs) hence "fv body \ {this} \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def) with RedCall.hyps show ?case by fastforce qed auto (*>*) lemma Red_dom_lcl: "P \ \e,(h,l)\ \ \e',(h',l')\ \ dom l' \ dom l \ fv e" and "P \ \es,(h,l)\ [\] \es',(h',l')\ \ dom l' \ dom l \ fvs es" (*<*) proof (induct rule:red_reds_inducts) case RedLAss thus ?case by(force split:if_splits) next case CallParams thus ?case by(force split:if_splits) next case BlockRedNone thus ?case by clarsimp (fastforce split:if_splits) next case BlockRedSome thus ?case by clarsimp (fastforce split:if_splits) next case InitBlockRed thus ?case by clarsimp (fastforce split:if_splits) qed auto (*>*) lemma Reds_dom_lcl: assumes wf: "wwf_J_prog P" shows "P \ \e,(h,l)\ \* \e',(h',l')\ \ dom l' \ dom l \ fv e" (*<*) proof(induct rule: converse_rtrancl_induct_red) case 1 then show ?case by blast next case 2 then show ?case using wf by(blast dest: Red_fv Red_dom_lcl) qed (*>*) text\Now a few lemmas on the behaviour of blocks during reduction.\ (* If you want to avoid the premise "distinct" further down \ consts upd_vals :: "locals \ vname list \ val list \ val list" primrec "upd_vals l [] vs = []" "upd_vals l (V#Vs) vs = (if V \ set Vs then hd vs else the(l V)) # upd_vals l Vs (tl vs)" lemma [simp]: "\vs. length(upd_vals l Vs vs) = length Vs" by(induct Vs, auto) *) lemma override_on_upd_lemma: "(override_on f (g(a\b)) A)(a := g a) = override_on f g (insert a A)" (*<*)by(rule ext) (simp add:override_on_def) declare fun_upd_apply[simp del] map_upds_twist[simp del] (*>*) lemma blocksReds: "\l. \ length Vs = length Ts; length vs = length Ts; distinct Vs; P \ \e, (h,l(Vs [\] vs))\ \* \e', (h',l')\ \ \ P \ \blocks(Vs,Ts,vs,e), (h,l)\ \* \blocks(Vs,Ts,map (the \ l') Vs,e'), (h',override_on l' l (set Vs))\" (*<*) proof(induct Vs Ts vs e rule:blocks_induct) case (1 V Vs T Ts v vs e) show ?case using InitBlockReds[OF "1.hyps"[of "l(V\v)"]] "1.prems" by(auto simp:override_on_upd_lemma) qed auto (*>*) lemma blocksFinal: "\l. \ length Vs = length Ts; length vs = length Ts; final e \ \ P \ \blocks(Vs,Ts,vs,e), (h,l)\ \* \e, (h,l)\" (*<*) proof(induct Vs Ts vs e rule:blocks_induct) case 1 show ?case using "1.prems" InitBlockReds[OF "1.hyps"] by(fastforce elim!:finalE elim: rtrancl_into_rtrancl[OF _ RedInitBlock] rtrancl_into_rtrancl[OF _ InitBlockThrow]) qed auto (*>*) lemma blocksRedsFinal: assumes wf: "length Vs = length Ts" "length vs = length Ts" "distinct Vs" and reds: "P \ \e, (h,l(Vs [\] vs))\ \* \e', (h',l')\" and fin: "final e'" and l'': "l'' = override_on l' l (set Vs)" shows "P \ \blocks(Vs,Ts,vs,e), (h,l)\ \* \e', (h',l'')\" (*<*) proof - let ?bv = "blocks(Vs,Ts,map (the o l') Vs,e')" have "P \ \blocks(Vs,Ts,vs,e), (h,l)\ \* \?bv, (h',l'')\" using l'' by simp (rule blocksReds[OF wf reds]) also have "P \ \?bv, (h',l'')\ \* \e', (h',l'')\" using wf by(fastforce intro:blocksFinal fin) finally show ?thesis . qed (*>*) text\An now the actual method call reduction lemmas.\ lemma CallRedsObj: "P \ \e,s\ \* \e',s'\ \ P \ \e\M(es),s\ \* \e'\M(es),s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) CallObj[OF step(2)]]) qed (*>*) lemma CallRedsParams: "P \ \es,s\ [\]* \es',s'\ \ P \ \(Val v)\M(es),s\ \* \(Val v)\M(es'),s'\" (*<*) proof(induct rule: rtrancl_induct2) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) CallParams[OF step(2)]]) qed (*>*) lemma CallRedsFinal: assumes wwf: "wwf_J_prog P" and "P \ \e,s\<^sub>0\ \* \addr a,s\<^sub>1\" "P \ \es,s\<^sub>1\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2)\" "h\<^sub>2 a = Some(C,fs)" "P \ C sees M:Ts\T = (pns,body) in D" "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [this \ Addr a, pns[\]vs]" and body: "P \ \body,(h\<^sub>2,l\<^sub>2')\ \* \ef,(h\<^sub>3,l\<^sub>3)\" and "final ef" shows "P \ \e\M(es), s\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2)\" (*<*) proof - have wf: "size Ts = size pns \ distinct pns \ this \ set pns" and wt: "fv body \ {this} \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ from body[THEN Red_lcl_add, of l\<^sub>2] have body': "P \ \body,(h\<^sub>2,l\<^sub>2(this\ Addr a, pns[\]vs))\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3)\" by (simp add:l\<^sub>2') have "dom l\<^sub>3 \ {this} \ set pns" using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns) = l\<^sub>2" by(fastforce simp add:map_add_def override_on_def fun_eq_iff) have "P \ \e\M(es),s\<^sub>0\ \* \(addr a)\M(es),s\<^sub>1\" by(rule CallRedsObj)(rule assms(2)) also have "P \ \(addr a)\M(es),s\<^sub>1\ \* \(addr a)\M(map Val vs),(h\<^sub>2,l\<^sub>2)\" by(rule CallRedsParams)(rule assms(3)) also have "P \ \(addr a)\M(map Val vs), (h\<^sub>2,l\<^sub>2)\ \ \blocks(this#pns, Class D#Ts, Addr a#vs, body), (h\<^sub>2,l\<^sub>2)\" by(rule RedCall)(auto simp: assms wf, rule assms(5)) also (rtrancl_into_rtrancl) have "P \ \blocks(this#pns, Class D#Ts, Addr a#vs, body), (h\<^sub>2,l\<^sub>2)\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns))\" by(rule blocksRedsFinal, insert assms wf body', simp_all) finally show ?thesis using eql\<^sub>2 by simp qed (*>*) lemma CallRedsThrowParams: assumes e_steps: "P \ \e,s\<^sub>0\ \* \Val v,s\<^sub>1\" and es_steps: "P \ \es,s\<^sub>1\ [\]* \map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2\" shows "P \ \e\M(es),s\<^sub>0\ \* \throw a,s\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(Val v\M(es),s\<^sub>1)" let ?y' = "(Val v\M(map Val vs\<^sub>1 @ throw a # es\<^sub>2),s\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) also have "(?y', ?z) \ (red P)\<^sup>*" using CallThrowParams by fast ultimately show ?thesis by simp qed (*>*) lemma CallRedsThrowObj: "P \ \e,s0\ \* \throw a,s\<^sub>1\ \ P \ \e\M(es),s0\ \* \throw a,s\<^sub>1\" (*<*)by(rule CallRedsObj[THEN rtrancl_into_rtrancl, OF _ CallThrowObj])(*>*) lemma CallRedsNull: assumes e_steps: "P \ \e,s\<^sub>0\ \* \null,s\<^sub>1\" and es_steps: "P \ \es,s\<^sub>1\ [\]* \map Val vs,s\<^sub>2\" shows "P \ \e\M(es),s\<^sub>0\ \* \THROW NullPointer,s\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(null\M(es),s\<^sub>1)" let ?y' = "(null\M(map Val vs),s\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) also have "(?y', ?z) \ (red P)" by(rule RedCallNull) ultimately show ?thesis by simp qed (*>*) subsubsection "The main Theorem" lemma assumes wwf: "wwf_J_prog P" shows big_by_small: "P \ \e,s\ \ \e',s'\ \ P \ \e,s\ \* \e',s'\" and bigs_by_smalls: "P \ \es,s\ [\] \es',s'\ \ P \ \es,s\ [\]* \es',s'\" (*<*) proof (induct rule: eval_evals.inducts) case New thus ?case by (auto simp:RedNew) next case NewFail thus ?case by (auto simp:RedNewFail) next case Cast thus ?case by(fastforce intro:CastRedsAddr) next case CastNull thus ?case by(simp add:CastRedsNull) next case CastFail thus ?case by(fastforce intro!:CastRedsFail) next case CastThrow thus ?case by(auto dest!:eval_final simp:CastRedsThrow) next case Val thus ?case by simp next case BinOp thus ?case by(auto simp:BinOpRedsVal) next case BinOpThrow1 thus ?case by(auto dest!:eval_final simp: BinOpRedsThrow1) next case BinOpThrow2 thus ?case by(auto dest!:eval_final simp: BinOpRedsThrow2) next case Var thus ?case by (auto simp:RedVar) next case LAss thus ?case by(auto simp: LAssRedsVal) next case LAssThrow thus ?case by(auto dest!:eval_final simp: LAssRedsThrow) next case FAcc thus ?case by(auto intro:FAccRedsVal) next case FAccNull thus ?case by(simp add:FAccRedsNull) next case FAccThrow thus ?case by(auto dest!:eval_final simp:FAccRedsThrow) next case FAss thus ?case by(auto simp:FAssRedsVal) next case FAssNull thus ?case by(auto simp:FAssRedsNull) next case FAssThrow1 thus ?case by(auto dest!:eval_final simp:FAssRedsThrow1) next case FAssThrow2 thus ?case by(auto dest!:eval_final simp:FAssRedsThrow2) next case CallObjThrow thus ?case by(auto dest!:eval_final simp:CallRedsThrowObj) next case CallNull thus ?case by(simp add:CallRedsNull) next case CallParamsThrow thus ?case by(auto dest!:evals_final simp:CallRedsThrowParams) next case (Call e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C fs M Ts T pns body D l\<^sub>2' e' h\<^sub>3 l\<^sub>3) have IHe: "P \ \e,s\<^sub>0\ \* \addr a,s\<^sub>1\" and IHes: "P \ \ps,s\<^sub>1\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2)\" and h\<^sub>2a: "h\<^sub>2 a = Some(C,fs)" and "method": "P \ C sees M:Ts\T = (pns,body) in D" and same_length: "length vs = length pns" and l\<^sub>2': "l\<^sub>2' = [this \ Addr a, pns[\]vs]" and eval_body: "P \ \body,(h\<^sub>2, l\<^sub>2')\ \ \e',(h\<^sub>3, l\<^sub>3)\" and IHbody: "P \ \body,(h\<^sub>2,l\<^sub>2')\ \* \e',(h\<^sub>3,l\<^sub>3)\" by fact+ show "P \ \e\M(ps),s\<^sub>0\ \* \e',(h\<^sub>3, l\<^sub>2)\" using "method" same_length l\<^sub>2' h\<^sub>2a IHbody eval_final[OF eval_body] by(fastforce intro:CallRedsFinal[OF wwf IHe IHes]) next case Block thus ?case by(auto simp: BlockRedsFinal dest:eval_final) next case Seq thus ?case by(auto simp:SeqReds2) next case SeqThrow thus ?case by(auto dest!:eval_final simp: SeqRedsThrow) next case CondT thus ?case by(auto simp:CondReds2T) next case CondF thus ?case by(auto simp:CondReds2F) next case CondThrow thus ?case by(auto dest!:eval_final simp:CondRedsThrow) next case WhileF thus ?case by(auto simp:WhileFReds) next case WhileT thus ?case by(auto simp: WhileTReds) next case WhileCondThrow thus ?case by(auto dest!:eval_final simp: WhileRedsThrow) next case WhileBodyThrow thus ?case by(auto dest!:eval_final simp: WhileTRedsThrow) next case Throw thus ?case by(auto simp:ThrowReds) next case ThrowNull thus ?case by(auto simp:ThrowRedsNull) next case ThrowThrow thus ?case by(auto dest!:eval_final simp:ThrowRedsThrow) next case Try thus ?case by(simp add:TryRedsVal) next case TryCatch thus ?case by(fast intro!: TryCatchRedsFinal dest!:eval_final) next case TryThrow thus ?case by(fastforce intro!:TryRedsFail) next case Nil thus ?case by simp next case Cons thus ?case by(fastforce intro!:Cons_eq_appendI[OF refl refl] ListRedsVal) next case ConsThrow thus ?case by(fastforce elim: ListReds1) qed (*>*) subsection\Big steps simulates small step\ text\This direction was carried out by Norbert Schirmer and Daniel Wasserrab.\ text \The big step equivalent of \RedWhile\:\ lemma unfold_while: "P \ \while(b) c,s\ \ \e',s'\ = P \ \if(b) (c;;while(b) c) else (unit),s\ \ \e',s'\" (*<*) proof assume "P \ \while (b) c,s\ \ \e',s'\" thus "P \ \if (b) (c;; while (b) c) else unit,s\ \ \e',s'\" by cases (fastforce intro: eval_evals.intros)+ next assume "P \ \if (b) (c;; while (b) c) else unit,s\ \ \e',s'\" thus "P \ \while (b) c,s\ \ \e',s'\" proof (cases) fix a assume e': "e' = throw a" assume "P \ \b,s\ \ \throw a,s'\" hence "P \ \while(b) c,s\ \ \throw a,s'\" by (rule WhileCondThrow) with e' show ?thesis by simp next fix s\<^sub>1 assume eval_false: "P \ \b,s\ \ \false,s\<^sub>1\" and eval_unit: "P \ \unit,s\<^sub>1\ \ \e',s'\" with eval_unit have "s' = s\<^sub>1" "e' = unit" by (auto elim: eval_cases) moreover from eval_false have "P \ \while (b) c,s\ \ \unit,s\<^sub>1\" by - (rule WhileF, simp) ultimately show ?thesis by simp next fix s\<^sub>1 assume eval_true: "P \ \b,s\ \ \true,s\<^sub>1\" and eval_rest: "P \ \c;; while (b) c,s\<^sub>1\\\e',s'\" from eval_rest show ?thesis proof (cases) fix s\<^sub>2 v\<^sub>1 assume "P \ \c,s\<^sub>1\ \ \Val v\<^sub>1,s\<^sub>2\" "P \ \while (b) c,s\<^sub>2\ \ \e',s'\" with eval_true show "P \ \while(b) c,s\ \ \e',s'\" by (rule WhileT) next fix a assume "P \ \c,s\<^sub>1\ \ \throw a,s'\" "e' = throw a" with eval_true show "P \ \while(b) c,s\ \ \e',s'\" by (iprover intro: WhileBodyThrow) qed qed qed (*>*) lemma blocksEval: "\Ts vs l l'. \size ps = size Ts; size ps = size vs; P \ \blocks(ps,Ts,vs,e),(h,l)\ \ \e',(h',l')\ \ \ \ l''. P \ \e,(h,l(ps[\]vs))\ \ \e',(h',l'')\" (*<*) proof (induct ps) case Nil then show ?case by fastforce next case (Cons p ps') have length_eqs: "length (p # ps') = length Ts" "length (p # ps') = length vs" by fact+ then obtain T Ts' where Ts: "Ts = T#Ts'" by (cases "Ts") simp obtain v vs' where vs: "vs = v#vs'" using length_eqs by (cases "vs") simp have "P \ \blocks (p # ps', Ts, vs, e),(h,l)\ \ \e',(h', l')\" by fact with Ts vs have "P \ \{p:T := Val v; blocks (ps', Ts', vs', e)},(h,l)\ \ \e',(h', l')\" by simp then obtain l''' where eval_ps': "P \ \blocks (ps', Ts', vs', e),(h, l(p\v))\ \ \e',(h', l''')\" and l''': "l'=l'''(p:=l p)" by (auto elim!: eval_cases) then obtain l'' where - hyp: "P \ \e,(h, l(p\v)(ps'[\]vs'))\ \ \e',(h', l'')\" + hyp: "P \ \e,(h, l(p\v, ps'[\]vs'))\ \ \e',(h', l'')\" using length_eqs Ts vs Cons.hyps [OF _ _ eval_ps'] by auto from hyp show "\l''. P \ \e,(h, l(p # ps'[\]vs))\ \ \e',(h', l'')\" using Ts vs by auto qed (*>*) (* FIXME exercise: show precise relationship between l' and l'': lemma blocksEval: "\ Ts vs l l'. \length ps = length Ts; length ps = length vs; P\ \blocks(ps,Ts,vs,e),(h,l)\ \ \e',(h',l')\ \ \ \ l''. P \ \e,(h,l(ps[\]vs))\ \ \e',(h',l'')\ \ l'=l''(l|set ps)" proof (induct ps) case Nil then show ?case by simp next case (Cons p ps') have length_eqs: "length (p # ps') = length Ts" "length (p # ps') = length vs" . then obtain T Ts' where Ts: "Ts=T#Ts'" by (cases "Ts") simp obtain v vs' where vs: "vs=v#vs'" using length_eqs by (cases "vs") simp have "P \ \blocks (p # ps', Ts, vs, e),(h,l)\ \ \e',(h', l')\". with Ts vs have "P \ \{p:T := Val v; blocks (ps', Ts', vs', e)},(h,l)\ \ \e',(h', l')\" by simp then obtain l''' where eval_ps': "P \ \blocks (ps', Ts', vs', e),(h, l(p\v))\ \ \e',(h', l''')\" and l''': "l'=l'''(p:=l p)" by (cases) (auto elim: eval_cases) then obtain l'' where hyp: "P \ \e,(h, l(p\v)(ps'[\]vs'))\ \ \e',(h', l'')\" and l'': "l''' = l''(l(p\v)|set ps')" using length_eqs Ts vs Cons.hyps [OF _ _ eval_ps'] by auto have "l' = l''(l|set (p # ps'))" proof - have "(l''(l(p\v)|set ps'))(p := l p) = l''(l|insert p (set ps'))" by (induct ps') (auto intro: ext simp add: fun_upd_def override_on_def) with l''' l'' show ?thesis by simp qed with hyp show "\l''. P \ \e,(h, l(p # ps'[\]vs))\ \ \e',(h', l'')\ \ l' = l''(l|set (p # ps'))" using Ts vs by auto qed *) lemma assumes wf: "wwf_J_prog P" shows eval_restrict_lcl: "P \ \e,(h,l)\ \ \e',(h',l')\ \ (\W. fv e \ W \ P \ \e,(h,l|`W)\ \ \e',(h',l'|`W)\)" and "P \ \es,(h,l)\ [\] \es',(h',l')\ \ (\W. fvs es \ W \ P \ \es,(h,l|`W)\ [\] \es',(h',l'|`W)\)" (*<*) proof(induct rule:eval_evals_inducts) case (Block e\<^sub>0 h\<^sub>0 l\<^sub>0 V e\<^sub>1 h\<^sub>1 l\<^sub>1 T) have IH: "\W. fv e\<^sub>0 \ W \ P \ \e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None)|`W)\ \ \e\<^sub>1,(h\<^sub>1,l\<^sub>1|`W)\" by fact have "fv({V:T; e\<^sub>0}) \ W" by fact+ hence "fv e\<^sub>0 - {V} \ W" by simp_all hence "fv e\<^sub>0 \ insert V W" by fast from IH[OF this] have "P \ \e\<^sub>0,(h\<^sub>0, (l\<^sub>0|`W)(V := None))\ \ \e\<^sub>1,(h\<^sub>1, l\<^sub>1|`insert V W)\" by fastforce from eval_evals.Block[OF this] show ?case by fastforce next case Seq thus ?case by simp (blast intro:eval_evals.Seq) next case New thus ?case by(simp add:eval_evals.intros) next case NewFail thus ?case by(simp add:eval_evals.intros) next case Cast thus ?case by simp (blast intro:eval_evals.Cast) next case CastNull thus ?case by simp (blast intro:eval_evals.CastNull) next case CastFail thus ?case by simp (blast intro:eval_evals.CastFail) next case CastThrow thus ?case by(simp add:eval_evals.intros) next case Val thus ?case by(simp add:eval_evals.intros) next case BinOp thus ?case by simp (blast intro:eval_evals.BinOp) next case BinOpThrow1 thus ?case by simp (blast intro:eval_evals.BinOpThrow1) next case BinOpThrow2 thus ?case by simp (blast intro:eval_evals.BinOpThrow2) next case Var thus ?case by(simp add:eval_evals.intros) next case (LAss e h\<^sub>0 l\<^sub>0 v h l l' V) have IH: "\W. fv e \ W \ P \ \e,(h\<^sub>0,l\<^sub>0|`W)\ \ \Val v,(h,l|`W)\" and [simp]: "l' = l(V \ v)" by fact+ have "fv (V:=e) \ W" by fact hence fv: "fv e \ W" and VinW: "V \ W" by auto from eval_evals.LAss[OF IH[OF fv] refl, of V] VinW show ?case by simp next case LAssThrow thus ?case by(fastforce intro: eval_evals.LAssThrow) next case FAcc thus ?case by simp (blast intro: eval_evals.FAcc) next case FAccNull thus ?case by(fastforce intro: eval_evals.FAccNull) next case FAccThrow thus ?case by(fastforce intro: eval_evals.FAccThrow) next case FAss thus ?case by simp (blast intro: eval_evals.FAss) next case FAssNull thus ?case by simp (blast intro: eval_evals.FAssNull) next case FAssThrow1 thus ?case by simp (blast intro: eval_evals.FAssThrow1) next case FAssThrow2 thus ?case by simp (blast intro: eval_evals.FAssThrow2) next case CallObjThrow thus ?case by simp (blast intro: eval_evals.intros) next case CallNull thus ?case by simp (blast intro: eval_evals.CallNull) next case CallParamsThrow thus ?case by simp (blast intro: eval_evals.CallParamsThrow) next case (Call e h\<^sub>0 l\<^sub>0 a h\<^sub>1 l\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C fs M Ts T pns body D l\<^sub>2' e' h\<^sub>3 l\<^sub>3) have IHe: "\W. fv e \ W \ P \ \e,(h\<^sub>0,l\<^sub>0|`W)\ \ \addr a,(h\<^sub>1,l\<^sub>1|`W)\" and IHps: "\W. fvs ps \ W \ P \ \ps,(h\<^sub>1,l\<^sub>1|`W)\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2|`W)\" and IHbd: "\W. fv body \ W \ P \ \body,(h\<^sub>2,l\<^sub>2'|`W)\ \ \e',(h\<^sub>3,l\<^sub>3|`W)\" and h\<^sub>2a: "h\<^sub>2 a = Some (C, fs)" and "method": "P \ C sees M: Ts\T = (pns, body) in D" and same_len: "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [this \ Addr a, pns [\] vs]" by fact+ have "fv (e\M(ps)) \ W" by fact hence fve: "fv e \ W" and fvps: "fvs(ps) \ W" by auto have wfmethod: "size Ts = size pns \ this \ set pns" and fvbd: "fv body \ {this} \ set pns" using "method" wf by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ show ?case using IHbd[OF fvbd] l\<^sub>2' same_len wfmethod h\<^sub>2a eval_evals.Call[OF IHe[OF fve] IHps[OF fvps] _ "method" same_len l\<^sub>2'] by (simp add:subset_insertI) next case SeqThrow thus ?case by simp (blast intro: eval_evals.SeqThrow) next case CondT thus ?case by simp (blast intro: eval_evals.CondT) next case CondF thus ?case by simp (blast intro: eval_evals.CondF) next case CondThrow thus ?case by simp (blast intro: eval_evals.CondThrow) next case WhileF thus ?case by simp (blast intro: eval_evals.WhileF) next case WhileT thus ?case by simp (blast intro: eval_evals.WhileT) next case WhileCondThrow thus ?case by simp (blast intro: eval_evals.WhileCondThrow) next case WhileBodyThrow thus ?case by simp (blast intro: eval_evals.WhileBodyThrow) next case Throw thus ?case by simp (blast intro: eval_evals.Throw) next case ThrowNull thus ?case by simp (blast intro: eval_evals.ThrowNull) next case ThrowThrow thus ?case by simp (blast intro: eval_evals.ThrowThrow) next case Try thus ?case by simp (blast intro: eval_evals.Try) next case (TryCatch e\<^sub>1 h\<^sub>0 l\<^sub>0 a h\<^sub>1 l\<^sub>1 D fs C e\<^sub>2 V e\<^sub>2' h\<^sub>2 l\<^sub>2) have IH\<^sub>1: "\W. fv e\<^sub>1 \ W \ P \ \e\<^sub>1,(h\<^sub>0,l\<^sub>0|`W)\ \ \Throw a,(h\<^sub>1,l\<^sub>1|`W)\" and IH\<^sub>2: "\W. fv e\<^sub>2 \ W \ P \ \e\<^sub>2,(h\<^sub>1,l\<^sub>1(V\Addr a)|`W)\ \ \e\<^sub>2',(h\<^sub>2,l\<^sub>2|`W)\" and lookup: "h\<^sub>1 a = Some(D, fs)" and subtype: "P \ D \\<^sup>* C" by fact+ have "fv (try e\<^sub>1 catch(C V) e\<^sub>2) \ W" by fact hence fv\<^sub>1: "fv e\<^sub>1 \ W" and fv\<^sub>2: "fv e\<^sub>2 \ insert V W" by auto have IH\<^sub>2': "P \ \e\<^sub>2,(h\<^sub>1,(l\<^sub>1|`W)(V \ Addr a))\ \ \e\<^sub>2',(h\<^sub>2,l\<^sub>2|`insert V W)\" using IH\<^sub>2[OF fv\<^sub>2] fun_upd_restrict[of l\<^sub>1 W] (*FIXME just l|W instead of l|(W-V) in simp rule??*) by simp with eval_evals.TryCatch[OF IH\<^sub>1[OF fv\<^sub>1] _ subtype IH\<^sub>2'] lookup show ?case by fastforce next case TryThrow thus ?case by simp (blast intro: eval_evals.TryThrow) next case Nil thus ?case by (simp add: eval_evals.Nil) next case Cons thus ?case by simp (blast intro: eval_evals.Cons) next case ConsThrow thus ?case by simp (blast intro: eval_evals.ConsThrow) qed (*>*) lemma eval_notfree_unchanged: "P \ \e,(h,l)\ \ \e',(h',l')\ \ (\V. V \ fv e \ l' V = l V)" and "P \ \es,(h,l)\ [\] \es',(h',l')\ \ (\V. V \ fvs es \ l' V = l V)" (*<*) proof(induct rule:eval_evals_inducts) case LAss thus ?case by(simp add:fun_upd_apply) next case Block thus ?case by (simp only:fun_upd_apply split:if_splits) fastforce next case TryCatch thus ?case by (simp only:fun_upd_apply split:if_splits) fastforce qed simp_all (*>*) lemma eval_closed_lcl_unchanged: "\ P \ \e,(h,l)\ \ \e',(h',l')\; fv e = {} \ \ l' = l" (*<*)by(fastforce dest:eval_notfree_unchanged simp add:fun_eq_iff [where 'b="val option"])(*>*) lemma list_eval_Throw: assumes eval_e: "P \ \throw x,s\ \ \e',s'\" shows "P \ \map Val vs @ throw x # es',s\ [\] \map Val vs @ e' # es',s'\" (*<*) proof - from eval_e obtain a where e': "e' = Throw a" by (cases) (auto dest!: eval_final) { fix es have "\vs. es = map Val vs @ throw x # es' \ P \ \es,s\[\]\map Val vs @ e' # es',s'\" proof (induct es type: list) case Nil thus ?case by simp next case (Cons e es vs) have e_es: "e # es = map Val vs @ throw x # es'" by fact show "P \ \e # es,s\ [\] \map Val vs @ e' # es',s'\" proof (cases vs) case Nil with e_es obtain "e=throw x" "es=es'" by simp moreover from eval_e e' have "P \ \throw x # es,s\ [\] \Throw a # es,s'\" by (iprover intro: ConsThrow) ultimately show ?thesis using Nil e' by simp next case (Cons v vs') have vs: "vs = v # vs'" by fact with e_es obtain e: "e=Val v" and es:"es= map Val vs' @ throw x # es'" by simp from e have "P \ \e,s\ \ \Val v,s\" by (iprover intro: eval_evals.Val) moreover from es have "P \ \es,s\ [\] \map Val vs' @ e' # es',s'\" by (rule Cons.hyps) ultimately show "P \ \e#es,s\ [\] \map Val vs @ e' # es',s'\" using vs by (auto intro: eval_evals.Cons) qed qed } thus ?thesis by simp qed (*>*) (* Hiermit kann man die ganze pair-Splitterei in den automatischen Taktiken abschalten. Wieder anschalten siehe nach dem Beweis. *) (*<*) declare split_paired_All [simp del] split_paired_Ex [simp del] (*>*) (* FIXME exercise 1: define a big step semantics where the body of a procedure can access not juts this and pns but all of the enclosing l. What exactly is fed in? What exactly is returned at the end? Notion: "dynamic binding" excercise 2: the semantics of exercise 1 is closer to the small step semantics. Reformulate equivalence proof by modifying call lemmas. *) text \The key lemma:\ lemma assumes wf: "wwf_J_prog P" shows extend_1_eval: "P \ \e,s\ \ \e'',s''\ \ (\s' e'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\)" and extend_1_evals: "P \ \es,t\ [\] \es'',t''\ \ (\t' es'. P \ \es'',t''\ [\] \es',t'\ \ P \ \es,t\ [\] \es',t'\)" (*<*) proof (induct rule: red_reds.inducts) case (RedCall s a C fs M Ts T pns body D vs s' e') have "P \ \addr a,s\ \ \addr a,s\" by (rule eval_evals.intros) moreover have finals: "finals(map Val vs)" by simp obtain h\<^sub>2 l\<^sub>2 where s: "s = (h\<^sub>2,l\<^sub>2)" by (cases s) with finals have "P \ \map Val vs,s\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2)\" by (iprover intro: eval_finalsId) moreover from s have "h\<^sub>2 a = Some (C, fs)" using RedCall by simp moreover have "method": "P \ C sees M: Ts\T = (pns, body) in D" by fact moreover have same_len\<^sub>1: "length Ts = length pns" and this_distinct: "this \ set pns" and fv: "fv body \ {this} \ set pns" using "method" wf by (fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ have same_len: "length vs = length pns" by fact moreover obtain l\<^sub>2' where l\<^sub>2': "l\<^sub>2' = [this\Addr a,pns[\]vs]" by simp moreover obtain h\<^sub>3 l\<^sub>3 where s': "s' = (h\<^sub>3,l\<^sub>3)" by (cases s') have eval_blocks: "P \ \blocks (this # pns, Class D # Ts, Addr a # vs, body),s\ \ \e',s'\" by fact hence id: "l\<^sub>3 = l\<^sub>2" using fv s s' same_len\<^sub>1 same_len by(fastforce elim: eval_closed_lcl_unchanged) from eval_blocks obtain l\<^sub>3' where "P \ \body,(h\<^sub>2,l\<^sub>2')\ \ \e',(h\<^sub>3,l\<^sub>3')\" proof - from same_len\<^sub>1 have "length(this#pns) = length(Class D#Ts)" by simp moreover from same_len\<^sub>1 same_len have "length (this#pns) = length (Addr a#vs)" by simp moreover from eval_blocks have "P \ \blocks (this#pns,Class D#Ts,Addr a#vs,body),(h\<^sub>2,l\<^sub>2)\ \\e',(h\<^sub>3,l\<^sub>3)\" using s s' by simp ultimately obtain l'' where "P \ \body,(h\<^sub>2,l\<^sub>2(this # pns[\]Addr a # vs))\\\e',(h\<^sub>3, l'')\" by (blast dest:blocksEval) from eval_restrict_lcl[OF wf this fv] this_distinct same_len\<^sub>1 same_len have "P \ \body,(h\<^sub>2,[this # pns[\]Addr a # vs])\ \ \e',(h\<^sub>3, l''|`(set(this#pns)))\" by(simp add:subset_insert_iff insert_Diff_if) thus ?thesis by(fastforce intro!:that simp add: l\<^sub>2') qed ultimately have "P \ \(addr a)\M(map Val vs),s\ \ \e',(h\<^sub>3,l\<^sub>2)\" by (rule Call) with s' id show ?case by simp next case RedNew thus ?case by (iprover elim: eval_cases intro: eval_evals.intros) next case RedNewFail thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (CastRed e s e'' s'' C s' e') thus ?case by(cases s, cases s') (erule eval_cases, auto intro: eval_evals.intros) next case RedCastNull thus ?case by (iprover elim: eval_cases intro: eval_evals.intros) next case (RedCast s a D fs C s'' e'') thus ?case by (cases s) (auto elim: eval_cases intro: eval_evals.intros) next case (RedCastFail s a D fs C s'' e'') thus ?case by (cases s) (auto elim!: eval_cases intro: eval_evals.intros) next case (BinOpRed1 e s e' s' bop e\<^sub>2 s'' e') thus ?case by (cases s'')(erule eval_cases,auto intro: eval_evals.intros) next case BinOpRed2 thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId) next case RedBinOp thus ?case by (iprover elim: eval_cases intro: eval_evals.intros) next case (RedVar s V v s' e') thus ?case by (cases s)(fastforce elim: eval_cases intro: eval_evals.intros) next case (LAssRed e s e' s' V s'') thus ?case by (cases s'')(erule eval_cases,auto intro: eval_evals.intros) next case RedLAss thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (FAccRed e s e' s' F D s'') thus ?case by (cases s'')(erule eval_cases,auto intro: eval_evals.intros) next case (RedFAcc s a C fs F D v s' e') thus ?case by (cases s)(fastforce elim: eval_cases intro: eval_evals.intros) next case RedFAccNull thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case (FAssRed1 e s e' s'' F D e\<^sub>2 s' e') thus ?case by (cases s')(erule eval_cases, auto intro: eval_evals.intros) next case (FAssRed2 e s e' s'' v F D s' e') thus ?case by (cases s) (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId) next case RedFAss thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case RedFAssNull thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case CallObj thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case CallParams thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId) next case RedCallNull thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros eval_finalsId) next case InitBlockRed thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId simp add: map_upd_triv fun_upd_same) next case (RedInitBlock V T v u s s' e') have "P \ \Val u,s\ \ \e',s'\" by fact then obtain s': "s'=s" and e': "e'=Val u" by cases simp obtain h l where s: "s=(h,l)" by (cases s) have "P \ \{V:T :=Val v; Val u},(h,l)\ \ \Val u,(h, (l(V\v))(V:=l V))\" by (fastforce intro!: eval_evals.intros) thus "P \ \{V:T := Val v; Val u},s\ \ \e',s'\" using s s' e' by simp next case BlockRedNone thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros simp add: fun_upd_same fun_upd_idem) next case BlockRedSome thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros simp add: fun_upd_same fun_upd_idem) next case (RedBlock V T v s s' e') have "P \ \Val v,s\ \ \e',s'\" by fact then obtain s': "s'=s" and e': "e'=Val v" by cases simp obtain h l where s: "s=(h,l)" by (cases s) have "P \ \Val v,(h,l(V:=None))\ \ \Val v,(h,l(V:=None))\" by (rule eval_evals.intros) hence "P \ \{V:T;Val v},(h,l)\ \ \Val v,(h,(l(V:=None))(V:=l V))\" by (rule eval_evals.Block) thus "P \ \{V:T; Val v},s\ \ \e',s'\" using s s' e' by simp next case SeqRed thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedSeq thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case CondRed thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedCondT thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedCondF thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedWhile thus ?case by (auto simp add: unfold_while intro:eval_evals.intros elim:eval_cases) next case ThrowRed thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedThrowNull thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (TryRed e s e' s' C V e\<^sub>2 s'' e') thus ?case by (cases s, cases s'', auto elim!: eval_cases intro: eval_evals.intros) next case RedTry thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedTryCatch thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (RedTryFail s a D fs C V e\<^sub>2 s' e') thus ?case by (cases s)(auto elim!: eval_cases intro: eval_evals.intros) next case ListRed1 thus ?case by (fastforce elim: evals_cases intro: eval_evals.intros) next case ListRed2 thus ?case by (fastforce elim!: evals_cases eval_cases intro: eval_evals.intros eval_finalId) next case CastThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case BinOpThrow1 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case BinOpThrow2 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case LAssThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAccThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAssThrow1 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAssThrow2 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case CallThrowObj thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (CallThrowParams es vs e es' v M s s' e') have "P \ \Val v,s\ \ \Val v,s\" by (rule eval_evals.intros) moreover have es: "es = map Val vs @ throw e # es'" by fact have eval_e: "P \ \throw e,s\ \ \e',s'\" by fact then obtain xa where e': "e' = Throw xa" by (cases) (auto dest!: eval_final) with list_eval_Throw [OF eval_e] es have "P \ \es,s\ [\] \map Val vs @ Throw xa # es',s'\" by simp ultimately have "P \ \Val v\M(es),s\ \ \Throw xa,s'\" by (rule eval_evals.CallParamsThrow) thus ?case using e' by simp next case (InitBlockThrow V T v a s s' e') have "P \ \Throw a,s\ \ \e',s'\" by fact then obtain s': "s' = s" and e': "e' = Throw a" by cases (auto elim!:eval_cases) obtain h l where s: "s = (h,l)" by (cases s) have "P \ \{V:T :=Val v; Throw a},(h,l)\ \ \Throw a, (h, (l(V\v))(V:=l V))\" by(fastforce intro:eval_evals.intros) thus "P \ \{V:T := Val v; Throw a},s\ \ \e',s'\" using s s' e' by simp next case (BlockThrow V T a s s' e') have "P \ \Throw a, s\ \ \e',s'\" by fact then obtain s': "s' = s" and e': "e' = Throw a" by cases (auto elim!:eval_cases) obtain h l where s: "s=(h,l)" by (cases s) have "P \ \Throw a, (h,l(V:=None))\ \ \Throw a, (h,l(V:=None))\" by (simp add:eval_evals.intros eval_finalId) hence "P\\{V:T;Throw a},(h,l)\\\Throw a, (h,(l(V:=None))(V:=l V))\" by (rule eval_evals.Block) thus "P \ \{V:T; Throw a},s\ \ \e',s'\" using s s' e' by simp next case SeqThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case CondThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case ThrowThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) qed (*>*) (*<*) (* ... und wieder anschalten: *) declare split_paired_All [simp] split_paired_Ex [simp] (*>*) text \Its extension to \\*\:\ lemma extend_eval: assumes wf: "wwf_J_prog P" and reds: "P \ \e,s\ \* \e'',s''\" and eval_rest: "P \ \e'',s''\ \ \e',s'\" shows "P \ \e,s\ \ \e',s'\" (*<*) using reds eval_rest proof (induct rule: converse_rtrancl_induct2) case refl then show ?case by simp next case step show ?case using step extend_1_eval[OF wf step.hyps(1)] by simp qed (*>*) lemma extend_evals: assumes wf: "wwf_J_prog P" and reds: "P \ \es,s\ [\]* \es'',s''\" and eval_rest: "P \ \es'',s''\ [\] \es',s'\" shows "P \ \es,s\ [\] \es',s'\" (*<*) using reds eval_rest proof (induct rule: converse_rtrancl_induct2) case refl then show ?case by simp next case step show ?case using step extend_1_evals[OF wf step.hyps(1)] by simp qed (*>*) text \Finally, small step semantics can be simulated by big step semantics: \ theorem assumes wf: "wwf_J_prog P" shows small_by_big: "\P \ \e,s\ \* \e',s'\; final e'\ \ P \ \e,s\ \ \e',s'\" and "\P \ \es,s\ [\]* \es',s'\; finals es'\ \ P \ \es,s\ [\] \es',s'\" (*<*) proof - note wf moreover assume "P \ \e,s\ \* \e',s'\" moreover assume "final e'" then have "P \ \e',s'\ \ \e',s'\" by (rule eval_finalId) ultimately show "P \ \e,s\\\e',s'\" by (rule extend_eval) next note wf moreover assume "P \ \es,s\ [\]* \es',s'\" moreover assume "finals es'" then have "P \ \es',s'\ [\] \es',s'\" by (rule eval_finalsId) ultimately show "P \ \es,s\ [\] \es',s'\" by (rule extend_evals) qed (*>*) subsection "Equivalence" text\And now, the crowning achievement:\ corollary big_iff_small: "wwf_J_prog P \ P \ \e,s\ \ \e',s'\ = (P \ \e,s\ \* \e',s'\ \ final e')" (*<*)by(blast dest: big_by_small eval_final small_by_big)(*>*) end diff --git a/thys/JinjaDCI/Common/Exceptions.thy b/thys/JinjaDCI/Common/Exceptions.thy --- a/thys/JinjaDCI/Common/Exceptions.thy +++ b/thys/JinjaDCI/Common/Exceptions.thy @@ -1,188 +1,188 @@ (* Title: JinjaDCI/Common/Exceptions.thy Author: Gerwin Klein, Martin Strecker, Susannah Mansky Copyright 2002 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory Common/Exceptions.thy by Gerwin Klein and Martin Strecker *) section \ Exceptions \ theory Exceptions imports Objects begin definition ErrorCl :: "string" where "ErrorCl = ''Error''" definition ThrowCl :: "string" where "ThrowCl = ''Throwable''" definition NullPointer :: cname where "NullPointer \ ''NullPointer''" definition ClassCast :: cname where "ClassCast \ ''ClassCast''" definition OutOfMemory :: cname where "OutOfMemory \ ''OutOfMemory''" definition NoClassDefFoundError :: cname where "NoClassDefFoundError \ ''NoClassDefFoundError''" definition IncompatibleClassChangeError :: cname where "IncompatibleClassChangeError \ ''IncompatibleClassChangeError''" definition NoSuchFieldError :: cname where "NoSuchFieldError \ ''NoSuchFieldError''" definition NoSuchMethodError :: cname where "NoSuchMethodError \ ''NoSuchMethodError''" definition sys_xcpts :: "cname set" where "sys_xcpts \ {NullPointer, ClassCast, OutOfMemory, NoClassDefFoundError, IncompatibleClassChangeError, NoSuchFieldError, NoSuchMethodError}" definition addr_of_sys_xcpt :: "cname \ addr" where "addr_of_sys_xcpt s \ if s = NullPointer then 0 else if s = ClassCast then 1 else if s = OutOfMemory then 2 else if s = NoClassDefFoundError then 3 else if s = IncompatibleClassChangeError then 4 else if s = NoSuchFieldError then 5 else if s = NoSuchMethodError then 6 else undefined" lemmas sys_xcpts_defs = NullPointer_def ClassCast_def OutOfMemory_def NoClassDefFoundError_def IncompatibleClassChangeError_def NoSuchFieldError_def NoSuchMethodError_def lemma Start_nsys_xcpts: "Start \ sys_xcpts" by(simp add: Start_def sys_xcpts_def sys_xcpts_defs) lemma Start_nsys_xcpts1 [simp]: "Start \ NullPointer" "Start \ ClassCast" "Start \ OutOfMemory" "Start \ NoClassDefFoundError" "Start \ IncompatibleClassChangeError" "Start \ NoSuchFieldError" "Start \ NoSuchMethodError" using Start_nsys_xcpts by(auto simp: sys_xcpts_def) lemma Start_nsys_xcpts2 [simp]: "NullPointer \ Start" "ClassCast \ Start" "OutOfMemory \ Start" "NoClassDefFoundError \ Start" "IncompatibleClassChangeError \ Start" "NoSuchFieldError \ Start" "NoSuchMethodError \ Start" using Start_nsys_xcpts by(auto simp: sys_xcpts_def dest: sym) definition start_heap :: "'c prog \ heap" where - "start_heap G \ Map.empty (addr_of_sys_xcpt NullPointer \ blank G NullPointer) - (addr_of_sys_xcpt ClassCast \ blank G ClassCast) - (addr_of_sys_xcpt OutOfMemory \ blank G OutOfMemory) - (addr_of_sys_xcpt NoClassDefFoundError \ blank G NoClassDefFoundError) - (addr_of_sys_xcpt IncompatibleClassChangeError \ blank G IncompatibleClassChangeError) - (addr_of_sys_xcpt NoSuchFieldError \ blank G NoSuchFieldError) - (addr_of_sys_xcpt NoSuchMethodError \ blank G NoSuchMethodError)" + "start_heap G \ Map.empty (addr_of_sys_xcpt NullPointer \ blank G NullPointer, + addr_of_sys_xcpt ClassCast \ blank G ClassCast, + addr_of_sys_xcpt OutOfMemory \ blank G OutOfMemory, + addr_of_sys_xcpt NoClassDefFoundError \ blank G NoClassDefFoundError, + addr_of_sys_xcpt IncompatibleClassChangeError \ blank G IncompatibleClassChangeError, + addr_of_sys_xcpt NoSuchFieldError \ blank G NoSuchFieldError, + addr_of_sys_xcpt NoSuchMethodError \ blank G NoSuchMethodError)" definition preallocated :: "heap \ bool" where "preallocated h \ \C \ sys_xcpts. \fs. h(addr_of_sys_xcpt C) = Some (C,fs)" subsection "System exceptions" lemma sys_xcpts_incl [simp]: "NullPointer \ sys_xcpts \ OutOfMemory \ sys_xcpts \ ClassCast \ sys_xcpts \ NoClassDefFoundError \ sys_xcpts \ IncompatibleClassChangeError \ sys_xcpts \ NoSuchFieldError \ sys_xcpts \ NoSuchMethodError \ sys_xcpts" (*<*)by(simp add: sys_xcpts_def)(*>*) lemma sys_xcpts_cases [consumes 1, cases set]: "\ C \ sys_xcpts; P NullPointer; P OutOfMemory; P ClassCast; P NoClassDefFoundError; P IncompatibleClassChangeError; P NoSuchFieldError; P NoSuchMethodError \ \ P C" (*<*)by (auto simp: sys_xcpts_def)(*>*) subsection "Starting heap" lemma start_heap_sys_xcpts: assumes "C \ sys_xcpts" shows "start_heap P (addr_of_sys_xcpt C) = Some(blank P C)" by(rule sys_xcpts_cases[OF assms]) (auto simp add: start_heap_def sys_xcpts_def addr_of_sys_xcpt_def sys_xcpts_defs) lemma start_heap_classes: "start_heap P a = Some(C,fs) \ C \ sys_xcpts" by(simp add: start_heap_def blank_def split: if_split_asm) lemma start_heap_nStart: "start_heap P a = Some obj \ fst(obj) \ Start" by(cases obj, auto dest!: start_heap_classes simp: Start_nsys_xcpts) subsection "@{term preallocated}" lemma preallocated_dom [simp]: "\ preallocated h; C \ sys_xcpts \ \ addr_of_sys_xcpt C \ dom h" (*<*)by (fastforce simp:preallocated_def dom_def)(*>*) lemma preallocatedD: "\ preallocated h; C \ sys_xcpts \ \ \fs. h(addr_of_sys_xcpt C) = Some (C, fs)" (*<*)by(auto simp: preallocated_def sys_xcpts_def)(*>*) lemma preallocatedE [elim?]: "\ preallocated h; C \ sys_xcpts; \fs. h(addr_of_sys_xcpt C) = Some(C,fs) \ P h C\ \ P h C" (*<*)by (fast dest: preallocatedD)(*>*) lemma cname_of_xcp [simp]: "\ preallocated h; C \ sys_xcpts \ \ cname_of h (addr_of_sys_xcpt C) = C" (*<*)by (auto elim: preallocatedE)(*>*) lemma typeof_ClassCast [simp]: "preallocated h \ typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt ClassCast)) = Some(Class ClassCast)" (*<*)by (auto elim: preallocatedE)(*>*) lemma typeof_OutOfMemory [simp]: "preallocated h \ typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt OutOfMemory)) = Some(Class OutOfMemory)" (*<*)by (auto elim: preallocatedE)(*>*) lemma typeof_NullPointer [simp]: "preallocated h \ typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NullPointer)) = Some(Class NullPointer)" (*<*)by (auto elim: preallocatedE)(*>*) lemma typeof_NoClassDefFoundError [simp]: "preallocated h \ typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NoClassDefFoundError)) = Some(Class NoClassDefFoundError)" (*<*)by (auto elim: preallocatedE)(*>*) lemma typeof_IncompatibleClassChangeError [simp]: "preallocated h \ typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt IncompatibleClassChangeError)) = Some(Class IncompatibleClassChangeError)" (*<*)by (auto elim: preallocatedE)(*>*) lemma typeof_NoSuchFieldError [simp]: "preallocated h \ typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NoSuchFieldError)) = Some(Class NoSuchFieldError)" (*<*)by (auto elim: preallocatedE)(*>*) lemma typeof_NoSuchMethodError [simp]: "preallocated h \ typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NoSuchMethodError)) = Some(Class NoSuchMethodError)" (*<*)by (auto elim: preallocatedE)(*>*) lemma preallocated_hext: "\ preallocated h; h \ h' \ \ preallocated h'" (*<*)by (simp add: preallocated_def hext_def)(*>*) (*<*) lemmas preallocated_upd_obj = preallocated_hext [OF _ hext_upd_obj] lemmas preallocated_new = preallocated_hext [OF _ hext_new] (*>*) lemma preallocated_start: "preallocated (start_heap P)" by(auto simp: start_heap_sys_xcpts blank_def preallocated_def) end diff --git a/thys/JinjaDCI/Compiler/Hidden.thy b/thys/JinjaDCI/Compiler/Hidden.thy --- a/thys/JinjaDCI/Compiler/Hidden.thy +++ b/thys/JinjaDCI/Compiler/Hidden.thy @@ -1,52 +1,52 @@ theory Hidden imports "List-Index.List_Index" begin definition hidden :: "'a list \ nat \ bool" where "hidden xs i \ i < size xs \ xs!i \ set(drop (i+1) xs)" lemma hidden_last_index: "x \ set xs \ hidden (xs @ [x]) (last_index xs x)" by(auto simp add: hidden_def nth_append rev_nth[symmetric] dest: last_index_less[OF _ le_refl]) lemma hidden_inacc: "hidden xs i \ last_index xs x \ i" by(auto simp add: hidden_def last_index_drop last_index_less_size_conv) lemma [simp]: "hidden xs i \ hidden (xs@[x]) i" by(auto simp add:hidden_def nth_append) lemma fun_upds_apply: "(m(xs[\]ys)) x = (let xs' = take (size ys) xs in if x \ set xs' then Some(ys ! last_index xs' x) else m x)" proof(induct xs arbitrary: m ys) case Nil then show ?case by(simp add: Let_def) next case Cons show ?case proof(cases ys) case Nil then show ?thesis by(simp add:Let_def) next case Cons': Cons then show ?thesis using Cons by(simp add: Let_def last_index_Cons) qed qed lemma map_upds_apply_eq_Some: "((m(xs[\]ys)) x = Some y) = (let xs' = take (size ys) xs in if x \ set xs' then ys ! last_index xs' x = y else m x = Some y)" by(simp add:fun_upds_apply Let_def) lemma map_upds_upd_conv_last_index: "\x \ set xs; size xs \ size ys \ - \ m(xs[\]ys)(x\y) = m(xs[\]ys[last_index xs x := y])" + \ m(xs[\]ys, x\y) = m(xs[\]ys[last_index xs x := y])" by(rule ext) (simp add:fun_upds_apply eq_sym_conv Let_def) end diff --git a/thys/JinjaDCI/J/Equivalence.thy b/thys/JinjaDCI/J/Equivalence.thy --- a/thys/JinjaDCI/J/Equivalence.thy +++ b/thys/JinjaDCI/J/Equivalence.thy @@ -1,4191 +1,4191 @@ (* Title: JinjaDCI/J/Equivalence.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory J/Equivalence.thy by Tobias Nipkow *) section \ Equivalence of Big Step and Small Step Semantics \ theory Equivalence imports TypeSafe WWellForm begin subsection\Small steps simulate big step\ subsubsection "Init" text "The reduction of initialization expressions doesn't touch or use their on-hold expressions (the subexpression to the right of @{text \}) until the initialization operation completes. This function is used to prove this and related properties. It is then used for general reduction of initialization expressions separately from their on-hold expressions by using the on-hold expression @{term unit}, then putting the real on-hold expression back in at the end." fun init_switch :: "'a exp \ 'a exp \ 'a exp" where "init_switch (INIT C (Cs,b) \ e\<^sub>i) e = (INIT C (Cs,b) \ e)" | "init_switch (RI(C,e');Cs \ e\<^sub>i) e = (RI(C,e');Cs \ e)" | "init_switch e' e = e'" fun INIT_class :: "'a exp \ cname option" where "INIT_class (INIT C (Cs,b) \ e) = (if C = last (C#Cs) then Some C else None)" | "INIT_class (RI(C,e\<^sub>0);Cs \ e) = Some (last (C#Cs))" | "INIT_class _ = None" lemma init_red_init: "\ init_exp_of e\<^sub>0 = \e\; P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\ \ \ (init_exp_of e\<^sub>1 = \e\ \ (INIT_class e\<^sub>0 = \C\ \ INIT_class e\<^sub>1 = \C\)) \ (e\<^sub>1 = e \ b\<^sub>1 = icheck P (the(INIT_class e\<^sub>0)) e) \ (\a. e\<^sub>1 = throw a)" by(erule red.cases, auto) lemma init_exp_switch[simp]: "init_exp_of e\<^sub>0 = \e\ \ init_exp_of (init_switch e\<^sub>0 e') = \e'\" by(cases e\<^sub>0, simp_all) lemma init_red_sync: "\ P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\; init_exp_of e\<^sub>0 = \e\; e\<^sub>1 \ e \ \ (\e'. P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \ \init_switch e\<^sub>1 e',s\<^sub>1,b\<^sub>1\)" proof(induct rule: red.cases) qed(simp_all add: red_reds.intros) lemma init_red_sync_end: "\ P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\; init_exp_of e\<^sub>0 = \e\; e\<^sub>1 = e; throw_of e = None \ \ (\e'. \sub_RI e' \ P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \ \e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\)" proof(induct rule: red.cases) qed(simp_all add: red_reds.intros) lemma init_reds_sync_unit': "\ P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \Val v',s\<^sub>1,b\<^sub>1\; init_exp_of e\<^sub>0 = \unit\; INIT_class e\<^sub>0 = \C\ \ \ (\e'. \sub_RI e' \ P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \* \e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\)" proof(induct rule:converse_rtrancl_induct3) case refl then show ?case by simp next case (step e0 s0 b0 e1 s1 b1) have "(init_exp_of e1 = \unit\ \ (INIT_class e0 = \C\ \ INIT_class e1 = \C\)) \ (e1 = unit \ b1 = icheck P (the(INIT_class e0)) unit) \ (\a. e1 = throw a)" using init_red_init[OF step.prems(1) step.hyps(1)] by simp then show ?case proof(rule disjE) assume assm: "init_exp_of e1 = \unit\ \ (INIT_class e0 = \C\ \ INIT_class e1 = \C\)" then have red: "P \ \init_switch e0 e',s0,b0\ \ \init_switch e1 e',s1,b1\" using init_red_sync[OF step.hyps(1) step.prems(1)] by simp have reds: "P \ \init_switch e1 e',s1,b1\ \* \e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\" using step.hyps(3)[OF _ _ step.prems(3)] assm step.prems(2) by simp show ?thesis by(rule converse_rtrancl_into_rtrancl[OF red reds]) next assume "(e1 = unit \ b1 = icheck P (the(INIT_class e0)) unit) \ (\a. e1 = throw a)" then show ?thesis proof(rule disjE) assume assm: "e1 = unit \ b1 = icheck P (the(INIT_class e0)) unit" then have e1: "e1 = unit" by simp have sb: "s1 = s\<^sub>1" "b1 = b\<^sub>1" using reds_final_same[OF step.hyps(2)] assm by(simp_all add: final_def) then have step': "P \ \init_switch e0 e',s0,b0\ \ \e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\" using init_red_sync_end[OF step.hyps(1) step.prems(1) e1 _ step.prems(3)] by auto then have "P \ \init_switch e0 e',s0,b0\ \* \e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\" using r_into_rtrancl by auto then show ?thesis using step assm sb by simp next assume "\a. e1 = throw a" then obtain a where "e1 = throw a" by clarsimp then have tof: "throw_of e1 = \a\" by simp then show ?thesis using reds_throw[OF step.hyps(2) tof] by simp qed qed qed lemma init_reds_sync_unit_throw': "\ P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\; init_exp_of e\<^sub>0 = \unit\ \ \ (\e'. P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\)" proof(induct rule:converse_rtrancl_induct3) case refl then show ?case by simp next case (step e0 s0 b0 e1 s1 b1) have "init_exp_of e1 = \unit\ \ (\C. INIT_class e0 = \C\ \ INIT_class e1 = \C\) \ e1 = unit \ b1 = icheck P (the (INIT_class e0)) unit \ (\a. e1 = throw a)" using init_red_init[OF step.prems(1) step.hyps(1)] by auto then show ?case proof(rule disjE) assume assm: "init_exp_of e1 = \unit\ \ (\C. INIT_class e0 = \C\ \ INIT_class e1 = \C\)" then have "P \ \init_switch e0 e',s0,b0\ \ \init_switch e1 e',s1,b1\" using step init_red_sync[OF step.hyps(1) step.prems] by simp then show ?thesis using step assm by (meson converse_rtrancl_into_rtrancl) next assume "e1 = unit \ b1 = icheck P (the (INIT_class e0)) unit \ (\a. e1 = throw a)" then show ?thesis proof(rule disjE) assume "e1 = unit \ b1 = icheck P (the (INIT_class e0)) unit" then show ?thesis using step using final_def reds_final_same by blast next assume assm: "\a. e1 = throw a" then have "P \ \init_switch e0 e',s0,b0\ \ \e1,s1,b1\" using init_red_sync[OF step.hyps(1) step.prems] by clarsimp then show ?thesis using step by simp qed qed qed lemma init_reds_sync_unit: assumes "P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \Val v',s\<^sub>1,b\<^sub>1\" and "init_exp_of e\<^sub>0 = \unit\" and "INIT_class e\<^sub>0 = \C\" and "\sub_RI e'" shows "P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \* \e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\" using init_reds_sync_unit'[OF assms] by clarsimp lemma init_reds_sync_unit_throw: assumes "P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" and "init_exp_of e\<^sub>0 = \unit\" shows "P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" using init_reds_sync_unit_throw'[OF assms] by clarsimp \ \ init reds lemmas \ lemma InitSeqReds: assumes "P \ \INIT C ([C],b) \ unit,s\<^sub>0,b\<^sub>0\ \* \Val v',s\<^sub>1,b\<^sub>1\" and "P \ \e,s\<^sub>1,icheck P C e\ \* \e\<^sub>2,s\<^sub>2,b\<^sub>2\" and "\sub_RI e" shows "P \ \INIT C ([C],b) \ e,s\<^sub>0,b\<^sub>0\ \* \e\<^sub>2,s\<^sub>2,b\<^sub>2\" using assms proof - have "P \ \init_switch (INIT C ([C],b) \ unit) e,s\<^sub>0,b\<^sub>0\ \* \e,s\<^sub>1,icheck P C e\" using init_reds_sync_unit[OF assms(1) _ _ assms(3)] by simp then show ?thesis using assms(2) by simp qed lemma InitSeqThrowReds: assumes "P \ \INIT C ([C],b) \ unit,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" shows "P \ \INIT C ([C],b) \ e,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" using assms proof - have "P \ \init_switch (INIT C ([C],b) \ unit) e,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" using init_reds_sync_unit_throw[OF assms(1)] by simp then show ?thesis by simp qed lemma InitNoneReds: "\ sh C = None; P \ \INIT C' (C # Cs,False) \ e,(h, l, sh(C \ (sblank P C, Prepared))),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*)by(rule InitNoneRed[THEN converse_rtrancl_into_rtrancl])(*>*) lemma InitDoneReds: "\ sh C = Some(sfs,Done); P \ \INIT C' (Cs,True) \ e,(h,l,sh),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*)by(rule RedInitDone[THEN converse_rtrancl_into_rtrancl])(*>*) lemma InitProcessingReds: "\ sh C = Some(sfs,Processing); P \ \INIT C' (Cs,True) \ e,(h,l,sh),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*)by(rule RedInitProcessing[THEN converse_rtrancl_into_rtrancl])(*>*) lemma InitErrorReds: "\ sh C = Some(sfs,Error); P \ \RI (C,THROW NoClassDefFoundError);Cs \ e,(h,l,sh),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*)by(rule RedInitError[THEN converse_rtrancl_into_rtrancl])(*>*) lemma InitObjectReds: "\ sh C = Some(sfs,Prepared); C = Object; sh' = sh(C \ (sfs,Processing)); P \ \INIT C' (C#Cs,True) \ e,(h,l,sh'),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*)by(rule InitObjectRed[THEN converse_rtrancl_into_rtrancl])(*>*) lemma InitNonObjectReds: "\ sh C = Some(sfs,Prepared); C \ Object; class P C = Some (D,r); sh' = sh(C \ (sfs,Processing)); P \ \INIT C' (D#C#Cs,False) \ e,(h,l,sh'),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*)by(rule InitNonObjectSuperRed[THEN converse_rtrancl_into_rtrancl])(*>*) lemma RedsInitRInit: "P \ \RI (C,C\\<^sub>sclinit([]));Cs \ e,(h,l,sh),b\ \* \e',s',b'\ \ P \ \INIT C' (C#Cs,True) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*)by(rule RedInitRInit[THEN converse_rtrancl_into_rtrancl])(*>*) lemmas rtrancl_induct3 = rtrancl_induct[of "(ax,ay,az)" "(bx,by,bz)", split_format (complete), consumes 1, case_names refl step] lemma RInitReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \RI (C,e);Cs \ e\<^sub>0, s, b\ \* \RI (C,e');Cs \ e\<^sub>0, s', b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) RInitRed[OF step(2)]]) qed (*>*) lemma RedsRInit: "\ P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\; sh\<^sub>1 C = Some (sfs, i); sh\<^sub>2 = sh\<^sub>1(C \ (sfs,Done)); C' = last(C#Cs); P \ \INIT C' (Cs,True) \ e,(h\<^sub>1,l\<^sub>1,sh\<^sub>2),b\<^sub>1\ \* \e',s',b'\ \ \ P \ \RI (C, e\<^sub>0);Cs \ e,s\<^sub>0,b\<^sub>0\ \* \e',s',b'\" (*<*) by(rule rtrancl_trans[OF RInitReds RedRInit[THEN converse_rtrancl_into_rtrancl]]) (*>*) lemma RInitInitThrowReds: "\ P \ \e,s,b\ \* \Throw a,(h',l',sh'),b'\; sh' C = Some (sfs, i); sh'' = sh'(C \ (sfs,Error)); P \ \RI (D,Throw a);Cs \ e\<^sub>0, (h',l',sh''),b'\ \* \e\<^sub>1,s\<^sub>1,b\<^sub>1\ \ \ P \ \RI (C, e);D#Cs \ e\<^sub>0,s,b\ \* \e\<^sub>1,s\<^sub>1,b\<^sub>1\" (*<*) by(rule rtrancl_trans[OF RInitReds RInitInitThrow[THEN converse_rtrancl_into_rtrancl]]) (*>*) lemma RInitThrowReds: "\ P \ \e,s,b\ \* \Throw a, (h',l',sh'),b'\; sh' C = Some(sfs, i); sh'' = sh'(C \ (sfs, Error)) \ \ P \ \RI (C,e);Nil \ e\<^sub>0,s,b\ \* \Throw a, (h',l',sh''),b'\" (*<*)by(rule RInitReds[THEN rtrancl_into_rtrancl, OF _ RInitThrow])(*>*) subsubsection "New" lemma NewInitDoneReds: "\ sh C = Some (sfs, Done); new_Addr h = Some a; P \ C has_fields FDTs; h' = h(a\blank P C) \ \ P \ \new C,(h,l,sh),False\ \* \addr a,(h',l,sh),False\" (*<*) by(rule NewInitDoneRed[THEN converse_rtrancl_into_rtrancl, OF _ RedNew[THEN r_into_rtrancl]]) (*>*) lemma NewInitDoneReds2: "\ sh C = Some (sfs, Done); new_Addr h = None; is_class P C \ \ P \ \new C,(h,l,sh),False\ \* \THROW OutOfMemory, (h,l,sh), False\" (*<*) by(rule NewInitDoneRed[THEN converse_rtrancl_into_rtrancl, OF _ RedNewFail[THEN r_into_rtrancl]]) (*>*) lemma NewInitReds: assumes nDone: "\sfs. shp s C = Some (sfs, Done)" and INIT_steps: "P \ \INIT C ([C],False) \ unit,s,False\ \* \Val v',(h',l',sh'),b'\" and Addr: "new_Addr h' = Some a" and has: "P \ C has_fields FDTs" and h'': "h'' = h'(a\blank P C)" and cls_C: "is_class P C" shows "P \ \new C,s,False\ \* \addr a,(h'',l',sh'),False\" (*<*)(is "(?a, ?c) \ (red P)\<^sup>*") proof - let ?b = "(INIT C ([C],False) \ new C,s,False)" let ?b' = "(new C,(h', l', sh'),icheck P C (new C::expr))" have b'c: "(?b', ?c) \ (red P)\<^sup>*" using RedNew[OF Addr has h'', THEN r_into_rtrancl] by simp obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp have "(?a, ?b) \ (red P)\<^sup>*" using NewInitRed[OF _ cls_C] nDone by fastforce also have "(?b, ?c) \ (red P)\<^sup>*" by(rule InitSeqReds[OF INIT_steps b'c]) simp ultimately show ?thesis by simp qed (*>*) lemma NewInitOOMReds: assumes nDone: "\sfs. shp s C = Some (sfs, Done)" and INIT_steps: "P \ \INIT C ([C],False) \ unit,s,False\ \* \Val v',(h',l',sh'),b'\" and Addr: "new_Addr h' = None" and cls_C: "is_class P C" shows "P \ \new C,s,False\ \* \THROW OutOfMemory,(h',l',sh'),False\" (*<*)(is "(?a, ?c) \ (red P)\<^sup>*") proof - let ?b = "(INIT C ([C],False) \ new C,s,False)" let ?b' = "(new C,(h', l', sh'),icheck P C (new C::expr))" have b'c: "(?b', ?c) \ (red P)\<^sup>*" using RedNewFail[OF Addr cls_C, THEN r_into_rtrancl] by simp obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp have "(?a, ?b) \ (red P)\<^sup>*" using NewInitRed[OF _ cls_C] nDone by fastforce also have "(?b, ?c) \ (red P)\<^sup>*" by(rule InitSeqReds[OF INIT_steps b'c]) simp ultimately show ?thesis by simp qed (*>*) lemma NewInitThrowReds: assumes nDone: "\sfs. shp s C = Some (sfs, Done)" and cls_C: "is_class P C" and INIT_steps: "P \ \INIT C ([C],False) \ unit,s,False\ \* \throw a,s',b'\" shows "P \ \new C,s,False\ \* \throw a,s',b'\" (*<*)(is "(?a, ?c) \ (red P)\<^sup>*") proof - let ?b = "(INIT C ([C],False) \ new C,s,False)" obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp have "(?a, ?b) \ (red P)\<^sup>*" using NewInitRed[OF _ cls_C] nDone by fastforce also have "(?b, ?c) \ (red P)\<^sup>*" using InitSeqThrowReds[OF INIT_steps] by simp ultimately show ?thesis by simp qed (*>*) subsubsection "Cast" lemma CastReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \Cast C e,s,b\ \* \Cast C e',s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) CastRed[OF step(2)]]) qed (*>*) lemma CastRedsNull: "P \ \e,s,b\ \* \null,s',b'\ \ P \ \Cast C e,s,b\ \* \null,s',b'\" (*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastNull])(*>*) lemma CastRedsAddr: "\ P \ \e,s,b\ \* \addr a,s',b'\; hp s' a = Some(D,fs); P \ D \\<^sup>* C \ \ P \ \Cast C e,s,b\ \* \addr a,s',b'\" (*<*)by(cases s', simp) (rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCast])(*>*) lemma CastRedsFail: "\ P \ \e,s,b\ \* \addr a,s',b'\; hp s' a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \Cast C e,s,b\ \* \THROW ClassCast,s',b'\" (*<*)by(cases s', simp) (rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastFail])(*>*) lemma CastRedsThrow: "\ P \ \e,s,b\ \* \throw a,s',b'\ \ \ P \ \Cast C e,s,b\ \* \throw a,s',b'\" (*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CastThrow])(*>*) subsubsection "LAss" lemma LAssReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \ V:=e,s,b\ \* \ V:=e',s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) LAssRed[OF step(2)]]) qed (*>*) lemma LAssRedsVal: "\ P \ \e,s,b\ \* \Val v,(h',l',sh'),b'\ \ \ P \ \ V:=e,s,b\ \* \unit,(h',l'(V\v),sh'),b'\" (*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ RedLAss])(*>*) lemma LAssRedsThrow: "\ P \ \e,s,b\ \* \throw a,s',b'\ \ \ P \ \ V:=e,s,b\ \* \throw a,s',b'\" (*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ red_reds.LAssThrow])(*>*) subsubsection "BinOp" lemma BinOp1Reds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \ e \bop\ e\<^sub>2, s,b\ \* \e' \bop\ e\<^sub>2, s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed1[OF step(2)]]) qed (*>*) lemma BinOp2Reds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \(Val v) \bop\ e, s,b\ \* \(Val v) \bop\ e', s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed2[OF step(2)]]) qed (*>*) lemma BinOpRedsVal: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1,b\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v\<^sub>2,s\<^sub>2,b\<^sub>2\" and op: "binop(bop,v\<^sub>1,v\<^sub>2) = Some v" shows "P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \Val v,s\<^sub>2,b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(Val v\<^sub>1 \bop\ e\<^sub>2, s\<^sub>1,b\<^sub>1)" let ?y' = "(Val v\<^sub>1 \bop\ Val v\<^sub>2, s\<^sub>2,b\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" by(rule RedBinOp[OF op]) ultimately show ?thesis by simp qed (*>*) lemma BinOpRedsThrow1: "P \ \e,s,b\ \* \throw e',s',b'\ \ P \ \e \bop\ e\<^sub>2, s,b\ \* \throw e', s',b'\" (*<*)by(rule BinOp1Reds[THEN rtrancl_into_rtrancl, OF _ red_reds.BinOpThrow1])(*>*) lemma BinOpRedsThrow2: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1,b\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \throw e,s\<^sub>2,b\<^sub>2\" shows "P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(Val v\<^sub>1 \bop\ e\<^sub>2, s\<^sub>1,b\<^sub>1)" let ?y' = "(Val v\<^sub>1 \bop\ throw e, s\<^sub>2,b\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" by(rule red_reds.BinOpThrow2) ultimately show ?thesis by simp qed (*>*) subsubsection "FAcc" lemma FAccReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \e\F{D}, s,b\ \* \e'\F{D}, s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) FAccRed[OF step(2)]]) qed (*>*) lemma FAccRedsVal: "\ P \ \e,s,b\ \* \addr a,s',b'\; hp s' a = Some(C,fs); fs(F,D) = Some v; P \ C has F,NonStatic:t in D \ \ P \ \e\F{D},s,b\ \* \Val v,s',b'\" (*<*)by(cases s',simp) (rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAcc])(*>*) lemma FAccRedsNull: "P \ \e,s,b\ \* \null,s',b'\ \ P \ \e\F{D},s,b\ \* \THROW NullPointer,s',b'\" (*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccNull])(*>*) lemma FAccRedsNone: "\ P \ \e,s,b\ \* \addr a,s',b'\; hp s' a = Some(C,fs); \(\b t. P \ C has F,b:t in D) \ \ P \ \e\F{D},s,b\ \* \THROW NoSuchFieldError,s',b'\" (*<*)by(cases s',simp) (auto intro: FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccNone])(*>*) lemma FAccRedsStatic: "\ P \ \e,s,b\ \* \addr a,s',b'\; hp s' a = Some(C,fs); P \ C has F,Static:t in D \ \ P \ \e\F{D},s,b\ \* \THROW IncompatibleClassChangeError,s',b'\" (*<*)by(cases s',simp) (rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccStatic])(*>*) lemma FAccRedsThrow: "P \ \e,s,b\ \* \throw a,s',b'\ \ P \ \e\F{D},s,b\ \* \throw a,s',b'\" (*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ red_reds.FAccThrow])(*>*) subsubsection "SFAcc" lemma SFAccReds: "\ P \ C has F,Static:t in D; shp s D = Some(sfs,Done); sfs F = Some v \ \ P \ \C\\<^sub>sF{D},s,True\ \* \Val v,s,False\" (*<*)by(cases s,simp) (rule RedSFAcc[THEN r_into_rtrancl])(*>*) lemma SFAccRedsNone: "\(\b t. P \ C has F,b:t in D) \ P \ \C\\<^sub>sF{D},s,b\ \* \THROW NoSuchFieldError,s,False\" (*<*)by(cases s,simp) (auto intro: RedSFAccNone[THEN r_into_rtrancl])(*>*) lemma SFAccRedsNonStatic: "P \ C has F,NonStatic:t in D \ P \ \C\\<^sub>sF{D},s,b\ \* \THROW IncompatibleClassChangeError,s,False\" (*<*)by(cases s,simp) (rule RedSFAccNonStatic[THEN r_into_rtrancl])(*>*) lemma SFAccInitDoneReds: assumes cF: "P \ C has F,Static:t in D" and shp: "shp s D = Some (sfs,Done)" and sfs: "sfs F = Some v" shows "P \ \C\\<^sub>sF{D}, s, b\ \* \Val v, s, False\" (*<*)(is "(?a, ?c) \ (red P)\<^sup>*") proof - obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp show ?thesis proof(cases b) case True then show ?thesis using assms by simp (rule RedSFAcc[THEN r_into_rtrancl]) next case False let ?b = "(C\\<^sub>sF{D},s,True)" have "(?a, ?b) \ (red P)\<^sup>*" using False SFAccInitDoneRed[where sh="shp s", OF cF shp] by fastforce also have "(?b, ?c) \ (red P)\<^sup>*" by(rule SFAccReds[OF assms]) ultimately show ?thesis by simp qed qed (*>*) lemma SFAccInitReds: assumes cF: "P \ C has F,Static:t in D" and nDone: "\sfs. shp s D = Some (sfs,Done)" and INIT_steps: "P \ \INIT D ([D],False) \ unit,s,False\ \* \Val v',s',b'\" and shp': "shp s' D = Some (sfs,i)" and sfs: "sfs F = Some v" shows "P \ \C\\<^sub>sF{D},s,False\ \* \Val v,s',False\" (*<*)(is "(?a, ?c) \ (red P)\<^sup>*") proof - let ?b = "(INIT D ([D],False) \ C\\<^sub>sF{D},s,False)" let ?b' = "(C\\<^sub>sF{D},s',icheck P D (C\\<^sub>sF{D}::expr))" obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp obtain h' l' sh' where [simp]: "s' = (h', l', sh')" by(cases s') simp have icheck: "icheck P D (C\\<^sub>sF{D}) = True" using cF by fastforce then have b'c: "(?b', ?c) \ (red P)\<^sup>*" using RedSFAcc[THEN r_into_rtrancl, OF cF] shp' sfs by simp have "(?a, ?b) \ (red P)" using SFAccInitRed[OF cF] nDone by simp also have "(?b, ?c) \ (red P)\<^sup>*" by(rule InitSeqReds[OF INIT_steps b'c]) simp ultimately show ?thesis by simp qed (*>*) lemma SFAccInitThrowReds: "\ P \ C has F,Static:t in D; \sfs. shp s D = Some (sfs,Done); P \ \INIT D ([D],False) \ unit,s,False\ \* \throw a,s',b'\ \ \ P \ \C\\<^sub>sF{D},s,False\ \* \throw a,s',b'\" (*<*) by(cases s, simp) (auto elim: converse_rtrancl_into_rtrancl[OF SFAccInitRed InitSeqThrowReds]) (*>*) subsubsection "FAss" lemma FAssReds1: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \e\F{D}:=e\<^sub>2, s,b\ \* \e'\F{D}:=e\<^sub>2, s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) FAssRed1[OF step(2)]]) qed (*>*) lemma FAssReds2: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \Val v\F{D}:=e, s,b\ \* \Val v\F{D}:=e', s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) FAssRed2[OF step(2)]]) qed (*>*) lemma FAssRedsVal: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \addr a,s\<^sub>1,b\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" and cF: "P \ C has F,NonStatic:t in D" and hC: "Some(C,fs) = h\<^sub>2 a" shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\ \* \unit, (h\<^sub>2(a\(C,fs((F,D)\v))),l\<^sub>2,sh\<^sub>2),b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(addr a\F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)" let ?y' = "(addr a\F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" using RedFAss[OF cF] hC by simp ultimately show ?thesis by simp qed (*>*) lemma FAssRedsNull: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \null,s\<^sub>1,b\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,s\<^sub>2,b\<^sub>2\" shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\ \* \THROW NullPointer, s\<^sub>2, b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(null\F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)" let ?y' = "(null\F{D}:=Val v::expr,s\<^sub>2,b\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" by(rule RedFAssNull) ultimately show ?thesis by simp qed (*>*) lemma FAssRedsThrow1: "P \ \e,s,b\ \* \throw e',s',b'\ \ P \ \e\F{D}:=e\<^sub>2, s,b\ \* \throw e', s',b'\" (*<*)by(rule FAssReds1[THEN rtrancl_into_rtrancl, OF _ red_reds.FAssThrow1])(*>*) lemma FAssRedsThrow2: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v,s\<^sub>1,b\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \throw e,s\<^sub>2,b\<^sub>2\" shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(Val v\F{D}:=e\<^sub>2,s\<^sub>1,b\<^sub>1)" let ?y' = "(Val v\F{D}:=throw e,s\<^sub>2,b\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" by(rule red_reds.FAssThrow2) ultimately show ?thesis by simp qed (*>*) lemma FAssRedsNone: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \addr a,s\<^sub>1,b\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" and hC: "h\<^sub>2 a = Some(C,fs)" and ncF: "\(\b t. P \ C has F,b:t in D)" shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\ \* \THROW NoSuchFieldError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(addr a\F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)" let ?y' = "(addr a\F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" using RedFAssNone[OF _ ncF] hC by simp ultimately show ?thesis by simp qed (*>*) lemma FAssRedsStatic: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \addr a,s\<^sub>1,b\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" and hC: "h\<^sub>2 a = Some(C,fs)" and cF_Static: "P \ C has F,Static:t in D" shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\ \* \THROW IncompatibleClassChangeError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(addr a\F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)" let ?y' = "(addr a\F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) also have "(?y', ?z) \ (red P)" using RedFAssStatic[OF _ cF_Static] hC by simp ultimately show ?thesis by simp qed (*>*) subsubsection "SFAss" lemma SFAssReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \C\\<^sub>sF{D}:=e,s,b\ \* \C\\<^sub>sF{D}:=e',s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) SFAssRed[OF step(2)]]) qed (*>*) lemma SFAssRedsVal: assumes e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" and cF: "P \ C has F,Static:t in D" and shD: "sh\<^sub>2 D = \(sfs,Done)\" shows "P \ \C\\<^sub>sF{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\ \* \unit, (h\<^sub>2,l\<^sub>2,sh\<^sub>2(D\(sfs(F\v), Done))),False\" (*<*)(is "(?a, ?c) \ (red P)\<^sup>*") proof - let ?b = "(C\\<^sub>sF{D}:=Val v::expr, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2)" have "(?a, ?b) \ (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps]) also have "(?b, ?c) \ (red P)\<^sup>*" proof(cases b\<^sub>2) case True then show ?thesis using RedSFAss[THEN r_into_rtrancl, OF cF] shD by simp next case False let ?b' = "(C\\<^sub>sF{D}:=Val v::expr, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), True)" have "(?b, ?b') \ (red P)\<^sup>*" using False SFAssInitDoneRed[THEN converse_rtrancl_into_rtrancl, OF cF] shD by simp also have "(?b', ?c) \ (red P)\<^sup>*" using RedSFAss[THEN r_into_rtrancl, OF cF] shD by simp ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed (*>*) lemma SFAssRedsThrow: "\ P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\ \ \ P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\" (*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ red_reds.SFAssThrow])(*>*) lemma SFAssRedsNone: "\ P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\; \(\b t. P \ C has F,b:t in D) \ \ P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \THROW NoSuchFieldError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" (*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ RedSFAssNone])(*>*) lemma SFAssRedsNonStatic: "\ P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\; P \ C has F,NonStatic:t in D \ \ P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" (*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ RedSFAssNonStatic])(*>*) lemma SFAssInitReds: assumes e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" and cF: "P \ C has F,Static:t in D" and nDone: "\sfs. sh\<^sub>2 D = Some (sfs, Done)" and INIT_steps: "P \ \INIT D ([D],False) \ unit,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\ \* \Val v',(h',l',sh'),b'\" and sh'D: "sh' D = Some(sfs,i)" and sfs': "sfs' = sfs(F\v)" and sh'': "sh'' = sh'(D\(sfs',i))" shows "P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \unit,(h',l',sh''),False\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(C\\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)" let ?y' = "(INIT D ([D],False) \ C\\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)" let ?y'' = "(C\\<^sub>sF{D} := Val v::expr,(h', l', sh'),icheck P D (C\\<^sub>sF{D} := Val v::expr))" have icheck: "icheck P D (C\\<^sub>sF{D} := Val v::expr)" using cF by fastforce then have y''z: "(?y'',?z) \ (red P)" using RedSFAss[OF cF _ sfs' sh''] sh'D by simp have "(?x, ?y) \ (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" using SFAssInitRed[THEN converse_rtrancl_into_rtrancl, OF cF] nDone by simp also have "(?y', ?z) \ (red P)\<^sup>*" using InitSeqReds[OF INIT_steps y''z[THEN r_into_rtrancl]] by simp ultimately show ?thesis by simp qed (*>*) lemma SFAssInitThrowReds: assumes e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" and cF: "P \ C has F,Static:t in D" and nDone: "\sfs. sh\<^sub>2 D = Some (sfs, Done)" and INIT_steps: "P \ \INIT D ([D],False) \ unit,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\ \* \throw a,s',b'\" shows "P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \throw a,s',b'\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(C\\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)" let ?y' = "(INIT D ([D],False) \ C\\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" using SFAssInitRed[THEN converse_rtrancl_into_rtrancl, OF cF] nDone by simp also have "(?y', ?z) \ (red P)\<^sup>*" using InitSeqThrowReds[OF INIT_steps] by simp ultimately show ?thesis by simp qed (*>*) subsubsection";;" lemma SeqReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \e;;e\<^sub>2, s,b\ \* \e';;e\<^sub>2, s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) SeqRed[OF step(2)]]) qed (*>*) lemma SeqRedsThrow: "P \ \e,s,b\ \* \throw e',s',b'\ \ P \ \e;;e\<^sub>2, s,b\ \* \throw e', s',b'\" (*<*)by(rule SeqReds[THEN rtrancl_into_rtrancl, OF _ red_reds.SeqThrow])(*>*) lemma SeqReds2: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1,b\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \e\<^sub>2',s\<^sub>2,b\<^sub>2\" shows "P \ \e\<^sub>1;;e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \e\<^sub>2',s\<^sub>2,b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(Val v\<^sub>1;; e\<^sub>2,s\<^sub>1,b\<^sub>1)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule SeqReds[OF e\<^sub>1_steps]) also have "(?y, ?z) \ (red P)\<^sup>*" by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps]) ultimately show ?thesis by simp qed (*>*) subsubsection"If" lemma CondReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \if (e) e\<^sub>1 else e\<^sub>2,s,b\ \* \if (e') e\<^sub>1 else e\<^sub>2,s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) CondRed[OF step(2)]]) qed (*>*) lemma CondRedsThrow: "P \ \e,s,b\ \* \throw a,s',b'\ \ P \ \if (e) e\<^sub>1 else e\<^sub>2, s,b\ \* \throw a,s',b'\" (*<*)by(rule CondReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CondThrow])(*>*) lemma CondReds2T: assumes e_steps: "P \ \e,s\<^sub>0,b\<^sub>0\ \* \true,s\<^sub>1,b\<^sub>1\" and e\<^sub>1_steps: "P \ \e\<^sub>1, s\<^sub>1,b\<^sub>1\ \* \e',s\<^sub>2,b\<^sub>2\" shows "P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \e',s\<^sub>2,b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(if (true) e\<^sub>1 else e\<^sub>2,s\<^sub>1,b\<^sub>1)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF e_steps]) also have "(?y, ?z) \ (red P)\<^sup>*" by(rule RedCondT[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>1_steps]) ultimately show ?thesis by simp qed (*>*) lemma CondReds2F: assumes e_steps: "P \ \e,s\<^sub>0,b\<^sub>0\ \* \false,s\<^sub>1,b\<^sub>1\" and e\<^sub>2_steps: "P \ \e\<^sub>2, s\<^sub>1,b\<^sub>1\ \* \e',s\<^sub>2,b\<^sub>2\" shows "P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \e',s\<^sub>2,b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(if (false) e\<^sub>1 else e\<^sub>2,s\<^sub>1,b\<^sub>1)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF e_steps]) also have "(?y, ?z) \ (red P)\<^sup>*" by(rule RedCondF[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps]) ultimately show ?thesis by simp qed (*>*) subsubsection "While" lemma WhileFReds: assumes b_steps: "P \ \b,s,b\<^sub>0\ \* \false,s',b'\" shows "P \ \while (b) c,s,b\<^sub>0\ \* \unit,s',b'\" (*<*) by(rule RedWhile[THEN converse_rtrancl_into_rtrancl, OF CondReds[THEN rtrancl_into_rtrancl, OF b_steps RedCondF]]) (*>*) lemma WhileRedsThrow: assumes b_steps: "P \ \b,s,b\<^sub>0\ \* \throw e,s',b'\" shows "P \ \while (b) c,s,b\<^sub>0\ \* \throw e,s',b'\" (*<*) by(rule RedWhile[THEN converse_rtrancl_into_rtrancl, OF CondReds[THEN rtrancl_into_rtrancl, OF b_steps red_reds.CondThrow]]) (*>*) lemma WhileTReds: assumes b_steps: "P \ \b,s\<^sub>0,b\<^sub>0\ \* \true,s\<^sub>1,b\<^sub>1\" and c_steps: "P \ \c,s\<^sub>1,b\<^sub>1\ \* \Val v\<^sub>1,s\<^sub>2,b\<^sub>2\" and while_steps: "P \ \while (b) c,s\<^sub>2,b\<^sub>2\ \* \e,s\<^sub>3,b\<^sub>3\" shows "P \ \while (b) c,s\<^sub>0,b\<^sub>0\ \* \e,s\<^sub>3,b\<^sub>3\" (*<*)(is "(?a, ?c) \ (red P)\<^sup>*") proof - let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0,b\<^sub>0)" let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1,b\<^sub>1)" let ?b' = "(c;; while (b) c,s\<^sub>1,b\<^sub>1)" let ?y' = "(Val v\<^sub>1;; while (b) c,s\<^sub>2,b\<^sub>2)" have "(?a, ?b) \ (red P)\<^sup>*" using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp also have "(?b, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF b_steps]) also have "(?y, ?b') \ (red P)\<^sup>*" using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp also have "(?b', ?y') \ (red P)\<^sup>*" by(rule SeqReds[OF c_steps]) also have "(?y', ?c) \ (red P)\<^sup>*" by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF while_steps]) ultimately show ?thesis by simp qed (*>*) lemma WhileTRedsThrow: assumes b_steps: "P \ \b,s\<^sub>0,b\<^sub>0\ \* \true,s\<^sub>1,b\<^sub>1\" and c_steps: "P \ \c,s\<^sub>1,b\<^sub>1\ \* \throw e,s\<^sub>2,b\<^sub>2\" shows "P \ \while (b) c,s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\" (*<*)(is "(?a, ?c) \ (red P)\<^sup>*") proof - let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0,b\<^sub>0)" let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1,b\<^sub>1)" let ?b' = "(c;; while (b) c,s\<^sub>1,b\<^sub>1)" let ?y' = "(throw e;; while (b) c,s\<^sub>2,b\<^sub>2)" have "(?a, ?b) \ (red P)\<^sup>*" using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp also have "(?b, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF b_steps]) also have "(?y, ?b') \ (red P)\<^sup>*" using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp also have "(?b', ?y') \ (red P)\<^sup>*" by(rule SeqReds[OF c_steps]) also have "(?y', ?c) \ (red P)" by(rule red_reds.SeqThrow) ultimately show ?thesis by simp qed (*>*) subsubsection"Throw" lemma ThrowReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \throw e,s,b\ \* \throw e',s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) ThrowRed[OF step(2)]]) qed (*>*) lemma ThrowRedsNull: "P \ \e,s,b\ \* \null,s',b'\ \ P \ \throw e,s,b\ \* \THROW NullPointer,s',b'\" (*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ RedThrowNull])(*>*) lemma ThrowRedsThrow: "P \ \e,s,b\ \* \throw a,s',b'\ \ P \ \throw e,s,b\ \* \throw a,s',b'\" (*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ red_reds.ThrowThrow])(*>*) subsubsection "InitBlock" lemma InitBlockReds_aux: "P \ \e,s,b\ \* \e',s',b'\ \ \h l sh h' l' sh' v. s = (h,l(V\v),sh) \ s' = (h',l',sh') \ P \ \{V:T := Val v; e},(h,l,sh),b\ \* \{V:T := Val(the(l' V)); e'},(h',l'(V:=(l V)),sh'),b'\" (*<*) proof(induct rule: converse_rtrancl_induct3) case refl then show ?case by(fastforce simp: fun_upd_same simp del:fun_upd_apply) next case (step e0 s0 b0 e1 s1 b1) obtain h1 l1 sh1 where s1[simp]: "s1 = (h1, l1, sh1)" by(cases s1) simp { fix h0 l0 sh0 h2 l2 sh2 v0 assume [simp]: "s0 = (h0, l0(V \ v0), sh0)" and s'[simp]: "s' = (h2, l2, sh2)" then have "V \ dom l1" using step(1) by(auto dest!: red_lcl_incr) then obtain v1 where l1V[simp]: "l1 V = \v1\" by blast let ?a = "({V:T; V:=Val v0;; e0},(h0, l0, sh0),b0)" let ?b = "({V:T; V:=Val v1;; e1},(h1, l1(V := l0 V), sh1),b1)" let ?c = "({V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l0 V), sh2),b')" let ?l = "l1(V := l0 V)" and ?v = v1 have e0_steps: "P \ \e0,(h0, l0(V \ v0), sh0),b0\ \ \e1,(h1, l1, sh1),b1\" using step(1) by simp have lv: "\l v. l1 = l(V \ v) \ P \ \{V:T; V:=Val v;; e1},(h1, l, sh1),b1\ \* \{V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l V), sh2),b'\" using step(3) s' s1 by blast moreover have "l1 = ?l(V \ ?v)" by(rule ext) (simp add:fun_upd_def) ultimately have "(?b, ?c) \ (red P)\<^sup>*" using lv[of ?l ?v] by simp then have "(?a, ?c) \ (red P)\<^sup>*" by(rule InitBlockRed[THEN converse_rtrancl_into_rtrancl, OF e0_steps l1V]) } then show ?case by blast qed (*>*) lemma InitBlockReds: "P \ \e, (h,l(V\v),sh),b\ \* \e', (h',l',sh'),b'\ \ P \ \{V:T := Val v; e}, (h,l,sh),b\ \* \{V:T := Val(the(l' V)); e'}, (h',l'(V:=(l V)),sh'),b'\" (*<*)by(blast dest:InitBlockReds_aux)(*>*) lemma InitBlockRedsFinal: "\ P \ \e,(h,l(V\v),sh),b\ \* \e',(h',l',sh'),b'\; final e' \ \ P \ \{V:T := Val v; e},(h,l,sh),b\ \* \e',(h', l'(V := l V),sh'),b'\" (*<*) by(fast elim!:InitBlockReds[THEN rtrancl_into_rtrancl] finalE intro:RedInitBlock InitBlockThrow) (*>*) subsubsection "Block" lemmas converse_rtranclE3 = converse_rtranclE [of "(xa,xb,xc)" "(za,zb,zc)", split_rule] lemma BlockRedsFinal: assumes reds: "P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" and fin: "final e\<^sub>2" shows "\h\<^sub>0 l\<^sub>0 sh\<^sub>0. s\<^sub>0 = (h\<^sub>0,l\<^sub>0(V:=None),sh\<^sub>0) \ P \ \{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\" (*<*) using reds proof (induct rule:converse_rtrancl_induct3) case refl thus ?case by(fastforce intro:finalE[OF fin] RedBlock BlockThrow simp del:fun_upd_apply) next case (step e\<^sub>0 s\<^sub>0 b\<^sub>0 e\<^sub>1 s\<^sub>1 b\<^sub>1) have red: "P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\" and reds: "P \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" and IH: "\h l sh. s\<^sub>1 = (h,l(V := None),sh) \ P \ \{V:T; e\<^sub>1},(h,l,sh),b\<^sub>1\ \* \e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l V),sh\<^sub>2),b\<^sub>2\" and s\<^sub>0: "s\<^sub>0 = (h\<^sub>0, l\<^sub>0(V := None),sh\<^sub>0)" by fact+ obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" using prod_cases3 by blast show ?case proof cases assume "assigned V e\<^sub>0" then obtain v e where e\<^sub>0: "e\<^sub>0 = V := Val v;; e" by (unfold assigned_def)blast from red e\<^sub>0 s\<^sub>0 have e\<^sub>1: "e\<^sub>1 = unit;;e" and s\<^sub>1: "s\<^sub>1 = (h\<^sub>0, l\<^sub>0(V \ v),sh\<^sub>0)" and b\<^sub>1: "b\<^sub>1 = b\<^sub>0" by auto from e\<^sub>1 fin have "e\<^sub>1 \ e\<^sub>2" by (auto simp:final_def) then obtain e' s' b' where red1: "P \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\ \ \e',s',b'\" and reds': "P \ \e',s',b'\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" using converse_rtranclE3[OF reds] by blast from red1 e\<^sub>1 b\<^sub>1 have es': "e' = e" "s' = s\<^sub>1" "b' = b\<^sub>0" by auto show ?case using e\<^sub>0 s\<^sub>1 es' reds' by(auto intro!: InitBlockRedsFinal[OF _ fin] simp del:fun_upd_apply) next assume unass: "\ assigned V e\<^sub>0" show ?thesis proof (cases "l\<^sub>1 V") assume None: "l\<^sub>1 V = None" hence "P \ \{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\ \ \{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\" using s\<^sub>0 s\<^sub>1 red by(simp add: BlockRedNone[OF _ _ unass]) moreover have "P \ \{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\ \* \e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\" using IH[of _ "l\<^sub>1(V := l\<^sub>0 V)"] s\<^sub>1 None by(simp add:fun_upd_idem) ultimately show ?case by(rule converse_rtrancl_into_rtrancl) next fix v assume Some: "l\<^sub>1 V = Some v" hence "P \ \{V:T;e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\ \ \{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\" using s\<^sub>0 s\<^sub>1 red by(simp add: BlockRedSome[OF _ _ unass]) moreover have "P \ \{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V:= l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\" using InitBlockRedsFinal[OF _ fin,of _ _ "l\<^sub>1(V:=l\<^sub>0 V)" V] Some reds s\<^sub>1 by(simp add:fun_upd_idem) ultimately show ?case by(rule converse_rtrancl_into_rtrancl) qed qed qed (*>*) subsubsection "try-catch" lemma TryReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \try e catch(C V) e\<^sub>2,s,b\ \* \try e' catch(C V) e\<^sub>2,s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) TryRed[OF step(2)]]) qed (*>*) lemma TryRedsVal: "P \ \e,s,b\ \* \Val v,s',b'\ \ P \ \try e catch(C V) e\<^sub>2,s,b\ \* \Val v,s',b'\" (*<*)by(rule TryReds[THEN rtrancl_into_rtrancl, OF _ RedTry])(*>*) lemma TryCatchRedsFinal: assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Throw a,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\" and h\<^sub>1a: "h\<^sub>1 a = Some(D,fs)" and sub: "P \ D \\<^sup>* C" and e\<^sub>2_steps: "P \ \e\<^sub>2, (h\<^sub>1, l\<^sub>1(V \ Addr a),sh\<^sub>1),b\<^sub>1\ \* \e\<^sub>2', (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2\" and final: "final e\<^sub>2'" shows "P \ \try e\<^sub>1 catch(C V) e\<^sub>2, s\<^sub>0, b\<^sub>0\ \* \e\<^sub>2', (h\<^sub>2, (l\<^sub>2::locals)(V := l\<^sub>1 V),sh\<^sub>2),b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(try Throw a catch(C V) e\<^sub>2,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\<^sub>1)" let ?b = "({V:Class C; V:=addr a;; e\<^sub>2},(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\<^sub>1)" have bz: "(?b, ?z) \ (red P)\<^sup>*" by(rule InitBlockRedsFinal[OF e\<^sub>2_steps final]) have hp: "hp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) a = \(D, fs)\" using h\<^sub>1a by simp have "(?x, ?y) \ (red P)\<^sup>*" by(rule TryReds[OF e\<^sub>1_steps]) also have "(?y, ?z) \ (red P)\<^sup>*" by(rule RedTryCatch[THEN converse_rtrancl_into_rtrancl, OF hp sub bz]) ultimately show ?thesis by simp qed (*>*) lemma TryRedsFail: "\ P \ \e\<^sub>1,s,b\ \* \Throw a,(h,l,sh),b'\; h a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \try e\<^sub>1 catch(C V) e\<^sub>2,s,b\ \* \Throw a,(h,l,sh),b'\" (*<*)by(fastforce intro!: TryReds[THEN rtrancl_into_rtrancl, OF _ RedTryFail])(*>*) subsubsection "List" lemma ListReds1: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \e#es,s,b\ [\]* \e' # es,s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) ListRed1[OF step(2)]]) qed (*>*) lemma ListReds2: "P \ \es,s,b\ [\]* \es',s',b'\ \ P \ \Val v # es,s,b\ [\]* \Val v # es',s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) ListRed2[OF step(2)]]) qed (*>*) lemma ListRedsVal: "\ P \ \e,s\<^sub>0,b\<^sub>0\ \* \Val v,s\<^sub>1,b\<^sub>1\; P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \es',s\<^sub>2,b\<^sub>2\ \ \ P \ \e#es,s\<^sub>0,b\<^sub>0\ [\]* \Val v # es',s\<^sub>2,b\<^sub>2\" (*<*)by(rule rtrancl_trans[OF ListReds1 ListReds2])(*>*) subsubsection"Call" text\ First a few lemmas on what happens to free variables during redction. \ lemma assumes wf: "wwf_J_prog P" shows Red_fv: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ fv e' \ fv e" and "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ fvs es' \ fvs es" (*<*) proof (induct rule:red_reds_inducts) case (RedCall h a C fs M Ts T pns body D vs l sh b) hence "fv body \ {this} \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def) with RedCall.hyps show ?case by fastforce next case (RedSCall C M Ts T pns body D vs) hence "fv body \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def) with RedSCall.hyps show ?case by fastforce qed auto (*>*) lemma Red_dom_lcl: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ dom l' \ dom l \ fv e" and "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ dom l' \ dom l \ fvs es" (*<*) proof (induct rule:red_reds_inducts) case RedLAss thus ?case by(force split:if_splits) next case CallParams thus ?case by(force split:if_splits) next case BlockRedNone thus ?case by clarsimp (fastforce split:if_splits) next case BlockRedSome thus ?case by clarsimp (fastforce split:if_splits) next case InitBlockRed thus ?case by clarsimp (fastforce split:if_splits) qed auto (*>*) lemma Reds_dom_lcl: assumes wf: "wwf_J_prog P" shows "P \ \e,(h,l,sh),b\ \* \e',(h',l',sh'),b'\ \ dom l' \ dom l \ fv e" (*<*) proof(induct rule: converse_rtrancl_induct_red) case 1 then show ?case by blast next case 2 then show ?case using wf by(blast dest: Red_fv Red_dom_lcl) qed (*>*) text\ Now a few lemmas on the behaviour of blocks during reduction. \ lemma override_on_upd_lemma: "(override_on f (g(a\b)) A)(a := g a) = override_on f g (insert a A)" (*<*)by(rule ext) (simp add:override_on_def) declare fun_upd_apply[simp del] map_upds_twist[simp del] (*>*) lemma blocksReds: "\l. \ length Vs = length Ts; length vs = length Ts; distinct Vs; P \ \e, (h,l(Vs [\] vs),sh),b\ \* \e', (h',l',sh'),b'\ \ \ P \ \blocks(Vs,Ts,vs,e), (h,l,sh),b\ \* \blocks(Vs,Ts,map (the \ l') Vs,e'), (h',override_on l' l (set Vs),sh'),b'\" (*<*) proof(induct Vs Ts vs e rule:blocks_induct) case (1 V Vs T Ts v vs e) show ?case using InitBlockReds[OF "1.hyps"[of "l(V\v)"]] "1.prems" by(auto simp:override_on_upd_lemma) qed auto (*>*) lemma blocksFinal: "\l. \ length Vs = length Ts; length vs = length Ts; final e \ \ P \ \blocks(Vs,Ts,vs,e), (h,l,sh),b\ \* \e, (h,l,sh),b\" (*<*) proof(induct Vs Ts vs e rule:blocks_induct) case 1 show ?case using "1.prems" InitBlockReds[OF "1.hyps"] by(fastforce elim!:finalE elim: rtrancl_into_rtrancl[OF _ RedInitBlock] rtrancl_into_rtrancl[OF _ InitBlockThrow]) qed auto (*>*) lemma blocksRedsFinal: assumes wf: "length Vs = length Ts" "length vs = length Ts" "distinct Vs" and reds: "P \ \e, (h,l(Vs [\] vs),sh),b\ \* \e', (h',l',sh'),b'\" and fin: "final e'" and l'': "l'' = override_on l' l (set Vs)" shows "P \ \blocks(Vs,Ts,vs,e), (h,l,sh),b\ \* \e', (h',l'',sh'),b'\" (*<*) proof - let ?bv = "blocks(Vs,Ts,map (the o l') Vs,e')" have "P \ \blocks(Vs,Ts,vs,e), (h,l,sh),b\ \* \?bv, (h',l'',sh'),b'\" using l'' by simp (rule blocksReds[OF wf reds]) also have "P \ \?bv, (h',l'',sh'),b'\ \* \e', (h',l'',sh'),b'\" using wf by(fastforce intro:blocksFinal fin) finally show ?thesis . qed (*>*) text\ An now the actual method call reduction lemmas. \ lemma CallRedsObj: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \e\M(es),s,b\ \* \e'\M(es),s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) CallObj[OF step(2)]]) qed (*>*) lemma CallRedsParams: "P \ \es,s,b\ [\]* \es',s',b'\ \ P \ \(Val v)\M(es),s,b\ \* \(Val v)\M(es'),s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) CallParams[OF step(2)]]) qed (*>*) lemma CallRedsFinal: assumes wwf: "wwf_J_prog P" and "P \ \e,s\<^sub>0,b\<^sub>0\ \* \addr a,s\<^sub>1,b\<^sub>1\" "P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" "h\<^sub>2 a = Some(C,fs)" "P \ C sees M,NonStatic:Ts\T = (pns,body) in D" "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [this \ Addr a, pns[\]vs]" and body: "P \ \body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),b\<^sub>2\ \* \ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\" and "final ef" shows "P \ \e\M(es), s\<^sub>0,b\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\" (*<*) proof - have wf: "size Ts = size pns \ distinct pns \ this \ set pns" and wt: "fv body \ {this} \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ from body[THEN Red_lcl_add, of l\<^sub>2] have body': "P \ \body,(h\<^sub>2,l\<^sub>2(this\ Addr a, pns[\]vs),sh\<^sub>2),b\<^sub>2\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\" by (simp add:l\<^sub>2') have "dom l\<^sub>3 \ {this} \ set pns" using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns) = l\<^sub>2" by(fastforce simp add:map_add_def override_on_def fun_eq_iff) have "P \ \e\M(es),s\<^sub>0,b\<^sub>0\ \* \(addr a)\M(es),s\<^sub>1,b\<^sub>1\" by(rule CallRedsObj)(rule assms(2)) also have "P \ \(addr a)\M(es),s\<^sub>1,b\<^sub>1\ \* \(addr a)\M(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" by(rule CallRedsParams)(rule assms(3)) also have "P \ \(addr a)\M(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \ \blocks(this#pns, Class D#Ts, Addr a#vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" by(rule RedCall)(auto simp: assms wf, rule assms(5)) also (rtrancl_into_rtrancl) have "P \ \blocks(this#pns, Class D#Ts, Addr a#vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns),sh\<^sub>3),b\<^sub>3\" by(rule blocksRedsFinal, insert assms wf body', simp_all) finally show ?thesis using eql\<^sub>2 by simp qed (*>*) lemma CallRedsThrowParams: assumes e_steps: "P \ \e,s\<^sub>0,b\<^sub>0\ \* \Val v,s\<^sub>1,b\<^sub>1\" and es_steps: "P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2,b\<^sub>2\" shows "P \ \e\M(es),s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>2,b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(Val v\M(es),s\<^sub>1,b\<^sub>1)" let ?y' = "(Val v\M(map Val vs\<^sub>1 @ throw a # es\<^sub>2),s\<^sub>2,b\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) also have "(?y', ?z) \ (red P)\<^sup>*" using CallThrowParams by fast ultimately show ?thesis by simp qed (*>*) lemma CallRedsThrowObj: "P \ \e,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\ \ P \ \e\M(es),s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" (*<*)by(rule CallRedsObj[THEN rtrancl_into_rtrancl, OF _ CallThrowObj])(*>*) lemma CallRedsNull: assumes e_steps: "P \ \e,s\<^sub>0,b\<^sub>0\ \* \null,s\<^sub>1,b\<^sub>1\" and es_steps: "P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs,s\<^sub>2,b\<^sub>2\" shows "P \ \e\M(es),s\<^sub>0,b\<^sub>0\ \* \THROW NullPointer,s\<^sub>2,b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(null\M(es),s\<^sub>1,b\<^sub>1)" let ?y' = "(null\M(map Val vs),s\<^sub>2,b\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) also have "(?y', ?z) \ (red P)" by(rule RedCallNull) ultimately show ?thesis by simp qed (*>*) lemma CallRedsNone: assumes e_steps: "P \ \e,s,b\ \* \addr a,s\<^sub>1,b\<^sub>1\" and es_steps: "P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs,s\<^sub>2,b\<^sub>2\" and hp\<^sub>2a: "hp s\<^sub>2 a = Some(C,fs)" and ncM: "\(\b Ts T m D. P \ C sees M,b:Ts\T = m in D)" shows "P \ \e\M(es),s,b\ \* \THROW NoSuchMethodError,s\<^sub>2,b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(addr a\M(es),s\<^sub>1,b\<^sub>1)" let ?y' = "(addr a\M(map Val vs),s\<^sub>2,b\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) also have "(?y', ?z) \ (red P)" using RedCallNone[OF _ ncM] hp\<^sub>2a by(cases s\<^sub>2) simp ultimately show ?thesis by simp qed (*>*) lemma CallRedsStatic: assumes e_steps: "P \ \e,s,b\ \* \addr a,s\<^sub>1,b\<^sub>1\" and es_steps: "P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs,s\<^sub>2,b\<^sub>2\" and hp\<^sub>2a: "hp s\<^sub>2 a = Some(C,fs)" and cM_Static: "P \ C sees M,Static:Ts\T = m in D" shows "P \ \e\M(es),s,b\ \* \THROW IncompatibleClassChangeError,s\<^sub>2,b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(addr a\M(es),s\<^sub>1,b\<^sub>1)" let ?y' = "(addr a\M(map Val vs),s\<^sub>2,b\<^sub>2)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) also have "(?y', ?z) \ (red P)" using RedCallStatic[OF _ cM_Static] hp\<^sub>2a by(cases s\<^sub>2) simp ultimately show ?thesis by simp qed (*>*) subsection\SCall\ lemma SCallRedsParams: "P \ \es,s,b\ [\]* \es',s',b'\ \ P \ \C\\<^sub>sM(es),s,b\ \* \C\\<^sub>sM(es'),s',b'\" (*<*) proof(induct rule: rtrancl_induct3) case refl show ?case by blast next case step show ?case by(rule rtrancl_into_rtrancl[OF step(3) SCallParams[OF step(2)]]) qed (*>*) lemma SCallRedsFinal: assumes wwf: "wwf_J_prog P" and "P \ \es,s\<^sub>0,b\<^sub>0\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" "P \ C sees M,Static:Ts\T = (pns,body) in D" "sh\<^sub>2 D = Some(sfs,Done) \ (M = clinit \ sh\<^sub>2 D = \(sfs, Processing)\)" "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [pns[\]vs]" and body: "P \ \body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\" and "final ef" shows "P \ \C\\<^sub>sM(es), s\<^sub>0,b\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\" (*<*) proof - have wf: "size Ts = size pns \ distinct pns" and wt: "fv body \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ from body[THEN Red_lcl_add, of l\<^sub>2] have body': "P \ \body,(h\<^sub>2,l\<^sub>2(pns[\]vs),sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\" by (simp add:l\<^sub>2') have "dom l\<^sub>3 \ set pns" using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2" by(fastforce simp add:map_add_def override_on_def fun_eq_iff) have b2T: "P \ \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \* \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),True\" proof(cases b\<^sub>2) case True then show ?thesis by simp next case False then show ?thesis using assms(3,4) by(auto elim: SCallInitDoneRed) qed have "P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \C\\<^sub>sM(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" by(rule SCallRedsParams)(rule assms(2)) also have "P \ \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \* \blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" by(auto intro!: rtrancl_into_rtrancl[OF b2T] RedSCall assms(3) simp: assms wf) also (rtrancl_trans) have "P \ \blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\" by(rule blocksRedsFinal, insert assms wf body', simp_all) finally show ?thesis using eql\<^sub>2 by simp qed (*>*) lemma SCallRedsThrowParams: "\ P \ \es,s\<^sub>0,b\<^sub>0\ [\]* \map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2,b\<^sub>2\ \ \ P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>2,b\<^sub>2\" (*<*) by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ SCallThrowParams]) simp (*>*) lemma SCallRedsNone: "\ P \ \es,s,b\ [\]* \map Val vs,s\<^sub>2,False\; \(\b Ts T m D. P \ C sees M,b:Ts\T = m in D) \ \ P \ \C\\<^sub>sM(es),s,b\ \* \THROW NoSuchMethodError,s\<^sub>2,False\" (*<*) by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ RedSCallNone]) simp (*>*) lemma SCallRedsNonStatic: "\ P \ \es,s,b\ [\]* \map Val vs,s\<^sub>2,False\; P \ C sees M,NonStatic:Ts\T = m in D \ \ P \ \C\\<^sub>sM(es),s,b\ \* \THROW IncompatibleClassChangeError,s\<^sub>2,False\" (*<*) by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ RedSCallNonStatic]) simp (*>*) lemma SCallInitThrowReds: assumes wwf: "wwf_J_prog P" and "P \ \es,s\<^sub>0,b\<^sub>0\ [\]* \map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\" "P \ C sees M,Static:Ts\T = (pns,body) in D" "\sfs. sh\<^sub>1 D = Some(sfs,Done)" "M \ clinit" and "P \ \INIT D ([D],False) \ unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\ \* \throw a,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" shows "P \ \C\\<^sub>sM(es), s\<^sub>0,b\<^sub>0\ \* \throw a,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" (*<*)(is "(?x, ?z) \ (red P)\<^sup>*") proof - let ?y = "(C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False)" let ?y' = "(INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False)" have "(?x, ?y) \ (red P)\<^sup>*" by(rule SCallRedsParams[OF assms(2)]) also have "(?y, ?y') \ (red P)\<^sup>*" using SCallInitRed[OF assms(3)] assms(4-5) by auto also have "(?y', ?z) \ (red P)\<^sup>*" by(rule InitSeqThrowReds[OF assms(6)]) finally show ?thesis by simp qed (*>*) lemma SCallInitReds: assumes wwf: "wwf_J_prog P" and "P \ \es,s\<^sub>0,b\<^sub>0\ [\]* \map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\" "P \ C sees M,Static:Ts\T = (pns,body) in D" "\sfs. sh\<^sub>1 D = Some(sfs,Done)" "M \ clinit" and "P \ \INIT D ([D],False) \ unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\ \* \Val v',(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" and "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [pns[\]vs]" and body: "P \ \body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\" and "final ef" shows "P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\" (*<*) proof - have wf: "size Ts = size pns \ distinct pns" and wt: "fv body \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ from body[THEN Red_lcl_add, of l\<^sub>2] have body': "P \ \body,(h\<^sub>2,l\<^sub>2(pns[\]vs),sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\" by (simp add:l\<^sub>2') have "dom l\<^sub>3 \ set pns" using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2" by(fastforce simp add:map_add_def override_on_def fun_eq_iff) have "icheck P D (C\\<^sub>sM(map Val vs)::'a exp)" using assms(3) by auto then have "P \ \C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),icheck P D (C\\<^sub>sM(map Val vs))\ \ \blocks(pns, Ts, vs, body), (h\<^sub>2, l\<^sub>2, sh\<^sub>2), False\" by (metis (full_types) assms(3) assms(7) local.wf red_reds.RedSCall) also have "P \ \blocks(pns, Ts, vs, body), (h\<^sub>2, l\<^sub>2, sh\<^sub>2), False\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\" by(rule blocksRedsFinal, insert assms wf body', simp_all) finally have trueReds: "P \ \C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),icheck P D (C\\<^sub>sM(map Val vs))\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\" by simp have "P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\" by(rule SCallRedsParams)(rule assms(2)) also have "P \ \C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\ \ \INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\" using SCallInitRed[OF assms(3)] assms(4-5) by auto also (rtrancl_into_rtrancl) have "P \ \INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\" using InitSeqReds[OF assms(6) trueReds] assms(5) by simp finally show ?thesis using eql\<^sub>2 by simp qed (*>*) lemma SCallInitProcessingReds: assumes wwf: "wwf_J_prog P" and "P \ \es,s\<^sub>0,b\<^sub>0\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" "P \ C sees M,Static:Ts\T = (pns,body) in D" "sh\<^sub>2 D = Some(sfs,Processing)" and "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [pns[\]vs]" and body: "P \ \body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\" and "final ef" shows "P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\" (*<*) proof - have wf: "size Ts = size pns \ distinct pns" and wt: "fv body \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ from body[THEN Red_lcl_add, of l\<^sub>2] have body': "P \ \body,(h\<^sub>2,l\<^sub>2(pns[\]vs),sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\" by (simp add:l\<^sub>2') have "dom l\<^sub>3 \ set pns" using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2" by(fastforce simp add:map_add_def override_on_def fun_eq_iff) have b2T: "P \ \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \* \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),True\" proof(cases b\<^sub>2) case True then show ?thesis by simp next case False show ?thesis proof(cases "M = clinit") case True then show ?thesis using False assms(3) red_reds.SCallInitDoneRed assms(4) by (simp add: r_into_rtrancl) next case nclinit: False have icheck: "icheck P D (C\\<^sub>sM(map Val vs))" using assms(3) by auto have "P \ \C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b\<^sub>2\ \ \INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SCallInitRed[OF assms(3)] assms(4) False nclinit by simp also have "P \ \INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\ \ \INIT D (Nil,True) \ C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using RedInitProcessing assms(4) by simp also have "P \ \INIT D (Nil,True) \ C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\ \ \C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),True\" using RedInit[of "C\\<^sub>sM(map Val vs)" D _ _ _ P] icheck nclinit by (metis (full_types) nsub_RI_Vals sub_RI.simps(12)) finally show ?thesis by simp qed qed have "P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \C\\<^sub>sM(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" by(rule SCallRedsParams)(rule assms(2)) also have "P \ \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \* \blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" by(auto intro!: rtrancl_into_rtrancl[OF b2T] RedSCall assms(3) simp: assms wf) also (rtrancl_trans) have "P \ \blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\" by(rule blocksRedsFinal, insert assms wf body', simp_all) finally show ?thesis using eql\<^sub>2 by simp qed (*>*) (***********************************) subsubsection "The main Theorem" lemma assumes wwf: "wwf_J_prog P" shows big_by_small: "P \ \e,s\ \ \e',s'\ \ (\b. iconf (shp s) e \ P,shp s \\<^sub>b (e,b) \ \ P \ \e,s,b\ \* \e',s',False\)" and bigs_by_smalls: "P \ \es,s\ [\] \es',s'\ \ (\b. iconfs (shp s) es \ P,shp s \\<^sub>b (es,b) \ \ P \ \es,s,b\ [\]* \es',s',False\)" (*<*) proof (induct rule: eval_evals.inducts) case New show ?case proof(cases b) case True then show ?thesis using RedNew[OF New.hyps(2-4)] by auto next case False then show ?thesis using New.hyps(1) NewInitDoneReds[OF _ New.hyps(2-4)] by auto qed next case NewFail show ?case proof(cases b) case True then show ?thesis using RedNewFail[OF NewFail.hyps(2)] NewFail.hyps(3) by fastforce next case False then show ?thesis using NewInitDoneReds2[OF _ NewFail.hyps(2)] NewFail by fastforce qed next case (NewInit sh C h l v' h' l' sh' a FDTs h'') show ?case proof(cases b) case True then obtain sfs where shC: "sh C = \(sfs, Processing)\" using NewInit.hyps(1) NewInit.prems by(clarsimp simp: bconf_def initPD_def) then have s': "(h',l',sh') = (h,l,sh)" using NewInit.hyps(2) init_ProcessingE by clarsimp then show ?thesis using RedNew[OF NewInit.hyps(4-6)] True by auto next case False then have init: "P \ \INIT C ([C],False) \ unit,(h, l, sh),False\ \* \Val v',(h', l', sh'),False\" using NewInit.hyps(3) by(auto simp: bconf_def) then show ?thesis using NewInit NewInitReds[OF _ init NewInit.hyps(4-6)] False has_fields_is_class[OF NewInit.hyps(5)] by auto qed next case (NewInitOOM sh C h l v' h' l' sh') show ?case proof(cases b) case True then obtain sfs where shC: "sh C = \(sfs, Processing)\" using NewInitOOM.hyps(1) NewInitOOM.prems by(clarsimp simp: bconf_def initPD_def) then have s': "(h',l',sh') = (h,l,sh)" using NewInitOOM.hyps(2) init_ProcessingE by clarsimp then show ?thesis using RedNewFail[OF NewInitOOM.hyps(4)] True r_into_rtrancl NewInitOOM.hyps(5) by auto next case False then have init: "P \ \INIT C ([C],False) \ unit,(h, l, sh),False\ \* \Val v',(h', l', sh'),False\" using NewInitOOM.hyps(3) by(auto simp: bconf_def) then show ?thesis using NewInitOOM.hyps(1) NewInitOOMReds[OF _ init NewInitOOM.hyps(4)] False NewInitOOM.hyps(5) by auto qed next case (NewInitThrow sh C h l a s') show ?case proof(cases b) case True then obtain sfs where shC: "sh C = \(sfs, Processing)\" using NewInitThrow.hyps(1) NewInitThrow.prems by(clarsimp simp: bconf_def initPD_def) then show ?thesis using NewInitThrow.hyps(2) init_ProcessingE by blast next case False then have init: "P \ \INIT C ([C],False) \ unit,(h, l, sh),b\ \* \throw a,s',False\" using NewInitThrow.hyps(3) by(auto simp: bconf_def) then show ?thesis using NewInitThrow NewInitThrowReds[of "(h,l,sh)" C P a s'] False by auto qed next case Cast then show ?case by(fastforce intro:CastRedsAddr) next case CastNull then show ?case by(fastforce intro: CastRedsNull) next case CastFail thus ?case by(fastforce intro!:CastRedsFail) next case CastThrow thus ?case by(fastforce dest!:eval_final intro:CastRedsThrow) next case Val then show ?case using exI[of _ b] by(simp add: bconf_def) next case (BinOp e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 v\<^sub>2 s\<^sub>2 bop v) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None BinOp.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None BinOp.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v\<^sub>1,s\<^sub>1,False\" using iconf BinOp.hyps(2) by auto have binop: "P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0,b\ \* \Val v\<^sub>1 \bop\ e\<^sub>2,s\<^sub>1,False\" by(rule BinOp1Reds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] None BinOp by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \Val v\<^sub>2,s\<^sub>2,False\" using BinOp.hyps(4)[OF iconf2'] by auto then show ?thesis using BinOpRedsVal[OF b1 b2 BinOp.hyps(5)] by fast next case (Some a) then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v\<^sub>1,s\<^sub>1,b1\" by (metis (no_types, lifting) BinOp.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have binop: "P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0,b\ \* \Val v\<^sub>1 \bop\ e\<^sub>2,s\<^sub>1,b1\" by(rule BinOp1Reds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] BinOp by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using BinOp.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf binop BinOp.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \Val v\<^sub>2,s\<^sub>2,False\" using BinOp.hyps(4)[OF iconf2'] by auto then show ?thesis using BinOpRedsVal[OF b1 b2 BinOp.hyps(5)] by fast qed next case (BinOpThrow1 e\<^sub>1 s\<^sub>0 e s\<^sub>1 bop e\<^sub>2) show ?case proof(cases "val_of e\<^sub>1") case None then have "iconf (shp s\<^sub>0) e\<^sub>1" and "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using BinOpThrow1.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \throw e,s\<^sub>1,False\" using BinOpThrow1.hyps(2) by auto then have "P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0,b\ \* \throw e,s\<^sub>1,False\" using BinOpThrow1 None by(auto dest!:eval_final simp: BinOpRedsThrow1[OF b1]) then show ?thesis by fast next case (Some a) then show ?thesis using eval_final_same[OF BinOpThrow1.hyps(1)] val_of_spec[OF Some] by auto qed next case (BinOpThrow2 e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 e s\<^sub>2 bop) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None BinOpThrow2.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None BinOpThrow2.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v\<^sub>1,s\<^sub>1,False\" using iconf BinOpThrow2.hyps(2) by auto have binop: "P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0,b\ \* \Val v\<^sub>1 \bop\ e\<^sub>2,s\<^sub>1,False\" by(rule BinOp1Reds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] None BinOpThrow2 by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \throw e,s\<^sub>2,False\" using BinOpThrow2.hyps(4)[OF iconf2'] by auto then show ?thesis using BinOpRedsThrow2[OF b1 b2] by fast next case (Some a) then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v\<^sub>1,s\<^sub>1,b1\" by (metis (no_types, lifting) BinOpThrow2.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have binop: "P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0,b\ \* \Val v\<^sub>1 \bop\ e\<^sub>2,s\<^sub>1,b1\" by(rule BinOp1Reds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] BinOpThrow2 by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using BinOpThrow2.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf binop BinOpThrow2.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \throw e,s\<^sub>2,False\" using BinOpThrow2.hyps(4)[OF iconf2'] by auto then show ?thesis using BinOpRedsThrow2[OF b1 b2] by fast qed next case Var thus ?case by(auto dest:RedVar simp: bconf_def) next case LAss thus ?case by(auto dest: LAssRedsVal) next case LAssThrow thus ?case by(auto dest!:eval_final dest: LAssRedsThrow) next case FAcc thus ?case by(fastforce intro:FAccRedsVal) next case FAccNull thus ?case by(auto dest:FAccRedsNull) next case FAccThrow thus ?case by(auto dest!:eval_final dest:FAccRedsThrow) next case FAccNone then show ?case by(fastforce intro: FAccRedsNone) next case FAccStatic then show ?case by(fastforce intro: FAccRedsStatic) next case SFAcc show ?case proof(cases b) case True then show ?thesis using RedSFAcc SFAcc.hyps by auto next case False then show ?thesis using SFAcc.hyps SFAccInitDoneReds[OF SFAcc.hyps(1)] by auto qed next case (SFAccInit C F t D sh h l v' h' l' sh' sfs i v) show ?case proof(cases b) case True then obtain sfs where shC: "sh D = \(sfs, Processing)\" using SFAccInit.hyps(2) SFAccInit.prems by(clarsimp simp: bconf_def initPD_def) then have s': "(h',l',sh') = (h,l,sh)" using SFAccInit.hyps(3) init_ProcessingE by clarsimp then show ?thesis using RedSFAcc SFAccInit.hyps True by auto next case False then have init: "P \ \INIT D ([D],False) \ unit,(h, l, sh),False\ \* \Val v',(h', l', sh'),False\" using SFAccInit.hyps(4) by(auto simp: bconf_def) then show ?thesis using SFAccInit SFAccInitReds[OF _ _ init] False by auto qed next case (SFAccInitThrow C F t D sh h l a s') show ?case proof(cases b) case True then obtain sfs where shC: "sh D = \(sfs, Processing)\" using SFAccInitThrow.hyps(2) SFAccInitThrow.prems(2) by(clarsimp simp: bconf_def initPD_def) then show ?thesis using SFAccInitThrow.hyps(3) init_ProcessingE by blast next case False then have init: "P \ \INIT D ([D],False) \ unit,(h, l, sh),b\ \* \throw a,s',False\" using SFAccInitThrow.hyps(4) by(auto simp: bconf_def) then show ?thesis using SFAccInitThrow SFAccInitThrowReds False by auto qed next case SFAccNone then show ?case by(fastforce intro: SFAccRedsNone) next case SFAccNonStatic then show ?case by(fastforce intro: SFAccRedsNonStatic) next case (FAss e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D fs' h\<^sub>2') show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAss.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None FAss.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using iconf FAss.hyps(2) by auto have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,False\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAss by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAss.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsVal[OF b1 b2 FAss.hyps(6) FAss.hyps(5)[THEN sym]] FAss.hyps(7,8) by fast next case (Some a') then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) FAss.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,b1\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAss by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using FAss.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf fass FAss.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAss.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsVal[OF b1 b2] FAss.hyps(5)[THEN sym] FAss.hyps(6-8) by fast qed next case (FAssNull e\<^sub>1 s\<^sub>0 s\<^sub>1 e\<^sub>2 v s\<^sub>2 F D) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using FAssNull.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using FAssNull.prems(2) None by simp then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \null,s\<^sub>1,False\" using FAssNull.hyps(2)[OF iconf] by auto have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \null\F{D} := e\<^sub>2,s\<^sub>1,False\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssNull by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \Val v,s\<^sub>2,False\" using FAssNull.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsNull[OF b1 b2] by fast next case (Some a') then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \null,s\<^sub>1,b1\" by (metis (no_types, lifting) FAssNull.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \null\F{D} := e\<^sub>2,s\<^sub>1,b1\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssNull by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using FAssNull.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf fass FAssNull.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \Val v,s\<^sub>2,False\" using FAssNull.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsNull[OF b1 b2] by fast qed next case (FAssThrow1 e\<^sub>1 s\<^sub>0 e' s\<^sub>1 F D e\<^sub>2) show ?case proof(cases "val_of e\<^sub>1") case None then have "iconf (shp s\<^sub>0) e\<^sub>1" and "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using FAssThrow1.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \throw e',s\<^sub>1,False\" using FAssThrow1.hyps(2) by auto then have "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \throw e',s\<^sub>1,False\" using FAssThrow1 None by(auto dest!:eval_final simp: FAssRedsThrow1[OF b1]) then show ?thesis by fast next case (Some a) then show ?thesis using eval_final_same[OF FAssThrow1.hyps(1)] val_of_spec[OF Some] by auto qed next case (FAssThrow2 e\<^sub>1 s\<^sub>0 v s\<^sub>1 e\<^sub>2 e' s\<^sub>2 F D) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssThrow2.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None FAssThrow2.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v,s\<^sub>1,False\" using iconf FAssThrow2.hyps(2) by auto have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \Val v\F{D} := e\<^sub>2,s\<^sub>1,False\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssThrow2 by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \throw e',s\<^sub>2,False\" using FAssThrow2.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsThrow2[OF b1 b2] by fast next case (Some a') then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v,s\<^sub>1,b1\" by (metis (no_types, lifting) FAssThrow2.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \Val v\F{D} := e\<^sub>2,s\<^sub>1,b1\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssThrow2 by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using FAssThrow2.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf fass FAssThrow2.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \throw e',s\<^sub>2,False\" using FAssThrow2.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsThrow2[OF b1 b2] by fast qed next case (FAssNone e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F D) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssNone.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None FAssNone.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using iconf FAssNone.hyps(2) by auto have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,False\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssNone by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAssNone.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsNone[OF b1 b2 FAssNone.hyps(5,6)] by fast next case (Some a') then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) FAssNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,b1\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssNone by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using FAssNone.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf fass FAssNone.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAssNone.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsNone[OF b1 b2 FAssNone.hyps(5,6)] by fast qed next case (FAssStatic e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssStatic.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None FAssStatic.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using iconf FAssStatic.hyps(2) by auto have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,False\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssStatic by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAssStatic.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsStatic[OF b1 b2 FAssStatic.hyps(5,6)] by fast next case (Some a') then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) FAssStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,b1\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssStatic by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using FAssStatic.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf fass FAssStatic.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAssStatic.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsStatic[OF b1 b2 FAssStatic.hyps(5,6)] by fast qed next case (SFAss e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D sfs sfs' sh\<^sub>1') show ?case proof(cases "val_of e\<^sub>2") case None then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAss.prems(2) by simp then have b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SFAss by auto thus ?thesis using SFAssRedsVal[OF b1 SFAss.hyps(3,4)] SFAss.hyps(5,6) by fast next case (Some a) then obtain b1 where b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b1\" by (metis (no_types, lifting) SFAss.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) thus ?thesis using SFAssRedsVal[OF b1 SFAss.hyps(3,4)] SFAss.hyps(5,6) by fast qed next case (SFAssInit e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D v' h' l' sh' sfs i sfs' sh'') then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" by simp show ?case proof(cases "val_of e\<^sub>2") case None then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAssInit.prems(2) by simp then have reds: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SFAssInit.hyps(2)[OF iconf bconf] by auto then have init: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \Val v',(h', l', sh'),False\" using SFAssInit.hyps(6) by(auto simp: bconf_def) then show ?thesis using SFAssInit SFAssInitReds[OF reds SFAssInit.hyps(3) _ init] by auto next case (Some v2) show ?thesis proof(cases b) case False then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" by(simp add: bconf_def) then have reds: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SFAssInit.hyps(2)[OF iconf bconf] by auto then have init: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \Val v',(h', l', sh'),False\" using SFAssInit.hyps(6) by(auto simp: bconf_def) then show ?thesis using SFAssInit SFAssInitReds[OF reds SFAssInit.hyps(3) _ init] by auto next case True have e\<^sub>2: "e\<^sub>2 = Val v2" using val_of_spec[OF Some] by simp then have vs: "v2 = v \ s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF SFAssInit.hyps(1)] by simp then obtain sfs where shC: "sh\<^sub>1 D = \(sfs, Processing)\" using SFAssInit.hyps(3,4) SFAssInit.prems(2) Some True by(cases e\<^sub>2, auto simp: bconf_def initPD_def dest: sees_method_fun) then have s': "(h',l',sh') = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using SFAssInit.hyps(5) init_ProcessingE by clarsimp then show ?thesis using SFAssInit.hyps(3,7-9) True e\<^sub>2 red_reds.RedSFAss vs by auto qed qed next case (SFAssInitThrow e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D a s') then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" by simp show ?case proof(cases "val_of e\<^sub>2") case None then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAssInitThrow.prems(2) by simp then have reds: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SFAssInitThrow.hyps(2)[OF iconf bconf] by auto then have init: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \throw a,s',False\" using SFAssInitThrow.hyps(6) by(auto simp: bconf_def) then show ?thesis using SFAssInitThrow SFAssInitThrowReds[OF reds _ _ init] by auto next case (Some v2) show ?thesis proof(cases b) case False then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" by(simp add: bconf_def) then have reds: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SFAssInitThrow.hyps(2)[OF iconf bconf] by auto then have init: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \throw a,s',False\" using SFAssInitThrow.hyps(6) by(auto simp: bconf_def) then show ?thesis using SFAssInitThrow SFAssInitThrowReds[OF reds _ _ init] by auto next case True obtain v2 where e\<^sub>2: "e\<^sub>2 = Val v2" using val_of_spec[OF Some] by simp then have vs: "v2 = v \ s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF SFAssInitThrow.hyps(1)] by simp then obtain sfs where shC: "sh\<^sub>1 D = \(sfs, Processing)\" using SFAssInitThrow.hyps(4) SFAssInitThrow.prems(2) Some True by(cases e\<^sub>2, auto simp: bconf_def initPD_def dest: sees_method_fun) then show ?thesis using SFAssInitThrow.hyps(5) init_ProcessingE by blast qed qed next case (SFAssThrow e\<^sub>2 s\<^sub>0 e' s\<^sub>2 C F D) show ?case proof(cases "val_of e\<^sub>2") case None then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAssThrow.prems(2) None by simp then have b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \throw e',s\<^sub>2,False\" using SFAssThrow by auto thus ?thesis using SFAssRedsThrow[OF b1] by fast next case (Some a) then show ?thesis using eval_final_same[OF SFAssThrow.hyps(1)] val_of_spec[OF Some] by auto qed next case (SFAssNone e\<^sub>2 s\<^sub>0 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F D) show ?case proof(cases "val_of e\<^sub>2") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" using SFAssNone by simp then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAssNone.prems(2) None by simp then have b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SFAssNone.hyps(2) iconf by auto thus ?thesis using SFAssRedsNone[OF b1 SFAssNone.hyps(3)] by fast next case (Some a) then obtain b1 where b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b1\" by (metis (no_types, lifting) SFAssNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) thus ?thesis using SFAssRedsNone[OF b1 SFAssNone.hyps(3)] by fast qed next case (SFAssNonStatic e\<^sub>2 s\<^sub>0 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F t D) show ?case proof(cases "val_of e\<^sub>2") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" using SFAssNonStatic by simp then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAssNonStatic.prems(2) None by simp then have b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SFAssNonStatic.hyps(2) iconf by auto thus ?thesis using SFAssRedsNonStatic[OF b1 SFAssNonStatic.hyps(3)] by fast next case (Some a) then obtain b' where b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b'\" by (metis (no_types, lifting) SFAssNonStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) thus ?thesis using SFAssRedsNonStatic[OF b1 SFAssNonStatic.hyps(3)] by fast qed next case (CallObjThrow e s\<^sub>0 e' s\<^sub>1 M ps) show ?case proof(cases "val_of e") case None then have "iconf (shp s\<^sub>0) e" and "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CallObjThrow.prems by auto then have b1: "P \ \e,s\<^sub>0,b\ \* \throw e',s\<^sub>1,False\" using CallObjThrow.hyps(2) by auto then have "P \ \e\M(ps),s\<^sub>0,b\ \* \throw e',s\<^sub>1,False\" using CallObjThrow None by(auto dest!:eval_final simp: CallRedsThrowObj[OF b1]) then show ?thesis by fast next case (Some a) then show ?thesis using eval_final_same[OF CallObjThrow.hyps(1)] val_of_spec[OF Some] by auto qed next case (CallNull e s\<^sub>0 s\<^sub>1 ps vs s\<^sub>2 M) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using CallNull.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CallNull.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \null,s\<^sub>1,False\" using CallNull.hyps(2)[OF iconf] by auto have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \null\M(ps),s\<^sub>1,False\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallNull by auto have "P,shp s\<^sub>1 \\<^sub>b (ps,False) \" by(simp add: bconfs_def) then have b2: "P \ \ps,s\<^sub>1,False\ [\]* \map Val vs,s\<^sub>2,False\" using CallNull.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsNull[OF b1 b2] by fast next case (Some a') then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \null,s\<^sub>1,b1\" by (metis (no_types, lifting) CallNull.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\M(ps),s\<^sub>0,b\ \* \null\M(ps),s\<^sub>1,b1\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf fass] CallNull by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using CallNull.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (ps,b1) \" using Red_preserves_bconf[OF wwf fass CallNull.prems] by simp then have b2: "P \ \ps,s\<^sub>1,b1\ [\]* \map Val vs,s\<^sub>2,False\" using CallNull.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsNull[OF b1 b2] by fast qed next case (CallParamsThrow e s\<^sub>0 v s\<^sub>1 es vs ex es' s\<^sub>2 M) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using CallParamsThrow.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CallParamsThrow.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \Val v,s\<^sub>1,False\" using CallParamsThrow.hyps(2)[OF iconf] by auto have call: "P \ \e\M(es),s\<^sub>0,b\ \* \Val v\M(es),s\<^sub>1,False\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) es" using Red_preserves_iconf[OF wwf call] None CallParamsThrow by auto have "P,shp s\<^sub>1 \\<^sub>b (es,False) \" by(simp add: bconfs_def) then have b2: "P \ \es,s\<^sub>1,False\ [\]* \map Val vs @ throw ex # es',s\<^sub>2,False\" using CallParamsThrow.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsThrowParams[OF b1 b2] by fast next case (Some a') then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \Val v,s\<^sub>1,b1\" by (metis (no_types, lifting) CallParamsThrow.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\M(es),s\<^sub>0,b\ \* \Val v\M(es),s\<^sub>1,b1\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) es" using Red_preserves_iconf[OF wwf fass] CallParamsThrow by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (es,b) \" using CallParamsThrow.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (es,b1) \" using Red_preserves_bconf[OF wwf fass CallParamsThrow.prems] by simp then have b2: "P \ \es,s\<^sub>1,b1\ [\]* \map Val vs @ throw ex # es',s\<^sub>2,False\" using CallParamsThrow.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsThrowParams[OF b1 b2] by fast qed next case (CallNone e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using CallNone.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CallNone.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using CallNone.hyps(2)[OF iconf] by auto have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,False\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallNone by auto have "P,shp s\<^sub>1 \\<^sub>b (ps,False) \" by(simp add: bconfs_def) then have b2: "P \ \ps,s\<^sub>1,False\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using CallNone.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsNone[OF b1 b2 _ CallNone.hyps(6)] CallNone.hyps(5) by fastforce next case (Some a') then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) CallNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,b1\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf fass] CallNone by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using CallNone.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (ps,b1) \" using Red_preserves_bconf[OF wwf fass CallNone.prems] by simp then have b2: "P \ \ps,s\<^sub>1,b1\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using CallNone.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsNone[OF b1 b2 _ CallNone.hyps(6)] CallNone.hyps(5) by fastforce qed next case (CallStatic e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T m D) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using CallStatic.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CallStatic.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using CallStatic.hyps(2)[OF iconf] by auto have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,False\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallStatic by auto have "P,shp s\<^sub>1 \\<^sub>b (ps,False) \" by(simp add: bconfs_def) then have b2: "P \ \ps,s\<^sub>1,False\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using CallStatic.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsStatic[OF b1 b2 _ CallStatic.hyps(6)] CallStatic.hyps(5) by fastforce next case (Some a') then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) CallStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,b1\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] CallStatic by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using CallStatic.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (ps,b1) \" using Red_preserves_bconf[OF wwf call CallStatic.prems] by simp then have b2: "P \ \ps,s\<^sub>1,b1\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using CallStatic.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsStatic[OF b1 b2 _ CallStatic.hyps(6)] CallStatic.hyps(5) by fastforce qed next case (Call e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using Call.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using Call.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using Call.hyps(2)[OF iconf] by auto have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,False\" by(rule CallRedsObj[OF b1]) then have iconf2: "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None Call by auto have "P,shp s\<^sub>1 \\<^sub>b (ps,False) \" by(simp add: bconfs_def) then have b2: "P \ \ps,s\<^sub>1,False\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using Call.hyps(4)[OF iconf2] by simp have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf Call.hyps(6)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b3: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using Call.hyps(10)[OF iconf3] by simp show ?thesis by(rule CallRedsFinal[OF wwf b1 b2 Call.hyps(5-8) b3 eval_final[OF Call.hyps(9)]]) next case (Some a') then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) Call.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,b1\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] Call by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using Call.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (ps,b1) \" using Red_preserves_bconf[OF wwf call Call.prems] by simp then have b2: "P \ \ps,s\<^sub>1,b1\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using Call.hyps(4)[OF iconf2'] by auto have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf Call.hyps(6)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b3: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using Call.hyps(10)[OF iconf3] by simp show ?thesis by(rule CallRedsFinal[OF wwf b1 b2 Call.hyps(5-8) b3 eval_final[OF Call.hyps(9)]]) qed next case (SCallParamsThrow es s\<^sub>0 vs ex es' s\<^sub>2 C M) show ?case proof(cases "map_vals_of es") case None then have iconf: "iconfs (shp s\<^sub>0) es" using SCallParamsThrow.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (es,b) \" using SCallParamsThrow.prems(2) None by simp then have b1: "P \ \es,s\<^sub>0,b\ [\]* \map Val vs @ throw ex # es',s\<^sub>2,False\" using SCallParamsThrow.hyps(2)[OF iconf] by simp show ?thesis using SCallRedsThrowParams[OF b1] by simp next case (Some vs') then have "es = map Val vs'" by(rule map_vals_of_spec) then show ?thesis using evals_finals_same[OF _ SCallParamsThrow.hyps(1)] map_Val_nthrow_neq by auto qed next case (SCallNone ps s\<^sub>0 vs s\<^sub>2 C M) show ?case proof(cases "map_vals_of ps") case None then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallNone.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using SCallNone.prems(2) None by simp then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,s\<^sub>2,False\" using SCallNone.hyps(2)[OF iconf] by auto then show ?thesis using SCallRedsNone[OF b1 SCallNone.hyps(3)] SCallNone.hyps(1) by simp next case (Some vs') then have ps: "ps = map Val vs'" by(rule map_vals_of_spec) then have s\<^sub>0: "s\<^sub>0 = s\<^sub>2" using SCallNone.hyps(1) evals_finals_same by blast then show ?thesis using RedSCallNone[OF SCallNone.hyps(3)] ps by(cases s\<^sub>2, auto) qed next case (SCallNonStatic ps s\<^sub>0 vs s\<^sub>2 C M Ts T m D) show ?case proof(cases "map_vals_of ps") case None then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallNonStatic.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using SCallNonStatic.prems(2) None by simp then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,s\<^sub>2,False\" using SCallNonStatic.hyps(2)[OF iconf] by auto then show ?thesis using SCallRedsNonStatic[OF b1 SCallNonStatic.hyps(3)] SCallNonStatic.hyps(1) by simp next case (Some vs') then have ps: "ps = map Val vs'" by(rule map_vals_of_spec) then have s\<^sub>0: "s\<^sub>0 = s\<^sub>2" using SCallNonStatic.hyps(1) evals_finals_same by blast then show ?thesis using RedSCallNonStatic[OF SCallNonStatic.hyps(3)] ps by(cases s\<^sub>2, auto) qed next case (SCallInitThrow ps s\<^sub>0 vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D a s') show ?case proof(cases "map_vals_of ps") case None then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallInitThrow.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using SCallInitThrow.prems(2) None by simp then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SCallInitThrow.hyps(2)[OF iconf] by auto have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \\<^sub>b (INIT D ([D],False) \ unit,False) \" by(simp add: bconf_def) then have b2: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \throw a,s',False\" using SCallInitThrow.hyps(7) by auto then show ?thesis using SCallInitThrowReds[OF wwf b1 SCallInitThrow.hyps(3-5)] by(cases s', auto) next case (Some vs') have ps: "ps = map Val vs'" by(rule map_vals_of_spec[OF Some]) then have vs: "vs = vs' \ s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using evals_finals_same[OF _ SCallInitThrow.hyps(1)] map_Val_eq by auto show ?thesis proof(cases b) case True obtain sfs where shC: "sh\<^sub>1 D = \(sfs, Processing)\" using SCallInitThrow.hyps(3,4) SCallInitThrow.prems(2) True Some vs by(auto simp: bconf_def initPD_def dest: sees_method_fun) then show ?thesis using init_ProcessingE[OF _ SCallInitThrow.hyps(6)] by blast next case False then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using ps vs by simp have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \\<^sub>b (INIT D ([D],False) \ unit,False) \" by(simp add: bconf_def) then have b2: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \throw a,s',False\" using SCallInitThrow.hyps(7) by auto then show ?thesis using SCallInitThrowReds[OF wwf b1 SCallInitThrow.hyps(3-5)] by(cases s', auto) qed qed next case (SCallInit ps s\<^sub>0 vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) show ?case proof(cases "map_vals_of ps") case None then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallInit.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using SCallInit.prems(2) None by simp then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SCallInit.hyps(2)[OF iconf] by auto have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \\<^sub>b (INIT D ([D],False) \ unit,False) \" by(simp add: bconf_def) then have b2: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SCallInit.hyps(7) by auto have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b3: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using SCallInit.hyps(11)[OF iconf3] by simp show ?thesis by(rule SCallInitReds[OF wwf b1 SCallInit.hyps(3-5) b2 SCallInit.hyps(8-9) b3 eval_final[OF SCallInit.hyps(10)]]) next case (Some vs') then have ps: "ps = map Val vs'" by(rule map_vals_of_spec) then have vs: "vs = vs' \ s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using evals_finals_same[OF _ SCallInit.hyps(1)] map_Val_eq by auto show ?thesis proof(cases b) case True then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\" using ps vs by simp obtain sfs where shC: "sh\<^sub>1 D = \(sfs, Processing)\" using SCallInit.hyps(3,4) SCallInit.prems(2) True Some vs by(auto simp: bconf_def initPD_def dest: sees_method_fun) then have s': "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)" using init_ProcessingE[OF _ SCallInit.hyps(6)] by simp have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b3: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using SCallInit.hyps(11)[OF iconf3] by simp then have b3': "P \ \body,(h\<^sub>1, l\<^sub>2', sh\<^sub>1),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using s' by simp then show ?thesis using SCallInitProcessingReds[OF wwf b1 SCallInit.hyps(3) shC SCallInit.hyps(8-9) b3' eval_final[OF SCallInit.hyps(10)]] s' by simp next case False then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using ps vs by simp have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \\<^sub>b (INIT D ([D],False) \ unit,False) \" by(simp add: bconf_def) then have b2: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SCallInit.hyps(7) by auto have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b3: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using SCallInit.hyps(11)[OF iconf3] by simp show ?thesis by(rule SCallInitReds[OF wwf b1 SCallInit.hyps(3-5) b2 SCallInit.hyps(8-9) b3 eval_final[OF SCallInit.hyps(10)]]) qed qed next case (SCall ps s\<^sub>0 vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M Ts T pns body D sfs l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) show ?case proof(cases "map_vals_of ps") case None then have iconf: "iconfs (shp s\<^sub>0) ps" using SCall.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using SCall.prems(2) None by simp then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SCall.hyps(2)[OF iconf] by auto have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCall.hyps(3)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b2: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using SCall.hyps(8)[OF iconf3] by simp show ?thesis by(rule SCallRedsFinal[OF wwf b1 SCall.hyps(3-6) b2 eval_final[OF SCall.hyps(7)]]) next case (Some vs') then have ps: "ps = map Val vs'" by(rule map_vals_of_spec) then have vs: "vs = vs' \ s\<^sub>0 = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)" using evals_finals_same[OF _ SCall.hyps(1)] map_Val_eq by auto then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b\" using ps by simp have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCall.hyps(3)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b2: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using SCall.hyps(8)[OF iconf3] by simp show ?thesis by(rule SCallRedsFinal[OF wwf b1 SCall.hyps(3-6) b2 eval_final[OF SCall.hyps(7)]]) qed next case (Block e\<^sub>0 h\<^sub>0 l\<^sub>0 V sh\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 T) have iconf: "iconf (shp (h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0)) e\<^sub>0" using Block.prems(1) by (auto simp: assigned_def) have bconf: "P,shp (h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0) \\<^sub>b (e\<^sub>0,b) \" using Block.prems(2) by(auto simp: bconf_def) then have b': "P \ \e\<^sub>0,(h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0),b\ \* \e\<^sub>1,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using Block.hyps(2)[OF iconf] by auto have fin: "final e\<^sub>1" using Block by(auto dest: eval_final) thus ?case using BlockRedsFinal[OF b' fin] by simp next case (Seq e\<^sub>0 s\<^sub>0 v s\<^sub>1 e\<^sub>1 e\<^sub>2 s\<^sub>2) then have iconf: "iconf (shp s\<^sub>0) e\<^sub>0" using Seq.prems(1) by(auto dest: val_of_spec lass_val_of_spec) have b1: "\b1. P \ \e\<^sub>0,s\<^sub>0,b\ \* \Val v,s\<^sub>1,b1\" proof(cases "val_of e\<^sub>0") case None show ?thesis proof(cases "lass_val_of e\<^sub>0") case lNone:None then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>0,b) \" using Seq.prems(2) None by simp then have "P \ \e\<^sub>0,s\<^sub>0,b\ \* \Val v,s\<^sub>1,False\" using iconf Seq.hyps(2) by auto then show ?thesis by fast next case (Some p) obtain V' v' where p: "p = (V',v')" and e\<^sub>0: "e\<^sub>0 = V':=Val v'" using lass_val_of_spec[OF Some] by(cases p, auto) obtain h l sh h' l' sh' where s\<^sub>0: "s\<^sub>0 = (h,l,sh)" and s\<^sub>1: "s\<^sub>1 = (h',l',sh')" by(cases s\<^sub>0, cases s\<^sub>1) then have eval: "P \ \e\<^sub>0,(h,l,sh)\ \ \Val v,(h',l',sh')\" using Seq.hyps(1) by simp then have s\<^sub>1': "Val v = unit \ h' = h \ l' = l(V' \ v') \ sh' = sh" using lass_val_of_eval[OF Some eval] p e\<^sub>0 by simp then have "P \ \e\<^sub>0,s\<^sub>0,b\ \ \Val v,s\<^sub>1,b\" using e\<^sub>0 s\<^sub>0 s\<^sub>1 by(auto intro: RedLAss) then show ?thesis by auto qed next case (Some a) then have "e\<^sub>0 = Val v" and "s\<^sub>0 = s\<^sub>1" using Seq.hyps(1) eval_cases(3) val_of_spec by blast+ then show ?thesis using Seq by auto qed then obtain b1 where b1': "P \ \e\<^sub>0,s\<^sub>0,b\ \* \Val v,s\<^sub>1,b1\" by clarsimp have seq: "P \ \e\<^sub>0;;e\<^sub>1,s\<^sub>0,b\ \* \Val v;;e\<^sub>1,s\<^sub>1,b1\" by(rule SeqReds[OF b1']) then have iconf2: "iconf (shp s\<^sub>1) e\<^sub>1" using Red_preserves_iconf[OF wwf seq] Seq nsub_RI_iconf by auto have "P,shp s\<^sub>1 \\<^sub>b (Val v;; e\<^sub>1,b1) \" by(rule Red_preserves_bconf[OF wwf seq Seq.prems]) then have bconf2: "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>1,b1) \" by simp have b2: "P \ \e\<^sub>1,s\<^sub>1,b1\ \* \e\<^sub>2,s\<^sub>2,False\" by(rule Seq.hyps(4)[OF iconf2 bconf2]) then show ?case using SeqReds2[OF b1' b2] by fast next case (SeqThrow e\<^sub>0 s\<^sub>0 a s\<^sub>1 e\<^sub>1 b) have notVal: "val_of e\<^sub>0 = None" "lass_val_of e\<^sub>0 = None" using SeqThrow.hyps(1) eval_throw_nonVal eval_throw_nonLAss by auto thus ?case using SeqThrow notVal by(auto dest!:eval_final dest: SeqRedsThrow) next case (CondT e s\<^sub>0 s\<^sub>1 e\<^sub>1 e' s\<^sub>2 e\<^sub>2) then have iconf: "iconf (shp s\<^sub>0) e" using CondT.prems(1) by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CondT.prems(2) by auto then have b1: "P \ \e,s\<^sub>0,b\ \* \true,s\<^sub>1,False\" using iconf CondT.hyps(2) by auto have cond: "P \ \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0,b\ \* \if (true) e\<^sub>1 else e\<^sub>2,s\<^sub>1,False\" by(rule CondReds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>1" using Red_preserves_iconf[OF wwf cond] CondT nsub_RI_iconf by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>1,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>1,s\<^sub>1,False\ \* \e',s\<^sub>2,False\" using CondT.hyps(4)[OF iconf2'] by auto then show ?case using CondReds2T[OF b1 b2] by fast next case (CondF e s\<^sub>0 s\<^sub>1 e\<^sub>2 e' s\<^sub>2 e\<^sub>1) then have iconf: "iconf (shp s\<^sub>0) e" using CondF.prems(1) by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CondF.prems(2) by auto then have b1: "P \ \e,s\<^sub>0,b\ \* \false,s\<^sub>1,False\" using iconf CondF.hyps(2) by auto have cond: "P \ \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0,b\ \* \if (false) e\<^sub>1 else e\<^sub>2,s\<^sub>1,False\" by(rule CondReds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf cond] CondF nsub_RI_iconf by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \e',s\<^sub>2,False\" using CondF.hyps(4)[OF iconf2'] by auto then show ?case using CondReds2F[OF b1 b2] by fast next case CondThrow thus ?case by(auto dest!:eval_final dest:CondRedsThrow) next case (WhileF e s\<^sub>0 s\<^sub>1 c) then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using WhileF.prems(2) by(simp add: bconf_def) then have b': "P \ \e,s\<^sub>0,b\ \* \false,s\<^sub>1,False\" using WhileF.hyps(2) iconf by auto thus ?case using WhileFReds[OF b'] by fast next case (WhileT e s\<^sub>0 s\<^sub>1 c v\<^sub>1 s\<^sub>2 e\<^sub>3 s\<^sub>3) then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using WhileT.prems(2) by(simp add: bconf_def) then have b1: "P \ \e,s\<^sub>0,b\ \* \true,s\<^sub>1,False\" using WhileT.hyps(2) iconf by auto have iconf2: "iconf (shp s\<^sub>1) c" using WhileT.prems(1) nsub_RI_iconf by auto have bconf2: "P,shp s\<^sub>1 \\<^sub>b (c,False) \" by(simp add: bconf_def) then have b2: "P \ \c,s\<^sub>1,False\ \* \Val v\<^sub>1,s\<^sub>2,False\" using WhileT.hyps(4) iconf2 by auto have iconf3: "iconf (shp s\<^sub>2) (while (e) c)" using WhileT.prems(1) by auto have "P,shp s\<^sub>2 \\<^sub>b (while (e) c,False) \" by(simp add: bconf_def) then have b3: "P \ \while (e) c,s\<^sub>2,False\ \* \e\<^sub>3,s\<^sub>3,False\" using WhileT.hyps(6) iconf3 by auto show ?case using WhileTReds[OF b1 b2 b3] by fast next case WhileCondThrow thus ?case by (metis (no_types, lifting) WhileRedsThrow iconf.simps(16) bconf_While bconf_def nsub_RI_iconf) next case (WhileBodyThrow e s\<^sub>0 s\<^sub>1 c e' s\<^sub>2) then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using WhileBodyThrow.prems(2) by(simp add: bconf_def) then have b1: "P \ \e,s\<^sub>0,b\ \* \true,s\<^sub>1,False\" using WhileBodyThrow.hyps(2) iconf by auto have iconf2: "iconf (shp s\<^sub>1) c" using WhileBodyThrow.prems(1) nsub_RI_iconf by auto have bconf2: "P,shp s\<^sub>1 \\<^sub>b (c,False) \" by(simp add: bconf_def) then have b2: "P \ \c,s\<^sub>1,False\ \* \throw e',s\<^sub>2,False\" using WhileBodyThrow.hyps(4) iconf2 by auto show ?case using WhileTRedsThrow[OF b1 b2] by fast next case Throw thus ?case by (meson ThrowReds iconf.simps(17) bconf_Throw) next case ThrowNull thus ?case by (meson ThrowRedsNull iconf.simps(17) bconf_Throw) next case ThrowThrow thus ?case by (meson ThrowRedsThrow iconf.simps(17) bconf_Throw) next case Try thus ?case by (meson TryRedsVal iconf.simps(18) bconf_Try) next case (TryCatch e\<^sub>1 s\<^sub>0 a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e\<^sub>2' h\<^sub>2 l\<^sub>2 sh\<^sub>2) then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Throw a,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" by auto have Try: "P \ \try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0,b\ \* \try (Throw a) catch(C V) e\<^sub>2,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" by(rule TryReds[OF b1]) have iconf: "iconf sh\<^sub>1 e\<^sub>2" using Red_preserves_iconf[OF wwf Try] TryCatch nsub_RI_iconf by auto have bconf: "P,shp (h\<^sub>1, l\<^sub>1(V \ Addr a), sh\<^sub>1) \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,(h\<^sub>1, l\<^sub>1(V \ Addr a), sh\<^sub>1),False\ \* \e\<^sub>2',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using TryCatch.hyps(6) iconf by auto thus ?case using TryCatchRedsFinal[OF b1] TryCatch.hyps(3-5) by (meson eval_final) next case TryThrow thus ?case by (meson TryRedsFail iconf.simps(18) bconf_Try) next case Nil thus ?case by(auto simp: bconfs_def) next case (Cons e s\<^sub>0 v s\<^sub>1 es es' s\<^sub>2) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using Cons.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using Cons.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \Val v,s\<^sub>1,False\" using Cons.hyps(2) iconf by auto have cons: "P \ \e # es,s\<^sub>0,b\ [\]* \Val v # es,s\<^sub>1,False\" by(rule ListReds1[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) es" using Reds_preserves_iconf[OF wwf cons] None Cons by auto have "P,shp s\<^sub>1 \\<^sub>b (es,False) \" by(simp add: bconfs_def) then have b2: "P \ \es,s\<^sub>1,False\ [\]* \es',s\<^sub>2,False\" using Cons.hyps(4)[OF iconf2'] by auto show ?thesis using ListRedsVal[OF b1 b2] by auto next case (Some a) then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \Val v,s\<^sub>1,b1\" by (metis (no_types, lifting) Cons.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have cons: "P \ \e # es,s\<^sub>0,b\ [\]* \Val v # es,s\<^sub>1,b1\" by(rule ListReds1[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) es" using Reds_preserves_iconf[OF wwf cons] Cons by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (es,b) \" using Cons.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (es,b1) \" using Reds_preserves_bconf[OF wwf cons Cons.prems] by simp then have b2: "P \ \es,s\<^sub>1,b1\ [\]* \es',s\<^sub>2,False\" using Cons.hyps(4)[OF iconf2'] by auto show ?thesis using ListRedsVal[OF b1 b2] by auto qed next case (ConsThrow e s\<^sub>0 e' s\<^sub>1 es) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using ConsThrow.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using ConsThrow.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \throw e',s\<^sub>1,False\" using ConsThrow.hyps(2) iconf by auto have cons: "P \ \e # es,s\<^sub>0,b\ [\]* \throw e' # es,s\<^sub>1,False\" by(rule ListReds1[OF b1]) then show ?thesis by fast next case (Some a) then show ?thesis using eval_final_same[OF ConsThrow.hyps(1)] val_of_spec[OF Some] by auto qed next case (InitFinal e s e' s' C b') then have "\sub_RI e" by simp then show ?case using InitFinal RedInit[of e C b' s b P] by (meson converse_rtrancl_into_rtrancl nsub_RI_iconf red_preserves_bconf RedInit) next case (InitNone sh C C' Cs e h l e' s') then have init: "P \ \INIT C' (C # Cs,False) \ e,(h, l, sh(C \ (sblank P C, Prepared))),b\ \* \e',s',False\" by(simp add: bconf_def) show ?case by(rule InitNoneReds[OF InitNone.hyps(1) init]) next case (InitDone sh C sfs C' Cs e h l e' s') then have "iconf (shp (h, l, sh)) (INIT C' (Cs,True) \ e)" using InitDone.hyps(1) proof(cases Cs) case Nil then have "C = C'" "\sub_RI e" using InitDone.prems(1) by simp+ then show ?thesis using Nil InitDone.hyps(1) by(simp add: initPD_def) qed(auto) then have init: "P \ \INIT C' (Cs,True) \ e,(h, l, sh),b\ \* \e',s',False\" using InitDone by(simp add: bconf_def) show ?case by(rule InitDoneReds[OF InitDone.hyps(1) init]) next case (InitProcessing sh C sfs C' Cs e h l e' s') then have "iconf (shp (h, l, sh)) (INIT C' (Cs,True) \ e)" using InitProcessing.hyps(1) proof(cases Cs) case Nil then have "C = C'" "\sub_RI e" using InitProcessing.prems(1) by simp+ then show ?thesis using Nil InitProcessing.hyps(1) by(simp add: initPD_def) qed(auto) then have init: "P \ \INIT C' (Cs,True) \ e,(h, l, sh),b\ \* \e',s',False\" using InitProcessing by(simp add: bconf_def) show ?case by(rule InitProcessingReds[OF InitProcessing.hyps(1) init]) next case InitError thus ?case by(fastforce intro: InitErrorReds simp: bconf_def) next case InitObject thus ?case by(fastforce intro: InitObjectReds simp: bconf_def) next case InitNonObject thus ?case by(fastforce intro: InitNonObjectReds simp: bconf_def) next case InitRInit thus ?case by(fastforce intro: RedsInitRInit simp: bconf_def) next case (RInit e s v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 s\<^sub>1) then have iconf2: "iconf (shp (h', l', sh'')) (INIT C' (Cs,True) \ e')" by(auto simp: initPD_def fun_upd_same list_nonempty_induct) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s) e" using RInit.prems(1) by simp have bconf: "P,shp s \\<^sub>b (e,b) \" using RInit.prems(2) None by simp then have b1: "P \ \e,s,b\ \* \Val v,(h',l',sh'),False\" using RInit.hyps(2)[OF iconf] by auto have "P,shp (h', l', sh'') \\<^sub>b (INIT C' (Cs,True) \ e',False) \" by(simp add: bconf_def) then have b2: "P \ \INIT C' (Cs,True) \ e',(h',l',sh''),False\ \* \e\<^sub>1,s\<^sub>1,False\" using RInit.hyps(7)[OF iconf2] by auto then show ?thesis using RedsRInit[OF b1 RInit.hyps(3-5) b2] by fast next case (Some a') then obtain b1 where b1: "P \ \e,s,b\ \* \Val v,(h',l',sh'),b1\" by (metis (no_types, lifting) RInit.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fin: "final e" by(simp add: val_of_spec[OF Some]) have "\b" using RInit.prems(2) Some by(simp add: bconf_def) then have nb1: "\b1" using reds_final_same[OF b1 fin] by simp have "P,shp (h', l', sh'') \\<^sub>b (INIT C' (Cs,True) \ e',b1) \" using nb1 by(simp add: bconf_def) then have b2: "P \ \INIT C' (Cs,True) \ e',(h', l', sh''),b1\ \* \e\<^sub>1,s\<^sub>1,False\" using RInit.hyps(7)[OF iconf2] by auto then show ?thesis using RedsRInit[OF b1 RInit.hyps(3-5) b2] by fast qed next case (RInitInitFail e s a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 s\<^sub>1) have fin: "final (throw a)" using eval_final[OF RInitInitFail.hyps(1)] by simp then obtain a' where a': "throw a = Throw a'" by auto have iconf2: "iconf (shp (h', l', sh'')) (RI (D,Throw a') ; Cs \ e')" using RInitInitFail.prems(1) by auto show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s) e" using RInitInitFail.prems(1) by simp have bconf: "P,shp s \\<^sub>b (e,b) \" using RInitInitFail.prems(2) None by simp then have b1: "P \ \e,s,b\ \* \Throw a',(h',l',sh'),False\" using RInitInitFail.hyps(2)[OF iconf] a' by auto have "P,shp (h', l', sh'') \\<^sub>b (RI (D,Throw a') ; Cs \ e',False) \" by(simp add: bconf_def) then have b2: "P \ \RI (D,Throw a') ; Cs \ e',(h',l',sh''),False\ \* \e\<^sub>1,s\<^sub>1,False\" using RInitInitFail.hyps(6) iconf2 a' by auto show ?thesis using RInitInitThrowReds[OF b1 RInitInitFail.hyps(3-4) b2] by fast next case (Some a1) then obtain b1 where b1: "P \ \e,s,b\ \* \Throw a',(h',l',sh'),b1\" using a' by (metis (no_types, lifting) RInitInitFail.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fin: "final e" by(simp add: val_of_spec[OF Some]) have "\b" using RInitInitFail.prems(2) Some by(simp add: bconf_def) then have nb1: "\b1" using reds_final_same[OF b1 fin] by simp have "P,shp (h', l', sh'') \\<^sub>b (RI (D,Throw a') ; Cs \ e',b1) \" using nb1 by(simp add: bconf_def) then have b2: "P \ \RI (D,Throw a') ; Cs \ e',(h', l', sh''),b1\ \* \e\<^sub>1,s\<^sub>1,False\" using RInitInitFail.hyps(6) iconf2 a' by auto show ?thesis using RInitInitThrowReds[OF b1 RInitInitFail.hyps(3-4) b2] by fast qed next case (RInitFailFinal e s a h' l' sh' C sfs i sh'' e') have fin: "final (throw a)" using eval_final[OF RInitFailFinal.hyps(1)] by simp then obtain a' where a': "throw a = Throw a'" by auto show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s) e" using RInitFailFinal.prems(1) by simp have bconf: "P,shp s \\<^sub>b (e,b) \" using RInitFailFinal.prems(2) None by simp then have b1: "P \ \e,s,b\ \* \Throw a',(h',l',sh'),False\" using RInitFailFinal.hyps(2)[OF iconf] a' by auto show ?thesis using RInitThrowReds[OF b1 RInitFailFinal.hyps(3-4)] a' by fast next case (Some a1) then show ?thesis using eval_final_same[OF RInitFailFinal.hyps(1)] val_of_spec[OF Some] by auto qed qed (*>*) subsection\Big steps simulates small step\ text\ This direction was carried out by Norbert Schirmer and Daniel Wasserrab (and modified to include statics and DCI by Susannah Mansky). \ text \ The big step equivalent of @{text RedWhile}: \ lemma unfold_while: "P \ \while(b) c,s\ \ \e',s'\ = P \ \if(b) (c;;while(b) c) else (unit),s\ \ \e',s'\" (*<*) proof assume "P \ \while (b) c,s\ \ \e',s'\" thus "P \ \if (b) (c;; while (b) c) else unit,s\ \ \e',s'\" by cases (fastforce intro: eval_evals.intros)+ next assume "P \ \if (b) (c;; while (b) c) else unit,s\ \ \e',s'\" thus "P \ \while (b) c,s\ \ \e',s'\" proof (cases) fix a assume e': "e' = throw a" assume "P \ \b,s\ \ \throw a,s'\" hence "P \ \while(b) c,s\ \ \throw a,s'\" by (rule WhileCondThrow) with e' show ?thesis by simp next fix s\<^sub>1 assume eval_false: "P \ \b,s\ \ \false,s\<^sub>1\" and eval_unit: "P \ \unit,s\<^sub>1\ \ \e',s'\" with eval_unit have "s' = s\<^sub>1" "e' = unit" by (auto elim: eval_cases) moreover from eval_false have "P \ \while (b) c,s\ \ \unit,s\<^sub>1\" by - (rule WhileF, simp) ultimately show ?thesis by simp next fix s\<^sub>1 assume eval_true: "P \ \b,s\ \ \true,s\<^sub>1\" and eval_rest: "P \ \c;; while (b) c,s\<^sub>1\\\e',s'\" from eval_rest show ?thesis proof (cases) fix s\<^sub>2 v\<^sub>1 assume "P \ \c,s\<^sub>1\ \ \Val v\<^sub>1,s\<^sub>2\" "P \ \while (b) c,s\<^sub>2\ \ \e',s'\" with eval_true show "P \ \while(b) c,s\ \ \e',s'\" by (rule WhileT) next fix a assume "P \ \c,s\<^sub>1\ \ \throw a,s'\" "e' = throw a" with eval_true show "P \ \while(b) c,s\ \ \e',s'\" by (iprover intro: WhileBodyThrow) qed qed qed (*>*) lemma blocksEval: "\Ts vs l l'. \size ps = size Ts; size ps = size vs; P \ \blocks(ps,Ts,vs,e),(h,l,sh)\ \ \e',(h',l',sh')\ \ \ \ l''. P \ \e,(h,l(ps[\]vs),sh)\ \ \e',(h',l'',sh')\" (*<*) proof (induct ps) case Nil then show ?case by fastforce next case (Cons p ps') have length_eqs: "length (p # ps') = length Ts" "length (p # ps') = length vs" by fact+ then obtain T Ts' where Ts: "Ts = T#Ts'" by (cases "Ts") simp obtain v vs' where vs: "vs = v#vs'" using length_eqs by (cases "vs") simp have "P \ \blocks (p # ps', Ts, vs, e),(h,l,sh)\ \ \e',(h', l',sh')\" by fact with Ts vs have "P \ \{p:T := Val v; blocks (ps', Ts', vs', e)},(h,l,sh)\ \ \e',(h', l',sh')\" by simp then obtain l''' where eval_ps': "P \ \blocks (ps', Ts', vs', e),(h, l(p\v), sh)\ \ \e',(h', l''', sh')\" and l''': "l'=l'''(p:=l p)" by (auto elim!: eval_cases) then obtain l'' where - hyp: "P \ \e,(h, l(p\v)(ps'[\]vs'), sh)\ \ \e',(h', l'', sh')\" + hyp: "P \ \e,(h, l(p\v, ps'[\]vs'), sh)\ \ \e',(h', l'', sh')\" using length_eqs Ts vs Cons.hyps [OF _ _ eval_ps'] by auto from hyp show "\l''. P \ \e,(h, l(p # ps'[\]vs), sh)\ \ \e',(h', l'', sh')\" using Ts vs by auto qed (*>*) lemma assumes wf: "wwf_J_prog P" shows eval_restrict_lcl: "P \ \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ (\W. fv e \ W \ P \ \e,(h,l|`W,sh)\ \ \e',(h',l'|`W,sh')\)" and "P \ \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ (\W. fvs es \ W \ P \ \es,(h,l|`W,sh)\ [\] \es',(h',l'|`W,sh')\)" (*<*) proof(induct rule:eval_evals_inducts) case (Block e\<^sub>0 h\<^sub>0 l\<^sub>0 V sh\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 T) have IH: "\W. fv e\<^sub>0 \ W \ P \ \e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None)|`W,sh\<^sub>0)\ \ \e\<^sub>1,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\" by fact have "fv({V:T; e\<^sub>0}) \ W" by fact+ hence "fv e\<^sub>0 - {V} \ W" by simp_all hence "fv e\<^sub>0 \ insert V W" by fast from IH[OF this] have "P \ \e\<^sub>0,(h\<^sub>0, (l\<^sub>0|`W)(V := None), sh\<^sub>0)\ \ \e\<^sub>1,(h\<^sub>1, l\<^sub>1|`insert V W, sh\<^sub>1)\" by fastforce from eval_evals.Block[OF this] show ?case by fastforce next case Seq thus ?case by simp (blast intro:eval_evals.Seq) next case New thus ?case by(simp add:eval_evals.intros) next case NewFail thus ?case by(simp add:eval_evals.intros) next case (NewInit sh C h l v' h' l' sh' a h'') have "fv(INIT C ([C],False) \ unit) \ W" by simp then have "P \ \INIT C ([C],False) \ unit,(h, l |` W, sh)\ \ \Val v',(h', l' |` W, sh')\" by (simp add: NewInit.hyps(3)) thus ?case using NewInit.hyps(1,4-6) eval_evals.NewInit by blast next case (NewInitOOM sh C h l v' h' l' sh') have "fv(INIT C ([C],False) \ unit) \ W" by simp then have "P \ \INIT C ([C],False) \ unit,(h, l |` W, sh)\ \ \Val v',(h', l' |` W, sh')\" by (simp add: NewInitOOM.hyps(3)) thus ?case using NewInitOOM.hyps(1,4,5) eval_evals.NewInitOOM by auto next case NewInitThrow thus ?case by(simp add:eval_evals.intros) next case Cast thus ?case by simp (blast intro:eval_evals.Cast) next case CastNull thus ?case by simp (blast intro:eval_evals.CastNull) next case CastFail thus ?case by simp (blast intro:eval_evals.CastFail) next case CastThrow thus ?case by(simp add:eval_evals.intros) next case Val thus ?case by(simp add:eval_evals.intros) next case BinOp thus ?case by simp (blast intro:eval_evals.BinOp) next case BinOpThrow1 thus ?case by simp (blast intro:eval_evals.BinOpThrow1) next case BinOpThrow2 thus ?case by simp (blast intro:eval_evals.BinOpThrow2) next case Var thus ?case by(simp add:eval_evals.intros) next case (LAss e h\<^sub>0 l\<^sub>0 sh\<^sub>0 v h l sh l' V) have IH: "\W. fv e \ W \ P \ \e,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\ \ \Val v,(h,l|`W,sh)\" and [simp]: "l' = l(V \ v)" by fact+ have "fv (V:=e) \ W" by fact hence fv: "fv e \ W" and VinW: "V \ W" by auto from eval_evals.LAss[OF IH[OF fv] refl, of V] VinW show ?case by simp next case LAssThrow thus ?case by(fastforce intro: eval_evals.LAssThrow) next case FAcc thus ?case by simp (blast intro: eval_evals.FAcc) next case FAccNull thus ?case by(fastforce intro: eval_evals.FAccNull) next case FAccThrow thus ?case by(fastforce intro: eval_evals.FAccThrow) next case FAccNone thus ?case by(metis eval_evals.FAccNone fv.simps(7)) next case FAccStatic thus ?case by(metis eval_evals.FAccStatic fv.simps(7)) next case SFAcc thus ?case by simp (blast intro: eval_evals.SFAcc) next case SFAccInit thus ?case by simp (blast intro: eval_evals.SFAccInit) next case SFAccInitThrow thus ?case by simp (blast intro: eval_evals.SFAccInitThrow) next case SFAccNone thus ?case by simp (blast intro: eval_evals.SFAccNone) next case SFAccNonStatic thus ?case by simp (blast intro: eval_evals.SFAccNonStatic) next case FAss thus ?case by simp (blast intro: eval_evals.FAss) next case FAssNull thus ?case by simp (blast intro: eval_evals.FAssNull) next case FAssThrow1 thus ?case by simp (blast intro: eval_evals.FAssThrow1) next case FAssThrow2 thus ?case by simp (blast intro: eval_evals.FAssThrow2) next case FAssNone thus ?case by simp (blast intro: eval_evals.FAssNone) next case FAssStatic thus ?case by simp (blast intro: eval_evals.FAssStatic) next case SFAss thus ?case by simp (blast intro: eval_evals.SFAss) next case SFAssInit thus ?case by simp (blast intro: eval_evals.SFAssInit) next case SFAssInitThrow thus ?case by simp (blast intro: eval_evals.SFAssInitThrow) next case SFAssThrow thus ?case by simp (blast intro: eval_evals.SFAssThrow) next case SFAssNone thus ?case by simp (blast intro: eval_evals.SFAssNone) next case SFAssNonStatic thus ?case by simp (blast intro: eval_evals.SFAssNonStatic) next case CallObjThrow thus ?case by simp (blast intro: eval_evals.intros) next case CallNull thus ?case by simp (blast intro: eval_evals.CallNull) next case (CallNone e h l sh a h' l' sh' ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M) have f1: "P \ \e,(h, l |` W, sh)\ \ \addr a,(h', l' |` W, sh')\" by (metis (no_types) fv.simps(11) le_sup_iff local.CallNone(2) local.CallNone(7)) have "P \ \ps,(h', l' |` W, sh')\ [\] \map Val vs, (h\<^sub>2, l\<^sub>2 |` W, sh\<^sub>2)\" using local.CallNone(4) local.CallNone(7) by auto then show ?case using f1 eval_evals.CallNone local.CallNone(5) local.CallNone(6) by auto next case CallStatic thus ?case by (metis (no_types, lifting) eval_evals.CallStatic fv.simps(11) le_sup_iff) next case CallParamsThrow thus ?case by simp (blast intro: eval_evals.CallParamsThrow) next case (Call e h\<^sub>0 l\<^sub>0 sh\<^sub>0 a h\<^sub>1 l\<^sub>1 sh\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) have IHe: "\W. fv e \ W \ P \ \e,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\ \ \addr a,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\" and IHps: "\W. fvs ps \ W \ P \ \ps,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\" and IHbd: "\W. fv body \ W \ P \ \body,(h\<^sub>2,l\<^sub>2'|`W,sh\<^sub>2)\ \ \e',(h\<^sub>3,l\<^sub>3|`W,sh\<^sub>3)\" and h\<^sub>2a: "h\<^sub>2 a = Some (C, fs)" and "method": "P \ C sees M,NonStatic: Ts\T = (pns, body) in D" and same_len: "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [this \ Addr a, pns [\] vs]" by fact+ have "fv (e\M(ps)) \ W" by fact hence fve: "fv e \ W" and fvps: "fvs(ps) \ W" by auto have wfmethod: "size Ts = size pns \ this \ set pns" and fvbd: "fv body \ {this} \ set pns" using "method" wf by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ show ?case using IHbd[OF fvbd] l\<^sub>2' same_len wfmethod h\<^sub>2a eval_evals.Call[OF IHe[OF fve] IHps[OF fvps] _ "method" same_len l\<^sub>2'] by (simp add:subset_insertI) next case (SCallNone ps h l sh vs h' l' sh' C M) have "P \ \ps,(h, l |` W, sh)\ [\] \map Val vs,(h', l' |` W, sh')\" using SCallNone.hyps(2) SCallNone.prems by auto then show ?case using SCallNone.hyps(3) eval_evals.SCallNone by auto next case SCallNonStatic thus ?case by (metis eval_evals.SCallNonStatic fv.simps(12)) next case SCallParamsThrow thus ?case by simp (blast intro: eval_evals.SCallParamsThrow) next case SCallInitThrow thus ?case by simp (blast intro: eval_evals.SCallInitThrow) next case SCallInit thus ?case by simp (blast intro: eval_evals.SCallInit) next case (SCall ps h\<^sub>0 l\<^sub>0 sh\<^sub>0 vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M Ts T pns body D sfs l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) have IHps: "\W. fvs ps \ W \ P \ \ps,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\" and IHbd: "\W. fv body \ W \ P \ \body,(h\<^sub>2,l\<^sub>2'|`W,sh\<^sub>2)\ \ \e',(h\<^sub>3,l\<^sub>3|`W,sh\<^sub>3)\" and sh\<^sub>2D: "sh\<^sub>2 D = Some (sfs, Done) \ M = clinit \ sh\<^sub>2 D = \(sfs, Processing)\" and "method": "P \ C sees M,Static: Ts\T = (pns, body) in D" and same_len: "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [pns [\] vs]" by fact+ have "fv (C\\<^sub>sM(ps)) \ W" by fact hence fvps: "fvs(ps) \ W" by auto have wfmethod: "size Ts = size pns" and fvbd: "fv body \ set pns" using "method" wf by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ show ?case using IHbd[OF fvbd] l\<^sub>2' same_len wfmethod sh\<^sub>2D eval_evals.SCall[OF IHps[OF fvps] "method" _ same_len l\<^sub>2'] by (simp add:subset_insertI) next case SeqThrow thus ?case by simp (blast intro: eval_evals.SeqThrow) next case CondT thus ?case by simp (blast intro: eval_evals.CondT) next case CondF thus ?case by simp (blast intro: eval_evals.CondF) next case CondThrow thus ?case by simp (blast intro: eval_evals.CondThrow) next case WhileF thus ?case by simp (blast intro: eval_evals.WhileF) next case WhileT thus ?case by simp (blast intro: eval_evals.WhileT) next case WhileCondThrow thus ?case by simp (blast intro: eval_evals.WhileCondThrow) next case WhileBodyThrow thus ?case by simp (blast intro: eval_evals.WhileBodyThrow) next case Throw thus ?case by simp (blast intro: eval_evals.Throw) next case ThrowNull thus ?case by simp (blast intro: eval_evals.ThrowNull) next case ThrowThrow thus ?case by simp (blast intro: eval_evals.ThrowThrow) next case Try thus ?case by simp (blast intro: eval_evals.Try) next case (TryCatch e\<^sub>1 h\<^sub>0 l\<^sub>0 sh\<^sub>0 a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e\<^sub>2' h\<^sub>2 l\<^sub>2 sh\<^sub>2) have IH\<^sub>1: "\W. fv e\<^sub>1 \ W \ P \ \e\<^sub>1,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\ \ \Throw a,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\" and IH\<^sub>2: "\W. fv e\<^sub>2 \ W \ P \ \e\<^sub>2,(h\<^sub>1,l\<^sub>1(V\Addr a)|`W,sh\<^sub>1)\ \ \e\<^sub>2',(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\" and lookup: "h\<^sub>1 a = Some(D, fs)" and subtype: "P \ D \\<^sup>* C" by fact+ have "fv (try e\<^sub>1 catch(C V) e\<^sub>2) \ W" by fact hence fv\<^sub>1: "fv e\<^sub>1 \ W" and fv\<^sub>2: "fv e\<^sub>2 \ insert V W" by auto have IH\<^sub>2': "P \ \e\<^sub>2,(h\<^sub>1,(l\<^sub>1|`W)(V \ Addr a),sh\<^sub>1)\ \ \e\<^sub>2',(h\<^sub>2,l\<^sub>2|`insert V W,sh\<^sub>2)\" using IH\<^sub>2[OF fv\<^sub>2] fun_upd_restrict[of l\<^sub>1 W] (*FIXME just l|W instead of l|(W-V) in simp rule??*) by simp with eval_evals.TryCatch[OF IH\<^sub>1[OF fv\<^sub>1] _ subtype IH\<^sub>2'] lookup show ?case by fastforce next case TryThrow thus ?case by simp (blast intro: eval_evals.TryThrow) next case Nil thus ?case by (simp add: eval_evals.Nil) next case Cons thus ?case by simp (blast intro: eval_evals.Cons) next case ConsThrow thus ?case by simp (blast intro: eval_evals.ConsThrow) next case InitFinal thus ?case by (simp add: eval_evals.InitFinal) next case InitNone thus ?case by(blast intro: eval_evals.InitNone) next case InitDone thus ?case by (simp add: InitDone.hyps(2) InitDone.prems eval_evals.InitDone) next case InitProcessing thus ?case by (simp add: eval_evals.InitProcessing) next case InitError thus ?case using eval_evals.InitError by auto next case InitObject thus ?case by(simp add: eval_evals.InitObject) next case InitNonObject thus ?case by(simp add: eval_evals.InitNonObject) next case InitRInit thus ?case by(simp add: eval_evals.InitRInit) next case (RInit e h l sh v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) have f1: "fv e \ W \ fv (INIT C' (Cs,True) \ e') \ W" using RInit.prems by auto then have f2: "P \ \e,(h, l |` W, sh)\ \ \Val v,(h', l' |` W, sh')\" using RInit.hyps(2) by blast have "P \ \INIT C' (Cs,True) \ e', (h', l' |` W, sh'')\ \ \e\<^sub>1,(h\<^sub>1, l\<^sub>1 |` W, sh\<^sub>1)\" using f1 by (meson RInit.hyps(7)) then show ?case using f2 RInit.hyps(3) RInit.hyps(4) RInit.hyps(5) eval_evals.RInit by blast next case (RInitInitFail e h l sh a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) have f1: "fv e \ W" "fv e' \ W" using RInitInitFail.prems by auto then have f2: "P \ \e,(h, l |` W, sh)\ \ \throw a,(h', l' |` W, sh')\" using RInitInitFail.hyps(2) by blast then have f2': "fv (throw a) \ W" using eval_final[OF f2] by auto then have f1': "fv (RI (D,throw a);Cs \ e') \ W" using f1 by auto have "P \ \RI (D,throw a);Cs \ e', (h', l' |` W, sh'')\ \ \e\<^sub>1,(h\<^sub>1, l\<^sub>1 |` W, sh\<^sub>1)\" using f1' by (meson RInitInitFail.hyps(6)) then show ?case using f2 by (simp add: RInitInitFail.hyps(3,4) eval_evals.RInitInitFail) next case (RInitFailFinal e h l sh a h' l' sh' sh'' C) have f1: "fv e \ W" using RInitFailFinal.prems by auto then have f2: "P \ \e,(h, l |` W, sh)\ \ \throw a,(h', l' |` W, sh')\" using RInitFailFinal.hyps(2) by blast then have f2': "fv (throw a) \ W" using eval_final[OF f2] by auto then show ?case using f2 RInitFailFinal.hyps(3,4) eval_evals.RInitFailFinal by blast qed (*>*) lemma eval_notfree_unchanged: "P \ \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ (\V. V \ fv e \ l' V = l V)" and "P \ \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ (\V. V \ fvs es \ l' V = l V)" (*<*) proof(induct rule:eval_evals_inducts) case LAss thus ?case by(simp add:fun_upd_apply) next case Block thus ?case by (simp only:fun_upd_apply split:if_splits) fastforce next case TryCatch thus ?case by (simp only:fun_upd_apply split:if_splits) fastforce next case (RInitInitFail e h l sh a h' l' sh' C sfs i sh'' D Cs e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) have "fv (throw a) = {}" using RInitInitFail.hyps(1) eval_final final_fv by blast then have "fv e \ fv (RI (D,throw a) ; Cs \ unit) \ fv (RI (C,e) ; D#Cs \ unit)" by auto then show ?case using RInitInitFail.hyps(2,6) RInitInitFail.prems by fastforce qed simp_all (*>*) lemma eval_closed_lcl_unchanged: "\ P \ \e,(h,l,sh)\ \ \e',(h',l',sh')\; fv e = {} \ \ l' = l" (*<*)by(fastforce dest:eval_notfree_unchanged simp add:fun_eq_iff [where 'b="val option"])(*>*) lemma list_eval_Throw: assumes eval_e: "P \ \throw x,s\ \ \e',s'\" shows "P \ \map Val vs @ throw x # es',s\ [\] \map Val vs @ e' # es',s'\" (*<*) proof - from eval_e obtain a where e': "e' = Throw a" by (cases) (auto dest!: eval_final) { fix es have "\vs. es = map Val vs @ throw x # es' \ P \ \es,s\[\]\map Val vs @ e' # es',s'\" proof (induct es type: list) case Nil thus ?case by simp next case (Cons e es vs) have e_es: "e # es = map Val vs @ throw x # es'" by fact show "P \ \e # es,s\ [\] \map Val vs @ e' # es',s'\" proof (cases vs) case Nil with e_es obtain "e=throw x" "es=es'" by simp moreover from eval_e e' have "P \ \throw x # es,s\ [\] \Throw a # es,s'\" by (iprover intro: ConsThrow) ultimately show ?thesis using Nil e' by simp next case (Cons v vs') have vs: "vs = v # vs'" by fact with e_es obtain e: "e=Val v" and es:"es= map Val vs' @ throw x # es'" by simp from e have "P \ \e,s\ \ \Val v,s\" by (iprover intro: eval_evals.Val) moreover from es have "P \ \es,s\ [\] \map Val vs' @ e' # es',s'\" by (rule Cons.hyps) ultimately show "P \ \e#es,s\ [\] \map Val vs @ e' # es',s'\" using vs by (auto intro: eval_evals.Cons) qed qed } thus ?thesis by simp qed (*>*) \ \ separate evaluation of first subexp of a sequence \ lemma seq_ext: assumes IH: "\e' s'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\" and seq: "P \ \e'' ;; e\<^sub>0,s''\ \ \e',s'\" shows "P \ \e ;; e\<^sub>0,s\ \ \e',s'\" proof(rule eval_cases(14)[OF seq]) \ \ Seq \ fix v' s\<^sub>1 assume e'': "P \ \e'',s''\ \ \Val v',s\<^sub>1\" and estep: "P \ \e\<^sub>0,s\<^sub>1\ \ \e',s'\" have "P \ \e,s\ \ \Val v',s\<^sub>1\" using e'' IH by simp then show ?thesis using estep Seq by simp next fix e\<^sub>t assume e'': "P \ \e'',s''\ \ \throw e\<^sub>t,s'\" and e': "e' = throw e\<^sub>t" have "P \ \e,s\ \ \throw e\<^sub>t,s'\" using e'' IH by simp then show ?thesis using eval_evals.SeqThrow e' by simp qed \ \ separate evaluation of @{text RI} subexp, val case \ lemma rinit_Val_ext: assumes ri: "P \ \RI (C,e'') ; Cs \ e\<^sub>0,s''\ \ \Val v',s\<^sub>1\" and IH: "\e' s'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\" shows "P \ \RI (C,e) ; Cs \ e\<^sub>0,s\ \ \Val v',s\<^sub>1\" proof(rule eval_cases(20)[OF ri]) \ \ RI \ fix v'' h' l' sh' sfs i assume e''step: "P \ \e'',s''\ \ \Val v'',(h', l', sh')\" and shC: "sh' C = \(sfs, i)\" and init: "P \ \INIT (if Cs = [] then C else last Cs) (Cs,True) \ e\<^sub>0,(h', l', sh'(C \ (sfs, Done)))\ \ \Val v',s\<^sub>1\" have "P \ \e,s\ \ \Val v'',(h', l', sh')\" using IH[OF e''step] by simp then show ?thesis using RInit init shC by auto next fix a h' l' sh' sfs i D Cs' assume e''step: "P \ \e'',s''\ \ \throw a,(h', l', sh')\" and riD: "P \ \RI (D,throw a) ; Cs' \ e\<^sub>0,(h', l', sh'(C \ (sfs, Error)))\ \ \Val v',s\<^sub>1\" have "P \ \e,s\ \ \throw a,(h', l', sh')\" using IH[OF e''step] by simp then show ?thesis using riD rinit_throwE by blast qed(simp) \ \ separate evaluation of @{text RI} subexp, throw case \ lemma rinit_throw_ext: assumes ri: "P \ \RI (C,e'') ; Cs \ e\<^sub>0,s''\ \ \throw e\<^sub>t,s'\" and IH: "\e' s'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\" shows "P \ \RI (C,e) ; Cs \ e\<^sub>0,s\ \ \throw e\<^sub>t,s'\" proof(rule eval_cases(20)[OF ri]) \ \ RI \ fix v h' l' sh' sfs i assume e''step: "P \ \e'',s''\ \ \Val v,(h', l', sh')\" and shC: "sh' C = \(sfs, i)\" and init: "P \ \INIT (if Cs = [] then C else last Cs) (Cs,True) \ e\<^sub>0,(h', l', sh'(C \ (sfs, Done)))\ \ \throw e\<^sub>t,s'\" have "P \ \e,s\ \ \Val v,(h', l', sh')\" using IH[OF e''step] by simp then show ?thesis using RInit init shC by auto next fix a h' l' sh' sfs i D Cs' assume e''step: "P \ \e'',s''\ \ \throw a,(h', l', sh')\" and shC: "sh' C = \(sfs, i)\" and riD: "P \ \RI (D,throw a) ; Cs' \ e\<^sub>0,(h', l', sh'(C \ (sfs, Error)))\ \ \throw e\<^sub>t,s'\" and cons: "Cs = D # Cs'" have estep': "P \ \e,s\ \ \throw a,(h', l', sh')\" using IH[OF e''step] by simp then show ?thesis using RInitInitFail cons riD shC by simp next fix a h' l' sh' sfs i assume "throw e\<^sub>t = throw a" and "s' = (h', l', sh'(C \ (sfs, Error)))" and "P \ \e'',s''\ \ \throw a,(h', l', sh')\" and "sh' C = \(sfs, i)\" and "Cs = []" then show ?thesis using RInitFailFinal IH by auto qed \ \ separate evaluation of @{text RI} subexp \ lemma rinit_ext: assumes IH: "\e' s'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\" shows "\e' s'. P \ \RI (C,e'') ; Cs \ e\<^sub>0,s''\ \ \e',s'\ \ P \ \RI (C,e) ; Cs \ e\<^sub>0,s\ \ \e',s'\" proof - fix e' s' assume ri'': "P \ \RI (C,e'') ; Cs \ e\<^sub>0,s''\ \ \e',s'\" then have "final e'" using eval_final by simp then show "P \ \RI (C,e) ; Cs \ e\<^sub>0,s\ \ \e',s'\" proof(rule finalE) fix v assume "e' = Val v" then show ?thesis using rinit_Val_ext[OF _ IH] ri'' by simp next fix a assume "e' = throw a" then show ?thesis using rinit_throw_ext[OF _ IH] ri'' by simp qed qed \ \ @{text INIT} and @{text RI} return either @{text Val} with @{text Done} or @{text Processing} flag or @{text Throw} with @{text Error} flag \ lemma shows eval_init_return: "P \ \e,s\ \ \e',s'\ \ iconf (shp s) e \ (\Cs b. e = INIT C' (Cs,b) \ unit) \ (\C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs@[C'] \ unit) \ (\e\<^sub>0. e = RI(C',e\<^sub>0);Nil \ unit) \ (val_of e' = Some v \ (\sfs i. shp s' C' = \(sfs,i)\ \ (i = Done \ i = Processing))) \ (throw_of e' = Some a \ (\sfs i. shp s' C' = \(sfs,Error)\))" and "P \ \es,s\ [\] \es',s'\ \ True" proof(induct rule: eval_evals.inducts) case (InitFinal e s e' s' C b) then show ?case by(fastforce simp: initPD_def dest: eval_final_same) next case (InitDone sh C sfs C' Cs e h l e' s') then have "final e'" using eval_final by simp then show ?case proof(rule finalE) fix v assume e': "e' = Val v" then show ?thesis using InitDone initPD_def proof(cases Cs) qed(auto) next fix a assume e': "e' = throw a" then show ?thesis using InitDone initPD_def proof(cases Cs) qed(auto) qed next case (InitProcessing sh C sfs C' Cs e h l e' s') then have "final e'" using eval_final by simp then show ?case proof(rule finalE) fix v assume e': "e' = Val v" then show ?thesis using InitProcessing initPD_def proof(cases Cs) qed(auto) next fix a assume e': "e' = throw a" then show ?thesis using InitProcessing initPD_def proof(cases Cs) qed(auto) qed next case (InitError sh C sfs Cs e h l e' s' C') show ?case proof(cases Cs) case Nil then show ?thesis using InitError by simp next case (Cons C2 list) then have "final e'" using InitError eval_final by simp then show ?thesis proof(rule finalE) fix v assume e': "e' = Val v" then show ?thesis using InitError proof - obtain ccss :: "char list list" and bb :: bool where "INIT C' (C # Cs,False) \ e = INIT C' (ccss,bb) \ unit" using InitError.prems(2) by blast then show ?thesis using InitError.hyps(2) e' by(auto dest!: rinit_throwE) qed next fix a assume e': "e' = throw a" then show ?thesis using Cons InitError cons_to_append[of list] by clarsimp qed qed next case (InitRInit C Cs h l sh e' s' C') show ?case proof(cases Cs) case Nil then show ?thesis using InitRInit by simp next case (Cons C' list) then show ?thesis using InitRInit Cons cons_to_append[of list] by clarsimp qed next case (RInit e s v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 s\<^sub>1) then have final: "final e\<^sub>1" using eval_final by simp then show ?case proof(cases Cs) case Nil show ?thesis using final proof(rule finalE) fix v assume e': "e\<^sub>1 = Val v" show ?thesis using RInit Nil by(auto simp: fun_upd_same initPD_def) next fix a assume e': "e\<^sub>1 = throw a" show ?thesis using RInit Nil by(auto simp: fun_upd_same initPD_def) qed next case (Cons a list) show ?thesis proof(rule finalE[OF final]) fix v assume e': "e\<^sub>1 = Val v" then show ?thesis using RInit Cons by clarsimp (metis last.simps last_appendR list.distinct(1)) next fix a assume e': "e\<^sub>1 = throw a" then show ?thesis using RInit Cons by clarsimp (metis last.simps last_appendR list.distinct(1)) qed qed next case (RInitInitFail e s a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 s\<^sub>1) then have final: "final e\<^sub>1" using eval_final by simp then show ?case proof(rule finalE) fix v assume e': "e\<^sub>1 = Val v" then show ?thesis using RInitInitFail by clarsimp (meson exp.distinct(101) rinit_throwE) next fix a' assume e': "e\<^sub>1 = Throw a'" then have "iconf (sh'(C \ (sfs, Error))) a" using RInitInitFail.hyps(1) eval_final by fastforce then show ?thesis using RInitInitFail e' by clarsimp (meson Cons_eq_append_conv list.inject) qed qed(auto simp: fun_upd_same) lemma init_Val_PD: "P \ \INIT C' (Cs,b) \ unit,s\ \ \Val v,s'\ \ iconf (shp s) (INIT C' (Cs,b) \ unit) \ \sfs i. shp s' C' = \(sfs,i)\ \ (i = Done \ i = Processing)" by(drule_tac v = v in eval_init_return, simp+) lemma init_throw_PD: "P \ \INIT C' (Cs,b) \ unit,s\ \ \throw a,s'\ \ iconf (shp s) (INIT C' (Cs,b) \ unit) \ \sfs i. shp s' C' = \(sfs,Error)\" by(drule_tac a = a in eval_init_return, simp+) lemma rinit_Val_PD: "P \ \RI(C,e\<^sub>0);Cs \ unit,s\ \ \Val v,s'\ \ iconf (shp s) (RI(C,e\<^sub>0);Cs \ unit) \ last(C#Cs) = C' \ \sfs i. shp s' C' = \(sfs,i)\ \ (i = Done \ i = Processing)" by(auto dest!: eval_init_return [where C'=C'] append_butlast_last_id[THEN sym]) lemma rinit_throw_PD: "P \ \RI(C,e\<^sub>0);Cs \ unit,s\ \ \throw a,s'\ \ iconf (shp s) (RI(C,e\<^sub>0);Cs \ unit) \ last(C#Cs) = C' \ \sfs i. shp s' C' = \(sfs,Error)\" by(auto dest!: eval_init_return [where C'=C'] append_butlast_last_id[THEN sym]) \ \ combining the above to get evaluation of @{text INIT} in a sequence \ (* Hiermit kann man die ganze pair-Splitterei in den automatischen Taktiken abschalten. Wieder anschalten siehe nach dem Beweis. *) (*<*) declare split_paired_All [simp del] split_paired_Ex [simp del] (*>*) lemma eval_init_seq': "P \ \e,s\ \ \e',s'\ \ (\C Cs b e\<^sub>i. e = INIT C (Cs,b) \ e\<^sub>i) \ (\C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs \ e\<^sub>i) \ (\C Cs b e\<^sub>i. e = INIT C (Cs,b) \ e\<^sub>i \ P \ \(INIT C (Cs,b) \ unit);; e\<^sub>i,s\ \ \e',s'\) \ (\C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs \ e\<^sub>i \ P \ \(RI(C,e\<^sub>0);Cs \ unit);; e\<^sub>i,s\ \ \e',s'\)" and "P \ \es,s\ [\] \es',s'\ \ True" proof(induct rule: eval_evals.inducts) case InitFinal then show ?case by(auto simp: Seq[OF eval_evals.InitFinal[OF Val[where v=Unit]]]) next case (InitNone sh) then show ?case using seq_ext[OF eval_evals.InitNone[where sh=sh and e=unit, OF InitNone.hyps(1)]] by fastforce next case (InitDone sh) then show ?case using seq_ext[OF eval_evals.InitDone[where sh=sh and e=unit, OF InitDone.hyps(1)]] by fastforce next case (InitProcessing sh) then show ?case using seq_ext[OF eval_evals.InitProcessing[where sh=sh and e=unit, OF InitProcessing.hyps(1)]] by fastforce next case (InitError sh) then show ?case using seq_ext[OF eval_evals.InitError[where sh=sh and e=unit, OF InitError.hyps(1)]] by fastforce next case (InitObject sh) then show ?case using seq_ext[OF eval_evals.InitObject[where sh=sh and e=unit, OF InitObject.hyps(1)]] by fastforce next case (InitNonObject sh) then show ?case using seq_ext[OF eval_evals.InitNonObject[where sh=sh and e=unit, OF InitNonObject.hyps(1)]] by fastforce next case (InitRInit C Cs e h l sh e' s' C') then show ?case using seq_ext[OF eval_evals.InitRInit[where e=unit]] by fastforce next case RInit then show ?case using seq_ext[OF eval_evals.RInit[where e=unit, OF RInit.hyps(1)]] by fastforce next case RInitInitFail then show ?case using seq_ext[OF eval_evals.RInitInitFail[where e=unit, OF RInitInitFail.hyps(1)]] by fastforce next case RInitFailFinal then show ?case using eval_evals.RInitFailFinal eval_evals.SeqThrow by auto qed(auto) lemma eval_init_seq: "P \ \INIT C (Cs,b) \ e,(h, l, sh)\ \ \e',s'\ \ P \ \(INIT C (Cs,b) \ unit);; e,(h, l, sh)\ \ \e',s'\" by(auto dest: eval_init_seq') text \ The key lemma: \ lemma assumes wf: "wwf_J_prog P" shows extend_1_eval: "P \ \e,s,b\ \ \e'',s'',b''\ \ P,shp s \\<^sub>b (e,b) \ \ (\s' e'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\)" and extend_1_evals: "P \ \es,s,b\ [\] \es'',s'',b''\ \ P,shp s \\<^sub>b (es,b) \ \ (\s' es'. P \ \es'',s''\ [\] \es',s'\ \ P \ \es,s\ [\] \es',s'\)" proof (induct rule: red_reds.inducts) case (RedNew h a C FDTs h' l sh) then have e':"e' = addr a" and s':"s' = (h(a \ blank P C), l, sh)" using eval_cases(3) by fastforce+ obtain sfs i where shC: "sh C = \(sfs, i)\" and "i = Done \ i = Processing" using RedNew by (clarsimp simp: bconf_def initPD_def) then show ?case proof(cases i) case Done then show ?thesis using RedNew shC e' s' New by simp next case Processing then have shC': "\sfs. sh C = Some(sfs,Done)" and shP: "sh C = Some(sfs,Processing)" using shC by simp+ then have init: "P \ \INIT C ([C],False) \ unit,(h,l,sh)\ \ \unit,(h,l,sh)\" by(simp add: InitFinal InitProcessing Val) have "P \ \new C,(h, l, sh)\ \ \addr a,(h(a \ blank P C),l,sh)\" using RedNew shC' by(auto intro: NewInit[OF _ init]) then show ?thesis using e' s' by simp qed(auto) next case (RedNewFail h C l sh) then have e':"e' = THROW OutOfMemory" and s':"s' = (h, l, sh)" using eval_final_same final_def by fastforce+ obtain sfs i where shC: "sh C = \(sfs, i)\" and "i = Done \ i = Processing" using RedNewFail by (clarsimp simp: bconf_def initPD_def) then show ?case proof(cases i) case Done then show ?thesis using RedNewFail shC e' s' NewFail by simp next case Processing then have shC': "\sfs. sh C = Some(sfs,Done)" and shP: "sh C = Some(sfs,Processing)" using shC by simp+ then have init: "P \ \INIT C ([C],False) \ unit,(h,l,sh)\ \ \unit,(h,l,sh)\" by(simp add: InitFinal InitProcessing Val) have "P \ \new C,(h, l, sh)\ \ \THROW OutOfMemory,(h,l,sh)\" using RedNewFail shC' by(auto intro: NewInitOOM[OF _ init]) then show ?thesis using e' s' by simp qed(auto) next case (NewInitRed sh C h l) then have seq: "P \ \(INIT C ([C],False) \ unit);; new C,(h, l, sh)\ \ \e',s'\" using eval_init_seq by simp then show ?case proof(rule eval_cases(14)) \ \ Seq \ fix v s\<^sub>1 assume init: "P \ \INIT C ([C],False) \ unit,(h, l, sh)\ \ \Val v,s\<^sub>1\" and new: "P \ \new C,s\<^sub>1\ \ \e',s'\" obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1) then obtain sfs i where shC: "sh\<^sub>1 C = \(sfs, i)\" and iDP: "i = Done \ i = Processing" using init_Val_PD[OF init] by auto show ?thesis proof(rule eval_cases(1)[OF new]) \ \ New \ fix sha ha a FDTs la assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = addr a" and s': "s' = (ha(a \ blank P C), la, sha)" and addr: "new_Addr ha = \a\" and fields: "P \ C has_fields FDTs" then show ?thesis using NewInit[OF _ _ addr fields] NewInitRed.hyps init by simp next fix sha ha la assume "s\<^sub>1 = (ha, la, sha)" and "e' = THROW OutOfMemory" and "s' = (ha, la, sha)" and "new_Addr ha = None" then show ?thesis using NewInitOOM NewInitRed.hyps init by simp next fix sha ha la v' h' l' sh' a FDTs assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = addr a" and s': "s' = (h'(a \ blank P C), l', sh')" and shaC: "\sfs. sha C \ \(sfs, Done)\" and init': "P \ \INIT C ([C],False) \ unit,(ha, la, sha)\ \ \Val v',(h', l', sh')\" and addr: "new_Addr h' = \a\" and fields: "P \ C has_fields FDTs" then have i: "i = Processing" using iDP shC s\<^sub>1 by simp then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast then show ?thesis using NewInit NewInitRed.hyps s\<^sub>1a addr fields init shaC e' s' by auto next fix sha ha la v' h' l' sh' assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = THROW OutOfMemory" and s': "s' = (h', l', sh')" and "\sfs. sha C \ \(sfs, Done)\" and init': "P \ \INIT C ([C],False) \ unit,(ha, la, sha)\ \ \Val v',(h', l', sh')\" and addr: "new_Addr h' = None" then have i: "i = Processing" using iDP shC s\<^sub>1 by simp then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast then show ?thesis using NewInitOOM NewInitRed.hyps e' addr s' s\<^sub>1a init by auto next fix sha ha la a assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and "\sfs. sha C \ \(sfs, Done)\" and init': "P \ \INIT C ([C],False) \ unit,(ha, la, sha)\ \ \throw a,s'\" then have i: "i = Processing" using iDP shC s\<^sub>1 by simp then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast qed next fix e assume e': "e' = throw e" and init: "P \ \INIT C ([C],False) \ unit,(h, l, sh)\ \ \throw e,s'\" obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s') then obtain sfs i where shC: "sh' C = \(sfs, i)\" and iDP: "i = Error" using init_throw_PD[OF init] by auto then show ?thesis by (simp add: NewInitRed.hyps NewInitThrow e' init) qed next case CastRed then show ?case by(fastforce elim!: eval_cases intro: eval_evals.intros intro!: CastFail) next case RedCastNull then show ?case by simp (iprover elim: eval_cases intro: eval_evals.intros) next case RedCast then show ?case by (auto elim: eval_cases intro: eval_evals.intros) next case RedCastFail then show ?case by (auto elim!: eval_cases intro: eval_evals.intros) next case BinOpRed1 then show ?case by(fastforce elim!: eval_cases intro: eval_evals.intros simp: val_no_step) next case BinOpRed2 thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId) next case RedBinOp thus ?case by simp (iprover elim: eval_cases intro: eval_evals.intros) next case RedVar thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case LAssRed thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case RedLAss thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAccRed thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case RedFAcc then show ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedFAccNull then show ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case RedFAccNone thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case RedFAccStatic thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case (RedSFAcc C F t D sh sfs i v h l) then have e':"e' = Val v" and s':"s' = (h, l, sh)" using eval_cases(3) by fastforce+ have "i = Done \ i = Processing" using RedSFAcc by (clarsimp simp: bconf_def initPD_def) then show ?case proof(cases i) case Done then show ?thesis using RedSFAcc e' s' SFAcc by simp next case Processing then have shC': "\sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)" using RedSFAcc by simp+ then have init: "P \ \INIT D ([D],False) \ unit,(h,l,sh)\ \ \unit,(h,l,sh)\" by(simp add: InitFinal InitProcessing Val) have "P \ \C\\<^sub>sF{D},(h, l, sh)\ \ \Val v,(h,l,sh)\" by(rule SFAccInit[OF RedSFAcc.hyps(1) shC' init shP RedSFAcc.hyps(3)]) then show ?thesis using e' s' by simp qed(auto) next case (SFAccInitRed C F t D sh h l) then have seq: "P \ \(INIT D ([D],False) \ unit);; C\\<^sub>sF{D},(h, l, sh)\ \ \e',s'\" using eval_init_seq by simp then show ?case proof(rule eval_cases(14)) \ \ Seq \ fix v s\<^sub>1 assume init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \Val v,s\<^sub>1\" and acc: "P \ \C\\<^sub>sF{D},s\<^sub>1\ \ \e',s'\" obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1) then obtain sfs i where shD: "sh\<^sub>1 D = \(sfs, i)\" and iDP: "i = Done \ i = Processing" using init_Val_PD[OF init] by auto show ?thesis proof(rule eval_cases(8)[OF acc]) \ \ SFAcc \ fix t sha sfs v ha la assume "s\<^sub>1 = (ha, la, sha)" and "e' = Val v" and "s' = (ha, la, sha)" and "P \ C has F,Static:t in D" and "sha D = \(sfs, Done)\" and "sfs F = \v\" then show ?thesis using SFAccInit SFAccInitRed.hyps(2) init by auto next fix t sha ha la v' h' l' sh' sfs i' v assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = Val v" and s': "s' = (h', l', sh')" and field: "P \ C has F,Static:t in D" and "\sfs. sha D \ \(sfs, Done)\" and init': "P \ \INIT D ([D],False) \ unit,(ha, la, sha)\ \ \Val v',(h', l', sh')\" and shD': "sh' D = \(sfs, i')\" and sfsF: "sfs F = \v\" then have i: "i = Processing" using iDP shD s\<^sub>1 by simp then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast then show ?thesis using SFAccInit SFAccInitRed.hyps(2) e' s' field init s\<^sub>1a sfsF shD' by auto next fix t sha ha la a assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and "e' = throw a" and "P \ C has F,Static:t in D" and "\sfs. sha D \ \(sfs, Done)\" and init': "P \ \INIT D ([D],False) \ unit,(ha, la, sha)\ \ \throw a,s'\" then have i: "i = Processing" using iDP shD s\<^sub>1 by simp then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast next assume "\b t. \ P \ C has F,b:t in D" then show ?thesis using SFAccInitRed.hyps(1) by blast next fix t assume field: "P \ C has F,NonStatic:t in D" then show ?thesis using has_field_fun[OF SFAccInitRed.hyps(1) field] by simp qed next fix e assume e': "e' = throw e" and init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \throw e,s'\" obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s') then obtain sfs i where shC: "sh' D = \(sfs, i)\" and iDP: "i = Error" using init_throw_PD[OF init] by auto then show ?thesis using SFAccInitRed.hyps(1) SFAccInitRed.hyps(2) SFAccInitThrow e' init by auto qed next case RedSFAccNone thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case RedSFAccNonStatic thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case (FAssRed1 e s b e1 s1 b1 F D e\<^sub>2) obtain h' l' sh' where "s'=(h',l',sh')" by(cases s') with FAssRed1 show ?case by(fastforce elim!: eval_cases(9)[where e\<^sub>1=e1] intro: eval_evals.intros simp: val_no_step intro!: FAss FAssNull FAssNone FAssStatic FAssThrow2) next case FAssRed2 obtain h' l' sh' where "s'=(h',l',sh')" by(cases s') with FAssRed2 show ?case by(auto elim!: eval_cases intro: eval_evals.intros intro!: FAss FAssNull FAssNone FAssStatic FAssThrow2 Val) next case RedFAss thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case RedFAssNull thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case RedFAssNone then show ?case by(auto elim!: eval_cases intro: eval_evals.intros eval_finalId) next case RedFAssStatic then show ?case by(auto elim!: eval_cases intro: eval_evals.intros eval_finalId) next case (SFAssRed e s b e'' s'' b'' C F D) obtain h l sh where [simp]: "s = (h,l,sh)" by(cases s) obtain h' l' sh' where [simp]: "s'=(h',l',sh')" by(cases s') have "val_of e = None" using val_no_step SFAssRed.hyps(1) by(meson option.exhaust) then have bconf: "P,sh \\<^sub>b (e,b) \" using SFAssRed by simp show ?case using SFAssRed.prems(2) SFAssRed bconf proof cases case 2 with SFAssRed bconf show ?thesis by(auto intro!: SFAssInit) next case 3 with SFAssRed bconf show ?thesis by(auto intro!: SFAssInitThrow) qed(auto intro: eval_evals.intros intro!: SFAss SFAssInit SFAssNone SFAssNonStatic) next case (RedSFAss C F t D sh sfs i sfs' v sh' h l) let ?sfs' = "sfs(F \ v)" have e':"e' = unit" and s':"s' = (h, l, sh(D \ (?sfs', i)))" using RedSFAss eval_cases(3) by fastforce+ have "i = Done \ i = Processing" using RedSFAss by(clarsimp simp: bconf_def initPD_def) then show ?case proof(cases i) case Done then show ?thesis using RedSFAss e' s' SFAss Val by auto next case Processing then have shC': "\sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)" using RedSFAss by simp+ then have init: "P \ \INIT D ([D],False) \ unit,(h,l,sh)\ \ \unit,(h,l,sh)\" by(simp add: InitFinal InitProcessing Val) have "P \ \C\\<^sub>sF{D} := Val v,(h, l, sh)\ \ \unit,(h,l,sh(D \ (?sfs', i)))\" using Processing by(auto intro: SFAssInit[OF Val RedSFAss.hyps(1) shC' init shP]) then show ?thesis using e' s' by simp qed(auto) next case (SFAssInitRed C F t D sh v h l) then have seq: "P \ \(INIT D ([D],False) \ unit);; C\\<^sub>sF{D} := Val v,(h, l, sh)\ \ \e',s'\" using eval_init_seq by simp then show ?case proof(rule eval_cases(14)) \ \ Seq \ fix v' s\<^sub>1 assume init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \Val v',s\<^sub>1\" and acc: "P \ \C\\<^sub>sF{D} := Val v,s\<^sub>1\ \ \e',s'\" obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1) then obtain sfs i where shD: "sh\<^sub>1 D = \(sfs, i)\" and iDP: "i = Done \ i = Processing" using init_Val_PD[OF init] by auto show ?thesis proof(rule eval_cases(10)[OF acc]) \ \ SFAss \ fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t sfs assume e': "e' = unit" and s': "s' = (h\<^sub>1, l\<^sub>1, sh\<^sub>1(D \ (sfs(F \ va), Done)))" and val: "P \ \Val v,s\<^sub>1\ \ \Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\" and field: "P \ C has F,Static:t in D" and shD': "sh\<^sub>1 D = \(sfs, Done)\" have "v = va" and "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto then show ?thesis using SFAssInit field SFAssInitRed.hyps(2) shD' e' s' init val by (metis eval_final eval_finalId) next fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t v' h' l' sh' sfs i' assume e': "e' = unit" and s': "s' = (h', l', sh'(D \ (sfs(F \ va), i')))" and val: "P \ \Val v,s\<^sub>1\ \ \Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\" and field: "P \ C has F,Static:t in D" and nDone: "\sfs. sh\<^sub>1 D \ \(sfs, Done)\" and init': "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\ \ \Val v',(h', l', sh')\" and shD': "sh' D = \(sfs, i')\" have v: "v = va" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto then have i: "i = Processing" using iDP shD s\<^sub>1 nDone by simp then have "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = (h', l', sh')" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast then show ?thesis using SFAssInit SFAssInitRed.hyps(2) e' s' field init v s\<^sub>1a shD' val by (metis eval_final eval_finalId) next fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t a assume "e' = throw a" and val: "P \ \Val v,s\<^sub>1\ \ \Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\" and "P \ C has F,Static:t in D" and nDone: "\sfs. sh\<^sub>1 D \ \(sfs, Done)\" and init': "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\ \ \throw a,s'\" have v: "v = va" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto then have i: "i = Processing" using iDP shD s\<^sub>1 nDone by simp then have "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = s'" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD i by blast next fix e'' assume val:"P \ \Val v,s\<^sub>1\ \ \throw e'',s'\" then show ?thesis using eval_final_same[OF val] by simp next assume "\b t. \ P \ C has F,b:t in D" then show ?thesis using SFAssInitRed.hyps(1) by blast next fix t assume field: "P \ C has F,NonStatic:t in D" then show ?thesis using has_field_fun[OF SFAssInitRed.hyps(1) field] by simp qed next fix e assume e': "e' = throw e" and init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \throw e,s'\" obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s') then obtain sfs i where shC: "sh' D = \(sfs, i)\" and iDP: "i = Error" using init_throw_PD[OF init] by auto then show ?thesis using SFAssInitRed.hyps(1) SFAssInitRed.hyps(2) SFAssInitThrow Val by (metis e' init) qed next case (RedSFAssNone C F D v s b) then show ?case by(cases s) (auto elim!: eval_cases intro: eval_evals.intros eval_finalId) next case (RedSFAssNonStatic C F t D v s b) then show ?case by(cases s) (auto elim!: eval_cases intro: eval_evals.intros eval_finalId) next case CallObj note val_no_step[simp] from CallObj.prems(2) CallObj show ?case proof cases case 2 with CallObj show ?thesis by(fastforce intro!: CallParamsThrow) next case 3 with CallObj show ?thesis by(fastforce intro!: CallNull) next case 4 with CallObj show ?thesis by(fastforce intro!: CallNone) next case 5 with CallObj show ?thesis by(fastforce intro!: CallStatic) qed(fastforce intro!: CallObjThrow Call)+ next case (CallParams es s b es'' s'' b'' v M s') then obtain h' l' sh' where "s' = (h',l',sh')" by(cases s') with CallParams show ?case by(auto elim!: eval_cases intro!: CallNone eval_finalId CallStatic Val) (auto intro!: CallParamsThrow CallNull Call Val) next case (RedCall h a C fs M Ts T pns body D vs l sh b) have "P \ \addr a,(h,l,sh)\ \ \addr a,(h,l,sh)\" by (rule eval_evals.intros) moreover have finals: "finals(map Val vs)" by simp with finals have "P \ \map Val vs,(h,l,sh)\ [\] \map Val vs,(h,l,sh)\" by (iprover intro: eval_finalsId) moreover have "h a = Some (C, fs)" using RedCall by simp moreover have "method": "P \ C sees M,NonStatic: Ts\T = (pns, body) in D" by fact moreover have same_len\<^sub>1: "length Ts = length pns" and this_distinct: "this \ set pns" and fv: "fv (body) \ {this} \ set pns" using "method" wf by (fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ have same_len: "length vs = length pns" by fact moreover obtain l\<^sub>2' where l\<^sub>2': "l\<^sub>2' = [this\Addr a,pns[\]vs]" by simp moreover obtain h\<^sub>3 l\<^sub>3 sh\<^sub>3 where s': "s' = (h\<^sub>3,l\<^sub>3,sh\<^sub>3)" by (cases s') have eval_blocks: "P \ \(blocks (this # pns, Class D # Ts, Addr a # vs, body)),(h,l,sh)\ \ \e',s'\" by fact hence id: "l\<^sub>3 = l" using fv s' same_len\<^sub>1 same_len by(fastforce elim: eval_closed_lcl_unchanged) from eval_blocks obtain l\<^sub>3' where "P \ \body,(h,l\<^sub>2',sh)\ \ \e',(h\<^sub>3,l\<^sub>3',sh\<^sub>3)\" proof - from same_len\<^sub>1 have "length(this#pns) = length(Class D#Ts)" by simp moreover from same_len\<^sub>1 same_len have same_len\<^sub>2: "length (this#pns) = length (Addr a#vs)" by simp moreover from eval_blocks have "P \ \blocks (this#pns,Class D#Ts,Addr a#vs,body),(h,l,sh)\ \\e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\" using s' same_len\<^sub>1 same_len\<^sub>2 by simp ultimately obtain l'' where "P \ \body,(h,l(this # pns[\]Addr a # vs),sh)\\\e',(h\<^sub>3, l'',sh\<^sub>3)\" by (blast dest:blocksEval) from eval_restrict_lcl[OF wf this fv] this_distinct same_len\<^sub>1 same_len have "P \ \body,(h,[this # pns[\]Addr a # vs],sh)\ \ \e',(h\<^sub>3, l''|`(set(this#pns)),sh\<^sub>3)\" using wf method by(simp add:subset_insert_iff insert_Diff_if) thus ?thesis by(fastforce intro!:that simp add: l\<^sub>2') qed ultimately have "P \ \(addr a)\M(map Val vs),(h,l,sh)\ \ \e',(h\<^sub>3,l,sh\<^sub>3)\" by (rule Call) with s' id show ?case by simp next case RedCallNull thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros eval_finalsId) next case (RedCallNone h a C fs M vs l sh b) then have tes: "THROW NoSuchMethodError = e' \ (h,l,sh) = s'" using eval_final_same by simp have "P \ \addr a,(h,l,sh)\ \ \addr a,(h,l,sh)\" and "P \ \map Val vs,(h,l,sh)\ [\] \map Val vs,(h,l,sh)\" using eval_finalId eval_finalsId by auto then show ?case using RedCallNone CallNone tes by auto next case (RedCallStatic h a C fs M Ts T m D vs l sh b) then have tes: "THROW IncompatibleClassChangeError = e' \ (h,l,sh) = s'" using eval_final_same by simp have "P \ \addr a,(h,l,sh)\ \ \addr a,(h,l,sh)\" and "P \ \map Val vs,(h,l,sh)\ [\] \map Val vs,(h,l,sh)\" using eval_finalId eval_finalsId by auto then show ?case using RedCallStatic CallStatic tes by fastforce next case (SCallParams es s b es'' s'' b' C M s') obtain h' l' sh' where s'[simp]: "s' = (h',l',sh')" by(cases s') obtain h l sh where s[simp]: "s = (h,l,sh)" by(cases s) have es: "map_vals_of es = None" using vals_no_step SCallParams.hyps(1) by (meson not_Some_eq) have bconf: "P,sh \\<^sub>b (es,b) \" using s SCallParams.prems(1) by (simp add: bconf_SCall[OF es]) from SCallParams.prems(2) SCallParams bconf show ?case proof cases case 2 with SCallParams bconf show ?thesis by(auto intro!: SCallNone) next case 3 with SCallParams bconf show ?thesis by(auto intro!: SCallNonStatic) next case 4 with SCallParams bconf show ?thesis by(auto intro!: SCallInitThrow) next case 5 with SCallParams bconf show ?thesis by(auto intro!: SCallInit) qed(auto intro!: SCallParamsThrow SCall) next case (RedSCall C M Ts T pns body D vs s) then obtain h l sh where s:"s = (h,l,sh)" by(cases s) then obtain sfs i where shC: "sh D = \(sfs, i)\" and "i = Done \ i = Processing" using RedSCall by(auto simp: bconf_def initPD_def dest: sees_method_fun) have finals: "finals(map Val vs)" by simp with finals have mVs: "P \ \map Val vs,(h,l,sh)\ [\] \map Val vs,(h,l,sh)\" by (iprover intro: eval_finalsId) obtain sfs i where shC: "sh D = \(sfs, i)\" using RedSCall s by(auto simp: bconf_def initPD_def dest: sees_method_fun) then have iDP: "i = Done \ i = Processing" using RedSCall s by (auto simp: bconf_def initPD_def dest: sees_method_fun[OF RedSCall.hyps(1)]) have "method": "P \ C sees M,Static: Ts\T = (pns, body) in D" by fact have same_len\<^sub>1: "length Ts = length pns" and fv: "fv (body) \ set pns" using "method" wf by (fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ have same_len: "length vs = length pns" by fact obtain l\<^sub>2' where l\<^sub>2': "l\<^sub>2' = [pns[\]vs]" by simp obtain h\<^sub>3 l\<^sub>3 sh\<^sub>3 where s': "s' = (h\<^sub>3,l\<^sub>3,sh\<^sub>3)" by (cases s') have eval_blocks: "P \ \(blocks (pns, Ts, vs, body)),(h,l,sh)\ \ \e',s'\" using RedSCall.prems(2) s by simp hence id: "l\<^sub>3 = l" using fv s' same_len\<^sub>1 same_len by(fastforce elim: eval_closed_lcl_unchanged) from eval_blocks obtain l\<^sub>3' where body: "P \ \body,(h,l\<^sub>2',sh)\ \ \e',(h\<^sub>3,l\<^sub>3',sh\<^sub>3)\" proof - from eval_blocks have "P \ \blocks (pns,Ts,vs,body),(h,l,sh)\ \\e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\" using s' same_len same_len\<^sub>1 by simp then obtain l'' where "P \ \body,(h,l(pns[\]vs),sh)\ \ \e',(h\<^sub>3, l'',sh\<^sub>3)\" by (blast dest:blocksEval[OF same_len\<^sub>1[THEN sym] same_len[THEN sym]]) from eval_restrict_lcl[OF wf this fv] same_len\<^sub>1 same_len have "P \ \body,(h,[pns[\]vs],sh)\ \ \e',(h\<^sub>3, l''|`(set(pns)),sh\<^sub>3)\" using wf method by(simp add:subset_insert_iff insert_Diff_if) thus ?thesis by(fastforce intro!:that simp add: l\<^sub>2') qed show ?case using iDP proof(cases i) case Done then have shC': "sh D = \(sfs, Done)\ \ M = clinit \ sh D = \(sfs, Processing)\" using shC by simp have "P \ \C\\<^sub>sM(map Val vs),(h,l,sh)\ \ \e',(h\<^sub>3,l,sh\<^sub>3)\" by (rule SCall[OF mVs method shC' same_len l\<^sub>2' body]) with s s' id show ?thesis by simp next case Processing then have shC': "\sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)" using shC by simp+ then have init: "P \ \INIT D ([D],False) \ unit,(h,l,sh)\ \ \unit,(h,l,sh)\" by(simp add: InitFinal InitProcessing Val) have "P \ \C\\<^sub>sM(map Val vs),(h,l,sh)\ \ \e',(h\<^sub>3,l,sh\<^sub>3)\" proof(cases "M = clinit") case False show ?thesis by(rule SCallInit[OF mVs method shC' False init same_len l\<^sub>2' body]) next case True then have shC': "sh D = \(sfs, Done)\ \ M = clinit \ sh D = \(sfs, Processing)\" using shC Processing by simp have "P \ \C\\<^sub>sM(map Val vs),(h,l,sh)\ \ \e',(h\<^sub>3,l,sh\<^sub>3)\" by (rule SCall[OF mVs method shC' same_len l\<^sub>2' body]) with s s' id show ?thesis by simp qed with s s' id show ?thesis by simp qed(auto) next case (SCallInitRed C M Ts T pns body D sh vs h l) then have seq: "P \ \(INIT D ([D],False) \ unit);; C\\<^sub>sM(map Val vs),(h, l, sh)\ \ \e',s'\" using eval_init_seq by simp then show ?case proof(rule eval_cases(14)) \ \ Seq \ fix v' s\<^sub>1 assume init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \Val v',s\<^sub>1\" and call: "P \ \C\\<^sub>sM(map Val vs),s\<^sub>1\ \ \e',s'\" obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1) then obtain sfs i where shD: "sh\<^sub>1 D = \(sfs, i)\" and iDP: "i = Done \ i = Processing" using init_Val_PD[OF init] by auto show ?thesis proof(rule eval_cases(12)[OF call]) \ \ SCall \ fix vsa ex es' assume "P \ \map Val vs,s\<^sub>1\ [\] \map Val vsa @ throw ex # es',s'\" then show ?thesis using evals_finals_same by (meson finals_def map_Val_nthrow_neq) next assume "\b Ts T a ba x. \ P \ C sees M, b : Ts\T = (a, ba) in x" then show ?thesis using SCallInitRed.hyps(1) by auto next fix Ts T m D assume "P \ C sees M, NonStatic : Ts\T = m in D" then show ?thesis using sees_method_fun[OF SCallInitRed.hyps(1)] by blast next fix vsa h1 l1 sh1 Ts T pns body D' a assume "e' = throw a" and vals: "P \ \map Val vs,s\<^sub>1\ [\] \map Val vsa,(h1, l1, sh1)\" and method: "P \ C sees M, Static : Ts\T = (pns, body) in D'" and nDone: "\sfs. sh1 D' \ \(sfs, Done)\" and init': "P \ \INIT D' ([D'],False) \ unit,(h1, l1, sh1)\ \ \throw a,s'\" have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h1, l1, sh1)" using evals_finals_same[OF _ vals] map_Val_eq by auto have D: "D = D'" using sees_method_fun[OF SCallInitRed.hyps(1) method] by simp then have i: "i = Processing" using iDP shD s\<^sub>1 s\<^sub>1a nDone by simp then show ?thesis using D init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast next fix vsa h1 l1 sh1 Ts T pns' body' D' v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 h\<^sub>3 l\<^sub>3 sh\<^sub>3 assume s': "s' = (h\<^sub>3, l\<^sub>2, sh\<^sub>3)" and vals: "P \ \map Val vs,s\<^sub>1\ [\] \map Val vsa,(h1, l1, sh1)\" and method: "P \ C sees M, Static : Ts\T = (pns', body') in D'" and nDone: "\sfs. sh1 D' \ \(sfs, Done)\" and init': "P \ \INIT D' ([D'],False) \ unit,(h1, l1, sh1)\ \ \Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2)\" and len: "length vsa = length pns'" and bstep: "P \ \body',(h\<^sub>2, [pns' [\] vsa], sh\<^sub>2)\ \ \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3)\" have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h1, l1, sh1)" using evals_finals_same[OF _ vals] map_Val_eq by auto have D: "D = D'" and pns: "pns = pns'" and body: "body = body'" using sees_method_fun[OF SCallInitRed.hyps(1) method] by auto then have i: "i = Processing" using iDP shD s\<^sub>1 s\<^sub>1a nDone by simp then have s2: "(h\<^sub>2, l\<^sub>2, sh\<^sub>2) = s\<^sub>1" using D init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast then show ?thesis using eval_finalId SCallInit[OF eval_finalsId[of "map Val vs" P "(h,l,sh)"] SCallInitRed.hyps(1)] init init' len bstep nDone D pns body s' s\<^sub>1 s\<^sub>1a shD vals vs SCallInitRed.hyps(2-3) s2 by auto next fix vsa h\<^sub>2 l\<^sub>2 sh\<^sub>2 Ts T pns' body' D' sfs h\<^sub>3 l\<^sub>3 sh\<^sub>3 assume s': "s' = (h\<^sub>3, l\<^sub>2, sh\<^sub>3)" and vals: "P \ \map Val vs,s\<^sub>1\ [\] \map Val vsa,(h\<^sub>2, l\<^sub>2, sh\<^sub>2)\" and method: "P \ C sees M, Static : Ts\T = (pns', body') in D'" and "sh\<^sub>2 D' = \(sfs, Done)\ \ M = clinit \ sh\<^sub>2 D' = \(sfs, Processing)\" and len: "length vsa = length pns'" and bstep: "P \ \body',(h\<^sub>2, [pns' [\] vsa], sh\<^sub>2)\ \ \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3)\" have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)" using evals_finals_same[OF _ vals] map_Val_eq by auto have D: "D = D'" and pns: "pns = pns'" and body: "body = body'" using sees_method_fun[OF SCallInitRed.hyps(1) method] by auto then show ?thesis using SCallInit[OF eval_finalsId[of "map Val vs" P "(h,l,sh)"] SCallInitRed.hyps(1)] bstep SCallInitRed.hyps(2-3) len s' s\<^sub>1a vals vs init by auto qed next fix e assume e': "e' = throw e" and init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \throw e,s'\" obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s') then obtain sfs i where shC: "sh' D = \(sfs, i)\" and iDP: "i = Error" using init_throw_PD[OF init] by auto then show ?thesis using SCallInitRed.hyps(2-3) init e' SCallInitThrow[OF eval_finalsId[of "map Val vs" _] SCallInitRed.hyps(1)] by auto qed next case (RedSCallNone C M vs s b) then have tes: "THROW NoSuchMethodError = e' \ s = s'" using eval_final_same by simp have "P \ \map Val vs,s\ [\] \map Val vs,s\" using eval_finalsId by simp then show ?case using RedSCallNone eval_evals.SCallNone tes by auto next case (RedSCallNonStatic C M Ts T m D vs s b) then have tes: "THROW IncompatibleClassChangeError = e' \ s = s'" using eval_final_same by simp have "P \ \map Val vs,s\ [\] \map Val vs,s\" using eval_finalsId by simp then show ?case using RedSCallNonStatic eval_evals.SCallNonStatic tes by auto next case InitBlockRed thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId simp: assigned_def map_upd_triv fun_upd_same) next case (RedInitBlock V T v u s b) then have "P \ \Val u,s\ \ \e',s'\" by simp then obtain s': "s'=s" and e': "e'=Val u" by cases simp obtain h l sh where s: "s=(h,l,sh)" by (cases s) have "P \ \{V:T :=Val v; Val u},(h,l,sh)\ \ \Val u,(h, (l(V\v))(V:=l V), sh)\" by (fastforce intro!: eval_evals.intros) then have "P \ \{V:T := Val v; Val u},s\ \ \e',s'\" using s s' e' by simp then show ?case by simp next case BlockRedNone thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros simp add: fun_upd_same fun_upd_idem) next case BlockRedSome thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros simp add: fun_upd_same fun_upd_idem) next case (RedBlock V T v s b) then have "P \ \Val v,s\ \ \e',s'\" by simp then obtain s': "s'=s" and e': "e'=Val v" by cases simp obtain h l sh where s: "s=(h,l,sh)" by (cases s) have "P \ \Val v,(h,l(V:=None),sh)\ \ \Val v,(h,l(V:=None),sh)\" by (rule eval_evals.intros) hence "P \ \{V:T;Val v},(h,l,sh)\ \ \Val v,(h,(l(V:=None))(V:=l V),sh)\" by (rule eval_evals.Block) then have "P \ \{V:T; Val v},s\ \ \e',s'\" using s s' e' by simp then show ?case by simp next case (SeqRed e s b e1 s1 b1 e\<^sub>2) show ?case proof(cases "val_of e") case None show ?thesis proof(cases "lass_val_of e") case lNone:None then have bconf: "P,shp s \\<^sub>b (e,b) \" using SeqRed.prems(1) None by simp then show ?thesis using SeqRed using seq_ext by fastforce next case (Some p) obtain V' v' where p: "p = (V',v')" and e: "e = V':=Val v'" using lass_val_of_spec[OF Some] by(cases p, auto) obtain h l sh h' l' sh' where s: "s = (h,l,sh)" and s1: "s1 = (h',l',sh')" by(cases s, cases s1) then have red: "P \ \e,(h,l,sh),b\ \ \e1,(h',l',sh'),b1\" using SeqRed.hyps(1) by simp then have s\<^sub>1': "e1 = unit \ h' = h \ l' = l(V' \ v') \ sh' = sh" using lass_val_of_red[OF Some red] p e by simp then have eval: "P \ \e,s\ \ \e1,s1\" using e s s1 LAss Val by auto then show ?thesis by (metis SeqRed.prems(2) eval_final eval_final_same seq_ext) qed next case (Some a) then show ?thesis using SeqRed.hyps(1) val_no_step by blast qed next case RedSeq thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case CondRed thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros simp: val_no_step) next case RedCondT thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedCondF thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedWhile thus ?case by (auto simp add: unfold_while intro:eval_evals.intros elim:eval_cases) next case ThrowRed then show ?case by(fastforce elim: eval_cases simp: eval_evals.intros) next case RedThrowNull thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case TryRed thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case RedTryCatch thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (RedTryFail s a D fs C V e\<^sub>2 b) thus ?case by (cases s)(auto elim!: eval_cases intro: eval_evals.intros) next case ListRed1 thus ?case by (fastforce elim: evals_cases intro: eval_evals.intros simp: val_no_step) next case ListRed2 thus ?case by (fastforce elim!: evals_cases eval_cases intro: eval_evals.intros eval_finalId) next case (RedInit e1 C b s1 b') then show ?case using InitFinal by simp next case (InitNoneRed sh C C' Cs e h l b) show ?case using InitNone InitNoneRed.hyps InitNoneRed.prems(2) by auto next case (RedInitDone sh C sfs C' Cs e h l b) show ?case using InitDone RedInitDone.hyps RedInitDone.prems(2) by auto next case (RedInitProcessing sh C sfs C' Cs e h l b) show ?case using InitProcessing RedInitProcessing.hyps RedInitProcessing.prems(2) by auto next case (RedInitError sh C sfs C' Cs e h l b) show ?case using InitError RedInitError.hyps RedInitError.prems(2) by auto next case (InitObjectRed sh C sfs sh' C' Cs e h l b) show ?case using InitObject InitObjectRed by auto next case (InitNonObjectSuperRed sh C sfs D r sh' C' Cs e h l b) show ?case using InitNonObject InitNonObjectSuperRed by auto next case (RedInitRInit C' C Cs e h l sh b) show ?case using InitRInit RedInitRInit by auto next case (RInitRed e s b e'' s'' b'' C Cs e\<^sub>0) then have IH: "\e' s'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\" by simp show ?case using RInitRed rinit_ext[OF IH] by simp next case (RedRInit sh C sfs i sh' C' Cs v e h l b s' e') then have init: "P \ \(INIT C' (Cs,True) \ e), (h, l, sh(C \ (sfs, Done)))\ \ \e',s'\" using RedRInit by simp then show ?case using RInit RedRInit.hyps(1) RedRInit.hyps(3) Val by fastforce next case BinOpThrow2 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAssThrow2 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case SFAssThrow then show ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (CallThrowParams es vs e es' v M s b) have val: "P \ \Val v,s\ \ \Val v,s\" by (rule eval_evals.intros) have eval_e: "P \ \throw (e),s\ \ \e',s'\" using CallThrowParams by simp then obtain xa where e': "e' = Throw xa" by (cases) (auto dest!: eval_final) with list_eval_Throw [OF eval_e] have vals: "P \ \es,s\ [\] \map Val vs @ Throw xa # es',s'\" using CallThrowParams.hyps(1) eval_e list_eval_Throw by blast then have "P \ \Val v\M(es),s\ \ \Throw xa,s'\" using eval_evals.CallParamsThrow[OF val vals] by simp thus ?case using e' by simp next case (SCallThrowParams es vs e es' C M s b) have eval_e: "P \ \throw (e),s\ \ \e',s'\" using SCallThrowParams by simp then obtain xa where e': "e' = Throw xa" by (cases) (auto dest!: eval_final) then have "P \ \es,s\ [\] \map Val vs @ Throw xa # es',s'\" using SCallThrowParams.hyps(1) eval_e list_eval_Throw by blast then have "P \ \C\\<^sub>sM(es),s\ \ \Throw xa,s'\" by (rule eval_evals.SCallParamsThrow) thus ?case using e' by simp next case (BlockThrow V T a s b) then have "P \ \Throw a, s\ \ \e',s'\" by simp then obtain s': "s' = s" and e': "e' = Throw a" by cases (auto elim!:eval_cases) obtain h l sh where s: "s=(h,l,sh)" by (cases s) have "P \ \Throw a, (h,l(V:=None),sh)\ \ \Throw a, (h,l(V:=None),sh)\" by (simp add:eval_evals.intros eval_finalId) hence "P\\{V:T;Throw a},(h,l,sh)\ \ \Throw a, (h,(l(V:=None))(V:=l V),sh)\" by (rule eval_evals.Block) then have "P \ \{V:T; Throw a},s\ \ \e',s'\" using s s' e' by simp then show ?case by simp next case (InitBlockThrow V T v a s b) then have "P \ \Throw a,s\ \ \e',s'\" by simp then obtain s': "s' = s" and e': "e' = Throw a" by cases (auto elim!:eval_cases) obtain h l sh where s: "s = (h,l,sh)" by (cases s) have "P \ \{V:T :=Val v; Throw a},(h,l,sh)\ \ \Throw a, (h, (l(V\v))(V:=l V),sh)\" by(fastforce intro:eval_evals.intros) then have "P \ \{V:T := Val v; Throw a},s\ \ \e',s'\" using s s' e' by simp then show ?case by simp next case (RInitInitThrow sh C sfs i sh' a D Cs e h l b) have IH: "\e' s'. P \ \RI (D,Throw a) ; Cs \ e,(h, l, sh(C \ (sfs, Error)))\ \ \e',s'\ \ P \ \RI (C,Throw a) ; D # Cs \ e,(h, l, sh)\ \ \e',s'\" using RInitInitFail[OF eval_finalId] RInitInitThrow by simp then show ?case using RInitInitThrow.hyps(2) RInitInitThrow.prems(2) by auto next case (RInitThrow sh C sfs i sh' a e h l b) then have e': "e' = Throw a" and s': "s' = (h,l,sh')" using eval_final_same final_def by fastforce+ show ?case using RInitFailFinal RInitThrow.hyps(1) RInitThrow.hyps(2) e' eval_finalId s' by auto qed(auto elim: eval_cases simp: eval_evals.intros) (*>*) (*<*) (* ... und wieder anschalten: *) declare split_paired_All [simp] split_paired_Ex [simp] (*>*) text \ Its extension to @{text"\*"}: \ lemma extend_eval: assumes wf: "wwf_J_prog P" shows "\ P \ \e,s,b\ \* \e'',s'',b''\; P \ \e'',s''\ \ \e',s'\; iconf (shp s) e; P,shp s \\<^sub>b (e::expr,b) \ \ \ P \ \e,s\ \ \e',s'\" (*<*) proof (induct rule: converse_rtrancl_induct3) case refl then show ?case by simp next case (step e1 s1 b1 e2 s2 b2) then have ic: "iconf (shp s2) e2" using Red_preserves_iconf local.wf by blast then have ec: "P,shp s2 \\<^sub>b (e2,b2) \" using Red_preserves_bconf local.wf step.hyps(1) step.prems(2) step.prems(3) by blast show ?case using step ic ec extend_1_eval[OF wf step.hyps(1)] by simp qed (*>*) lemma extend_evals: assumes wf: "wwf_J_prog P" shows "\ P \ \es,s,b\ [\]* \es'',s'',b''\; P \ \es'',s''\ [\] \es',s'\; iconfs (shp s) es; P,shp s \\<^sub>b (es::expr list,b) \ \ \ P \ \es,s\ [\] \es',s'\" (*<*) proof (induct rule: converse_rtrancl_induct3) case refl then show ?case by simp next case (step es1 s1 b1 es2 s2 b2) then have ic: "iconfs (shp s2) es2" using Reds_preserves_iconf local.wf by blast then have ec: "P,shp s2 \\<^sub>b (es2,b2) \" using Reds_preserves_bconf local.wf step.hyps(1) step.prems(2) step.prems(3) by blast show ?case using step ic ec extend_1_evals[OF wf step.hyps(1)] by simp qed (*>*) text \ Finally, small step semantics can be simulated by big step semantics: \ theorem assumes wf: "wwf_J_prog P" shows small_by_big: "\P \ \e,s,b\ \* \e',s',b'\; iconf (shp s) e; P,shp s \\<^sub>b (e,b) \; final e'\ \ P \ \e,s\ \ \e',s'\" and "\P \ \es,s,b\ [\]* \es',s',b'\; iconfs (shp s) es; P,shp s \\<^sub>b (es,b) \; finals es'\ \ P \ \es,s\ [\] \es',s'\" (*<*) proof - note wf moreover assume "P \ \e,s,b\ \* \e',s',b'\" moreover assume "final e'" then have "P \ \e',s'\ \ \e',s'\" by (simp add: eval_finalId) moreover assume "iconf (shp s) e" moreover assume "P,shp s \\<^sub>b (e,b) \" ultimately show "P \ \e,s\ \ \e',s'\" by (rule extend_eval) next assume fins: "finals es'" note wf moreover assume "P \ \es,s,b\ [\]* \es',s',b'\" moreover have "P \ \es',s'\ [\] \es',s'\" using fins by (rule eval_finalsId) moreover assume "iconfs (shp s) es" moreover assume "P,shp s \\<^sub>b (es,b) \" ultimately show "P \ \es,s\ [\] \es',s'\" by (rule extend_evals) qed (*>*) subsection "Equivalence" text\ And now, the crowning achievement: \ corollary big_iff_small: "\ wwf_J_prog P; iconf (shp s) e; P,shp s \\<^sub>b (e::expr,b) \ \ \ P \ \e,s\ \ \e',s'\ = (P \ \e,s,b\ \* \e',s',False\ \ final e')" (*<*)by(blast dest: big_by_small eval_final small_by_big)(*>*) corollary big_iff_small_WT: "wwf_J_prog P \ P,E \ e::T \ P,shp s \\<^sub>b (e,b) \ \ P \ \e,s\ \ \e',s'\ = (P \ \e,s,b\ \* \e',s',False\ \ final e')" (*<*)by(blast dest: big_iff_small WT_nsub_RI nsub_RI_iconf)(*>*) subsection \ Lifting type safety to @{text"\"} \ text\ \dots and now to the big step semantics, just for fun. \ lemma eval_preserves_sconf: fixes s::state and s'::state assumes "wf_J_prog P" and "P \ \e,s\ \ \e',s'\" and "iconf (shp s) e" and "P,E \ e::T" and "P,E \ s\" shows "P,E \ s'\" (*<*) proof - have "P,shp s \\<^sub>b (e,False) \" by(simp add: bconf_def) with assms show ?thesis by(blast intro:Red_preserves_sconf Red_preserves_iconf Red_preserves_bconf big_by_small WT_implies_WTrt wf_prog_wwf_prog) qed (*>*) lemma eval_preserves_type: fixes s::state assumes wf: "wf_J_prog P" and "P \ \e,s\ \ \e',s'\" and "P,E \ s\" and "iconf (shp s) e" and "P,E \ e::T" shows "\T'. P \ T' \ T \ P,E,hp s',shp s' \ e':T'" (*<*) proof - have "P,shp s \\<^sub>b (e,False) \" by(simp add: bconf_def) with assms show ?thesis by(blast dest:big_by_small[OF wf_prog_wwf_prog[OF wf]] WT_implies_WTrt Red_preserves_type[OF wf]) qed (*>*) end diff --git a/thys/JinjaThreads/BV/BVProgressThreaded.thy b/thys/JinjaThreads/BV/BVProgressThreaded.thy --- a/thys/JinjaThreads/BV/BVProgressThreaded.thy +++ b/thys/JinjaThreads/BV/BVProgressThreaded.thy @@ -1,1088 +1,1088 @@ (* Title: JinjaThreads/BV/JVMDeadlocked.thy Author: Andreas Lochbihler *) section \Progress result for both of the multithreaded JVMs\ theory BVProgressThreaded imports "../Framework/FWProgress" "../Framework/FWLTS" BVNoTypeError "../JVM/JVMThreaded" begin lemma (in JVM_heap_conf_base') mexec_eq_mexecd: "\ wf_jvm_prog\<^bsub>\\<^esub> P; \ \ t: (xcp, h, frs) \ \ \ mexec P t ((xcp, frs), h) = mexecd P t ((xcp, frs), h)" apply(auto intro!: ext) apply(unfold exec_1_iff) apply(drule no_type_error) apply(assumption) apply(clarify) apply(rule exec_1_d_NormalI) apply(assumption) apply(simp add: exec_d_def split: if_split_asm) apply(erule jvmd_NormalE, auto) done (* conformance lifted to multithreaded case *) context JVM_heap_conf_base begin abbreviation correct_state_ts :: "ty\<^sub>P \ ('addr,'thread_id,'addr jvm_thread_state) thread_info \ 'heap \ bool" where "correct_state_ts \ \ ts_ok (\t (xcp, frstls) h. \ \ t: (xcp, h, frstls) \)" lemma correct_state_ts_thread_conf: "correct_state_ts \ (thr s) (shr s) \ thread_conf P (thr s) (shr s)" by(erule ts_ok_mono)(auto simp add: correct_state_def) lemma invoke_new_thread: assumes "wf_jvm_prog\<^bsub>\\<^esub> P" and "P \ C sees M:Ts\T=\(mxs,mxl0,ins,xt)\ in C" and "ins ! pc = Invoke Type.start 0" and "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and "\ \ t: (None, h, (stk, loc, C, M, pc) # frs) \" and "typeof_addr h (thread_id2addr a) = \Class_type D\" and "P \ D \\<^sup>* Thread" and "P \ D sees run:[]\Void=\(mxs', mxl0', ins',xt')\ in D'" shows "\ \ a: (None, h, [([], Addr (thread_id2addr a) # replicate mxl0' undefined_value, D', run, 0)]) \" proof - from \\ \ t: (None, h, (stk, loc, C, M, pc) # frs) \\ have "hconf h" and "preallocated h" by(simp_all add: correct_state_def) moreover from \P \ D sees run:[]\Void=\(mxs', mxl0', ins',xt')\ in D'\ have "P \ D' sees run:[]\Void=\(mxs', mxl0', ins',xt')\ in D'" by(rule sees_method_idemp) with \wf_jvm_prog\<^bsub>\\<^esub> P\ have "wt_start P D' [] mxl0' (\ D' run)" and "ins' \ []" by(auto dest: wt_jvm_prog_impl_wt_start) then obtain LT' where LT': "\ D' run ! 0 = Some ([], LT')" by (clarsimp simp add: wt_start_def defs1 sup_state_opt_any_Some) moreover have "conf_f P h ([], LT') ins' ([], Addr (thread_id2addr a) # replicate mxl0' undefined_value, D', run, 0)" proof - let ?LT = "OK (Class D') # (replicate mxl0' Err)" have "P,h \ replicate mxl0' undefined_value [:\\<^sub>\] replicate mxl0' Err" by simp also from \P \ D sees run:[]\Void=\(mxs', mxl0', ins',xt')\ in D'\ have "P \ D \\<^sup>* D'" by(rule sees_method_decl_above) with \typeof_addr h (thread_id2addr a) = \Class_type D\\ have "P,h \ Addr (thread_id2addr a) :\ Class D'" by(simp add: conf_def) ultimately have "P,h \ Addr (thread_id2addr a) # replicate mxl0' undefined_value [:\\<^sub>\] ?LT" by(simp) also from \wt_start P D' [] mxl0' (\ D' run)\ LT' have "P \ \ [\\<^sub>\] LT'" by(simp add: wt_start_def) finally have "P,h \ Addr (thread_id2addr a) # replicate mxl0' undefined_value [:\\<^sub>\] LT'" . with \ins' \ []\ show ?thesis by(simp add: conf_f_def) qed moreover from \typeof_addr h (thread_id2addr a) = \Class_type D\\ \P \ D \\<^sup>* Thread\ have "P,h \ a \t" by(rule tconfI) ultimately show ?thesis using \P \ D' sees run:[]\Void=\(mxs', mxl0', ins',xt')\ in D'\ by(fastforce simp add: correct_state_def) qed lemma exec_new_threadE: assumes "wf_jvm_prog\<^bsub>\\<^esub> P" and "P,t \ Normal \ -ta-jvmd\ Normal \'" and "\ \ t: \ \" and "\ta\\<^bsub>t\<^esub> \ []" obtains h frs a stk loc C M pc Ts T mxs mxl0 ins xt M' n Ta ta' va Us Us' U m' D' where "\ = (None, h, (stk, loc, C, M, pc) # frs)" and "(ta, \') \ exec P t (None, h, (stk, loc, C, M, pc) # frs)" and "P \ C sees M: Ts\T = \(mxs, mxl0, ins, xt)\ in C" and "stk ! n = Addr a" and "ins ! pc = Invoke M' n" and "n < length stk" and "typeof_addr h a = \Ta\" and "is_native P Ta M'" and "ta = extTA2JVM P ta'" and "\' = extRet2JVM n m' stk loc C M pc frs va" and "(ta', va, m') \ red_external_aggr P t a M' (rev (take n stk)) h" and "map typeof\<^bsub>h\<^esub> (rev (take n stk)) = map Some Us" and "P \ class_type_of Ta sees M':Us'\U = Native in D'" and "D'\M'(Us') :: U" and "P \ Us [\] Us'" proof - from \P,t \ Normal \ -ta-jvmd\ Normal \'\ obtain h f Frs xcp where check: "check P \" and exec: "(ta, \') \ exec P t \" and [simp]: "\ = (xcp, h, f # Frs)" by(rule jvmd_NormalE) obtain stk loc C M pc where [simp]: "f = (stk, loc, C, M, pc)" by(cases f, blast) from \\ta\\<^bsub>t\<^esub> \ []\ exec have [simp]: "xcp = None" by(cases xcp) auto from \\ \ t: \ \\ obtain Ts T mxs mxl0 ins xt ST LT where "hconf h" "preallocated h" and sees: "P \ C sees M: Ts\T = \(mxs, mxl0, ins, xt)\ in C" and "\ C M ! pc = \(ST, LT)\" and "conf_f P h (ST, LT) ins (stk, loc, C, M, pc)" and "conf_fs P h \ M (length Ts) T Frs" by(fastforce simp add: correct_state_def) from check \\ C M ! pc = \(ST, LT)\\ sees have checkins: "check_instr (ins ! pc) P h stk loc C M pc Frs" by(clarsimp simp add: check_def) from sees \\ta\\<^bsub>t\<^esub> \ []\ exec obtain M' n where [simp]: "ins ! pc = Invoke M' n" by(cases "ins ! pc", auto split: if_split_asm simp add: split_beta ta_upd_simps) from \wf_jvm_prog\<^bsub>\\<^esub> P\ obtain wfmd where wfp: "wf_prog wfmd P" by(auto dest: wt_jvm_progD) from checkins have "n < length stk" "is_Ref (stk ! n)" by auto moreover from exec sees \\ta\\<^bsub>t\<^esub> \ []\ have "stk ! n \ Null" by auto with \is_Ref (stk ! n)\ obtain a where "stk ! n = Addr a" by(auto simp add: is_Ref_def elim: is_AddrE) moreover with checkins obtain Ta where Ta: "typeof_addr h a = \Ta\" by(fastforce) moreover with checkins exec sees \n < length stk\ \\ta\\<^bsub>t\<^esub> \ []\ \stk ! n = Addr a\ obtain Us Us' U D' where "map typeof\<^bsub>h\<^esub> (rev (take n stk)) = map Some Us" and "P \ class_type_of Ta sees M':Us'\U = Native in D'" and "D'\M'(Us') :: U" and "P \ Us [\] Us'" by(auto simp add: confs_conv_map min_def split_beta has_method_def external_WT'_iff split: if_split_asm) moreover with \typeof_addr h a = \Ta\\ \n < length stk\ exec sees \stk ! n = Addr a\ obtain ta' va h' where "ta = extTA2JVM P ta'" "\' = extRet2JVM n h' stk loc C M pc Frs va" "(ta', va, h') \ red_external_aggr P t a M' (rev (take n stk)) h" by(fastforce simp add: min_def) ultimately show thesis using exec sees by-(rule that, auto intro!: is_native.intros) qed end context JVM_conf_read begin lemma correct_state_new_thread: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and red: "P,t \ Normal \ -ta-jvmd\ Normal \'" and cs: "\ \ t: \ \" and nt: "NewThread t'' (xcp, frs) h'' \ set \ta\\<^bsub>t\<^esub>" shows "\ \ t'': (xcp, h'', frs) \" proof - from wf obtain wt where wfp: "wf_prog wt P" by(blast dest: wt_jvm_progD) from nt have "\ta\\<^bsub>t\<^esub> \ []" by auto with wf red cs obtain h Frs a stk loc C M pc Ts T mxs mxl0 ins xt M' n Ta ta' va h' Us Us' U D' where [simp]: "\ = (None, h, (stk, loc, C, M, pc) # Frs)" and exec: "(ta, \') \ exec P t (None, h, (stk, loc, C, M, pc) # Frs)" and sees: "P \ C sees M: Ts\T = \(mxs, mxl0, ins, xt)\ in C" and [simp]: "stk ! n = Addr a" and [simp]: "ins ! pc = Invoke M' n" and n: "n < length stk" and Ta: "typeof_addr h a = \Ta\" and iec: "is_native P Ta M'" and ta: "ta = extTA2JVM P ta'" and \': "\' = extRet2JVM n h' stk loc C M pc Frs va" and rel: "(ta', va, h') \ red_external_aggr P t a M' (rev (take n stk)) h" and Us: "map typeof\<^bsub>h\<^esub> (rev (take n stk)) = map Some Us" and wtext: "P \ class_type_of Ta sees M':Us'\U = Native in D'" "D'\M'(Us') :: U" and sub: "P \ Us [\] Us'" by(rule exec_new_threadE) from cs have hconf: "hconf h" and preh: "preallocated h" and tconf: "P,h \ t \t" by(auto simp add: correct_state_def) from Ta Us wtext sub have wtext': "P,h \ a\M'(rev (take n stk)) : U" by(auto intro!: external_WT'.intros) from rel have red: "P,t \ \a\M'(rev (take n stk)), h\ -ta'\ext \va, h'\" by(unfold WT_red_external_list_conv[OF wfp wtext' tconf]) from ta nt obtain D M'' a' where nt': "NewThread t'' (D, M'', a') h'' \ set \ta'\\<^bsub>t\<^esub>" "(xcp, frs) = extNTA2JVM P (D, M'', a')" by auto with red have [simp]: "h'' = h'" by-(rule red_ext_new_thread_heap) from red_external_new_thread_sub_thread[OF red nt'(1)] have h't'': "typeof_addr h' a' = \Class_type D\" "P \ D \\<^sup>* Thread" and [simp]: "M'' = run" by auto from red_external_new_thread_exists_thread_object[OF red nt'(1)] have tconf': "P,h' \ t'' \t" by(auto intro: tconfI) from sub_Thread_sees_run[OF wfp \P \ D \\<^sup>* Thread\] obtain mxs' mxl0' ins' xt' D' where seesrun: "P \ D sees run: []\Void = \(mxs', mxl0', ins', xt')\ in D'" by auto with nt' ta nt have "xcp = None" "frs = [([],Addr a' # replicate mxl0' undefined_value,D',run,0)]" by(auto simp add: extNTA2JVM_def split_beta) moreover have "\ \ t'': (None, h', [([], Addr a' # replicate mxl0' undefined_value, D', run, 0)]) \" proof - from red wtext' \hconf h\ have "hconf h'" by(rule external_call_hconf) moreover from red have "h \ h'" by(rule red_external_hext) with preh have "preallocated h'" by(rule preallocated_hext) moreover from seesrun have seesrun': "P \ D' sees run: []\Void = \(mxs', mxl0', ins', xt')\ in D'" by(rule sees_method_idemp) moreover with \wf_jvm_prog\<^bsub>\\<^esub> P\ obtain "wt_start P D' [] mxl0' (\ D' run)" "ins' \ []" by (auto dest: wt_jvm_prog_impl_wt_start) then obtain LT' where "\ D' run ! 0 = Some ([], LT')" by (clarsimp simp add: wt_start_def defs1 sup_state_opt_any_Some) moreover have "conf_f P h' ([], LT') ins' ([], Addr a' # replicate mxl0' undefined_value, D', run, 0)" proof - let ?LT = "OK (Class D') # (replicate mxl0' Err)" from seesrun have "P \ D \\<^sup>* D'" by(rule sees_method_decl_above) hence "P,h' \ Addr a' # replicate mxl0' undefined_value [:\\<^sub>\] ?LT" using h't'' by(simp add: conf_def) also from \wt_start P D' [] mxl0' (\ D' run)\ \\ D' run ! 0 = Some ([], LT')\ have "P \ ?LT [\\<^sub>\] LT'" by(simp add: wt_start_def) finally have "P,h' \ Addr a' # replicate mxl0' undefined_value [:\\<^sub>\] LT'" . with \ins' \ []\ show ?thesis by(simp add: conf_f_def) qed ultimately show ?thesis using tconf' by(fastforce simp add: correct_state_def) qed ultimately show ?thesis by(clarsimp) qed lemma correct_state_heap_change: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and red: "P,t \ Normal (xcp, h, frs) -ta-jvmd\ Normal (xcp', h', frs')" and cs: "\ \ t: (xcp, h, frs) \" and cs'': "\ \ t'': (xcp'', h, frs'') \" shows "\ \ t'': (xcp'', h', frs'') \" proof(cases xcp) case None from cs have "P,h \ t \t" by(simp add: correct_state_def) with red have "hext h h'" by (auto intro: exec_1_d_hext simp add: tconf_def) from \wf_jvm_prog\<^bsub>\\<^esub> P\ cs red have "\ \ t: (xcp', h', frs') \" by(auto elim!: jvmd_NormalE intro: BV_correct_1 simp add: exec_1_iff) from cs'' have "P,h \ t'' \t" by(simp add: correct_state_def) with \h \ h'\ have tconf': "P,h' \ t'' \t" by-(rule tconf_hext_mono) from \\ \ t: (xcp', h', frs') \\ have hconf': "hconf h'" "preallocated h'" by(simp_all add: correct_state_def) show ?thesis proof(cases frs'') case Nil thus ?thesis using tconf' hconf' by(simp add: correct_state_def) next case (Cons f'' Frs'') obtain stk'' loc'' C0'' M0'' pc'' where "f'' = (stk'', loc'', C0'', M0'', pc'')" by(cases f'', blast) with \frs'' = f'' # Frs''\ cs'' obtain Ts'' T'' mxs'' mxl\<^sub>0'' ins'' xt'' ST'' LT'' where "hconf h" and sees'': "P \ C0'' sees M0'': Ts''\T'' = \(mxs'', mxl\<^sub>0'', ins'', xt'')\ in C0''" and "\ C0'' M0'' ! pc'' = \(ST'', LT'')\" and "conf_f P h (ST'', LT'') ins'' (stk'', loc'', C0'', M0'', pc'')" and "conf_fs P h \ M0'' (length Ts'') T'' Frs''" by(fastforce simp add: correct_state_def) show ?thesis using Cons \\ \ t'': (xcp'', h, frs'') \\ \hext h h'\ hconf' tconf' apply(cases xcp'') apply(auto simp add: correct_state_def) apply(blast dest: hext_objD intro: conf_fs_hext conf_f_hext)+ done qed next case (Some a) with \P,t \ Normal (xcp, h, frs) -ta-jvmd\ Normal (xcp', h', frs')\ have "h = h'" by(auto elim!: jvmd_NormalE) with \\ \ t'': (xcp'', h, frs'') \\ show ?thesis by simp qed lemma lifting_wf_correct_state_d: "wf_jvm_prog\<^bsub>\\<^esub> P \ lifting_wf JVM_final (mexecd P) (\t (xcp, frs) h. \ \ t: (xcp, h, frs) \)" by(unfold_locales)(auto intro: BV_correct_d_1 correct_state_new_thread correct_state_heap_change) lemma lifting_wf_correct_state: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "lifting_wf JVM_final (mexec P) (\t (xcp, frs) h. \ \ t: (xcp, h, frs) \)" proof(unfold_locales) fix t x m ta x' m' assume "mexec P t (x, m) ta (x', m')" and "(\(xcp, frs) h. \ \ t: (xcp, h, frs) \) x m" with wf show "(\(xcp, frs) h. \ \ t: (xcp, h, frs) \) x' m'" by(cases x)(cases x', simp add: welltyped_commute[symmetric, OF \wf_jvm_prog\<^bsub>\\<^esub> P\], rule BV_correct_d_1) next fix t x m ta x' m' t'' x'' assume "mexec P t (x, m) ta (x', m')" and "(\(xcp, frs) h. \ \ t: (xcp, h, frs) \) x m" and "NewThread t'' x'' m' \ set \ta\\<^bsub>t\<^esub>" with wf show "(\(xcp, frs) h. \ \ t'': (xcp, h, frs) \) x'' m'" apply(cases x, cases x', cases x'', clarify, unfold welltyped_commute[symmetric, OF \wf_jvm_prog\<^bsub>\\<^esub> P\]) by(rule correct_state_new_thread) next fix t x m ta x' m' t'' x'' assume "mexec P t (x, m) ta (x', m')" and "(\(xcp, frs) h. \ \ t: (xcp, h, frs) \) x m" and "(\(xcp, frs) h. \ \ t'': (xcp, h, frs) \) x'' m" with wf show "(\(xcp, frs) h. \ \ t'': (xcp, h, frs) \) x'' m'" by(cases x)(cases x', cases x'', clarify, unfold welltyped_commute[symmetric, OF \wf_jvm_prog\<^bsub>\\<^esub> P\], rule correct_state_heap_change) qed lemmas preserves_correct_state = FWLiftingSem.lifting_wf.RedT_preserves[OF lifting_wf_correct_state] lemmas preserves_correct_state_d = FWLiftingSem.lifting_wf.RedT_preserves[OF lifting_wf_correct_state_d] end context JVM_heap_conf_base begin definition correct_jvm_state :: "ty\<^sub>P \ ('addr,'thread_id,'addr jvm_thread_state,'heap,'addr) state set" where "correct_jvm_state \ = {s. correct_state_ts \ (thr s) (shr s) \ lock_thread_ok (locks s) (thr s)}" end context JVM_heap_conf begin lemma correct_jvm_state_initial: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and wf_start: "wf_start_state P C M vs" shows "JVM_start_state P C M vs \ correct_jvm_state \" proof - from wf_start obtain Ts T m D where "start_heap_ok" and "P \ C sees M:Ts\T = \m\ in D" and "P,start_heap \ vs [:\] Ts" by cases with wf BV_correct_initial[OF wf this] show ?thesis by(cases m)(auto simp add: correct_jvm_state_def start_state_def JVM_start_state'_def intro: lock_thread_okI ts_okI split: if_split_asm) qed end context JVM_conf_read begin lemma invariant3p_correct_jvm_state_mexecdT: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "invariant3p (mexecdT P) (correct_jvm_state \)" unfolding correct_jvm_state_def apply(rule invariant3pI) apply safe apply(erule (1) lifting_wf.redT_preserves[OF lifting_wf_correct_state_d[OF wf]]) apply(erule (1) execd_mthr.redT_preserves_lock_thread_ok) done lemma invariant3p_correct_jvm_state_mexecT: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "invariant3p (mexecT P) (correct_jvm_state \)" unfolding correct_jvm_state_def apply(rule invariant3pI) apply safe apply(erule (1) lifting_wf.redT_preserves[OF lifting_wf_correct_state[OF wf]]) apply(erule (1) exec_mthr.redT_preserves_lock_thread_ok) done lemma correct_jvm_state_preserved: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and correct: "s \ correct_jvm_state \" and red: "P \ s -\ttas\\<^bsub>jvm\<^esub>* s'" shows "s' \ correct_jvm_state \" using wf red correct unfolding exec_mthr.RedT_def by(rule invariant3p_rtrancl3p[OF invariant3p_correct_jvm_state_mexecT]) theorem jvm_typesafe: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and start: "wf_start_state P C M vs" and exec: "P \ JVM_start_state P C M vs -\ttas\\<^bsub>jvm\<^esub>* s'" shows "s' \ correct_jvm_state \" by(rule correct_jvm_state_preserved[OF wf _ exec])(rule correct_jvm_state_initial[OF wf start]) end declare (in JVM_typesafe) split_paired_Ex [simp del] context JVM_heap_conf_base' begin lemma execd_NewThread_Thread_Object: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and conf: "\ \ t': \ \" and red: "P,t' \ Normal \ -ta-jvmd\ Normal \'" and nt: "NewThread t x m \ set \ta\\<^bsub>t\<^esub>" shows "\C. typeof_addr (fst (snd \')) (thread_id2addr t) = \Class_type C\ \ P \ Class C \ Class Thread" proof - from wf obtain wfmd where wfp: "wf_prog wfmd P" by(blast dest: wt_jvm_progD) from red obtain h f Frs xcp where check: "check P \" and exec: "(ta, \') \ exec P t' \" and [simp]: "\ = (xcp, h, f # Frs)" by(rule jvmd_NormalE) obtain xcp' h' frs' where [simp]: "\' = (xcp', h', frs')" by(cases \', auto) obtain stk loc C M pc where [simp]: "f = (stk, loc, C, M, pc)" by(cases f, blast) from exec nt have [simp]: "xcp = None" by(cases xcp, auto) from \\ \ t': \ \\ obtain Ts T mxs mxl0 ins xt ST LT where "hconf h" and "P,h \ t' \t" and sees: "P \ C sees M: Ts\T = \(mxs, mxl0, ins, xt)\ in C" and "\ C M ! pc = \(ST, LT)\" and "conf_f P h (ST, LT) ins (stk, loc, C, M, pc)" and "conf_fs P h \ M (length Ts) T Frs" by(fastforce simp add: correct_state_def) from wf red conf nt obtain h frs a stk loc C M pc M' n ta' va h' where ha: "typeof_addr h a \ None" and ta: "ta = extTA2JVM P ta'" and \': "\' = extRet2JVM n h' stk loc C M pc frs va" and rel: "(ta', va, h') \ red_external_aggr P t' a M' (rev (take n stk)) h" by -(erule (2) exec_new_threadE, fastforce+) from nt ta obtain x' where "NewThread t x' m \ set \ta'\\<^bsub>t\<^esub>" by auto from red_external_aggr_new_thread_exists_thread_object[OF rel ha this] \' show ?thesis by(cases va) auto qed lemma mexecdT_NewThread_Thread_Object: "\ wf_jvm_prog\<^bsub>\\<^esub> P; correct_state_ts \ (thr s) (shr s); P \ s -t'\ta\\<^bsub>jvmd\<^esub> s'; NewThread t x m \ set \ta\\<^bsub>t\<^esub> \ \ \C. typeof_addr (shr s') (thread_id2addr t) = \Class_type C\ \ P \ C \\<^sup>* Thread" apply(frule correct_state_ts_thread_conf) apply(erule execd_mthr.redT.cases) apply(hypsubst) apply(frule (2) execd_tconf.redT_updTs_preserves[where ln'="redT_updLns (locks s) t' no_wait_locks \ta\\<^bsub>l\<^esub>"]) apply clarsimp apply(clarsimp) apply(drule execd_NewThread_Thread_Object) apply(drule (1) ts_okD) apply(fastforce) apply(assumption) apply(fastforce) apply(clarsimp) apply(simp) done end context JVM_heap begin lemma exec_ta_satisfiable: assumes "P,t \ s -ta-jvm\ s'" shows "\s. exec_mthr.actions_ok s t ta" proof - obtain xcp h frs where [simp]: "s = (xcp, h, frs)" by(cases s) from assms obtain stk loc C M pc frs' where [simp]: "frs = (stk, loc, C, M, pc) # frs'" by(cases frs)(auto simp add: exec_1_iff) show ?thesis proof(cases xcp) case Some with assms show ?thesis by(auto simp add: exec_1_iff lock_ok_las_def finfun_upd_apply split_paired_Ex) next case None with assms show ?thesis apply(cases "instrs_of P C M ! pc") apply(auto simp add: exec_1_iff lock_ok_las_def finfun_upd_apply split_beta final_thread.actions_ok_iff split: if_split_asm dest: red_external_aggr_ta_satisfiable[where final=JVM_final]) apply(fastforce simp add: final_thread.actions_ok_iff lock_ok_las_def dest: red_external_aggr_ta_satisfiable[where final=JVM_final]) apply(fastforce simp add: finfun_upd_apply intro: exI[where x="K$ None"] exI[where x="K$ \(t, 0)\"] may_lock.intros)+ done qed qed end context JVM_typesafe begin lemma execd_wf_progress: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "progress JVM_final (mexecd P) (execd_mthr.wset_Suspend_ok P (correct_jvm_state \))" (is "progress _ _ ?wf_state") proof { fix s t x ta x' m' w assume mexecd: "mexecd P t (x, shr s) ta (x', m')" and Suspend: "Suspend w \ set \ta\\<^bsub>w\<^esub>" from mexecd_Suspend_Invoke[OF mexecd Suspend] show "\ JVM_final x'" by auto } note Suspend_final = this { fix s assume s: "s \ ?wf_state" hence "lock_thread_ok (locks s) (thr s)" by(auto dest: execd_mthr.wset_Suspend_okD1 simp add: correct_jvm_state_def) moreover have "exec_mthr.wset_final_ok (wset s) (thr s)" proof(rule exec_mthr.wset_final_okI) fix t w assume "wset s t = \w\" from execd_mthr.wset_Suspend_okD2[OF s this] obtain x0 ta x m1 w' ln'' and s0 :: "('addr, 'thread_id, 'addr option \ 'addr frame list, 'heap, 'addr) state" where mexecd: "mexecd P t (x0, shr s0) ta (x, m1)" and Suspend: "Suspend w' \ set \ta\\<^bsub>w\<^esub>" and tst: "thr s t = \(x, ln'')\" by blast from Suspend_final[OF mexecd Suspend] tst show " \x ln. thr s t = \(x, ln)\ \ \ JVM_final x" by blast qed ultimately show "lock_thread_ok (locks s) (thr s) \ exec_mthr.wset_final_ok (wset s) (thr s)" .. } next fix s t x ta x' m' assume wfs: "s \ ?wf_state" and "thr s t = \(x, no_wait_locks)\" and "mexecd P t (x, shr s) ta (x', m')" and wait: "\ waiting (wset s t)" moreover obtain ls ts h ws "is" where s [simp]: "s = (ls, (ts, h), ws, is)" by(cases s) fastforce ultimately have "ts t = \(x, no_wait_locks)\" "mexecd P t (x, h) ta (x', m')" by auto from wfs have "correct_state_ts \ ts h" by(auto dest: execd_mthr.wset_Suspend_okD1 simp add: correct_jvm_state_def) from wf obtain wfmd where wfp: "wf_prog wfmd P" by(auto dest: wt_jvm_progD) from \ts t = \(x, no_wait_locks)\\ \mexecd P t (x, h) ta (x', m')\ obtain xcp frs xcp' frs' where "P,t \ Normal (xcp, h, frs) -ta-jvmd\ Normal (xcp', m', frs')" and [simp]: "x = (xcp, frs)" "x' = (xcp', frs')" by(cases x, auto) then obtain f Frs where check: "check P (xcp, h, f # Frs)" and [simp]: "frs = f # Frs" and exec: "(ta, xcp', m', frs') \ exec P t (xcp, h, f # Frs)" by(auto elim: jvmd_NormalE) with \ts t = \(x, no_wait_locks)\\ \correct_state_ts \ ts h\ have correct: "\ \ t: (xcp, h, f # Frs) \" by(auto dest: ts_okD) obtain stk loc C M pc where f [simp]: "f = (stk, loc, C, M, pc)" by (cases f) from correct obtain Ts T mxs mxl0 ins xt ST LT where hconf: "hconf h" and tconf: "P, h \ t \t" and sees: "P \ C sees M:Ts\T = \(mxs, mxl0, ins, xt)\ in C" and wt: "\ C M ! pc = \(ST, LT)\" and conf_f: "conf_f P h (ST, LT) ins (stk, loc, C, M, pc)" and confs: "conf_fs P h \ M (length Ts) T Frs" and confxcp: "conf_xcp P h xcp (ins ! pc)" and preh: "preallocated h" by(fastforce simp add: correct_state_def) have "\ta' \'. P,t \ Normal (xcp, h, (stk, loc, C, M, pc) # Frs) -ta'-jvmd\ Normal \' \ (final_thread.actions_ok JVM_final (ls, (ts, h), ws, is) t ta' \ final_thread.actions_ok' (ls, (ts, h), ws, is) t ta' \ final_thread.actions_subset ta' ta)" proof(cases "final_thread.actions_ok' (ls, (ts, h), ws, is) t ta") case True have "final_thread.actions_subset ta ta" .. with True \P,t \ Normal (xcp, h, frs) -ta-jvmd\ Normal (xcp', m', frs')\ show ?thesis by auto next case False note naok = this have ws: "wset s t = None \ (\n a T w. ins ! pc = Invoke wait n \ stk ! n = Addr a \ typeof_addr h a = \T\ \ is_native P T wait \ wset s t = \PostWS w\ \ xcp = None)" proof(cases "wset s t") case None thus ?thesis .. next case (Some w) from execd_mthr.wset_Suspend_okD2[OF wfs this] \ts t = \(x, no_wait_locks)\\ obtain xcp0 frs0 h0 ta0 w' s1 tta1 where red0: "mexecd P t ((xcp0, frs0), h0) ta0 ((xcp, frs), shr s1)" and Suspend: "Suspend w' \ set \ta0\\<^bsub>w\<^esub>" and s1: "P \ s1 -\tta1\\<^bsub>jvmd\<^esub>* s" by auto from mexecd_Suspend_Invoke[OF red0 Suspend] sees obtain n a T where [simp]: "ins ! pc = Invoke wait n" "xcp = None" "stk ! n = Addr a" and type: "typeof_addr h0 a = \T\" and iec: "is_native P T wait" by(auto simp add: is_native.simps) blast from red0 have "h0 \ shr s1" by(auto dest: exec_1_d_hext) also from s1 have "shr s1 \ shr s" by(rule Execd_hext) finally have "typeof_addr (shr s) a = \T\" using type by(rule typeof_addr_hext_mono) moreover from Some wait s obtain w' where "ws t = \PostWS w'\" by(auto simp add: not_waiting_iff) ultimately show ?thesis using iec s by auto qed from ws naok exec sees show ?thesis proof(cases "ins ! pc") case (Invoke M' n) from ws Invoke check exec sees naok obtain a Ts U Ta Us D D' where a: "stk ! n = Addr a" and n: "n < length stk" and Ta: "typeof_addr h a = \Ta\" and wtext: "P \ class_type_of Ta sees M':Us\U = Native in D'" "D'\M'(Us)::U" and sub: "P \ Ts [\] Us" and Ts: "map typeof\<^bsub>h\<^esub> (rev (take n stk)) = map Some Ts" and [simp]: "xcp = None" apply(cases xcp) apply(simp add: is_Ref_def has_method_def external_WT'_iff check_def lock_ok_las'_def confs_conv_map split_beta split: if_split_asm option.splits) apply(auto simp add: lock_ok_las'_def)[2] apply(fastforce simp add: is_native.simps lock_ok_las'_def dest: sees_method_fun)+ done from exec Ta n a sees Invoke wtext obtain ta' va m'' where exec': "(ta', va, m'') \ red_external_aggr P t a M' (rev (take n stk)) h" and ta: "ta = extTA2JVM P ta'" and va: "(xcp', m', frs') = extRet2JVM n m'' stk loc C M pc Frs va" by(auto) from va have [simp]: "m'' = m'" by(cases va) simp_all from Ta Ts wtext sub have wtext': "P,h \ a\M'(rev (take n stk)) : U" by(auto intro!: external_WT'.intros simp add: is_native.simps) with wfp exec' tconf have red: "P,t \ \a\M'(rev (take n stk)), h\ -ta'\ext \va, m'\" by(simp add: WT_red_external_list_conv) from ws Invoke have "wset s t = None \ M' = wait \ (\w. wset s t = \PostWS w\)" by auto with wfp red tconf hconf obtain ta'' va' h'' where red': "P,t \ \a\M'(rev (take n stk)),h\ -ta''\ext \va',h''\" and ok': "final_thread.actions_ok JVM_final s t ta'' \ final_thread.actions_ok' s t ta'' \ final_thread.actions_subset ta'' ta'" by(rule red_external_wf_red) from red' a n Ta Invoke sees wtext have "(extTA2JVM P ta'', extRet2JVM n h'' stk loc C M pc Frs va') \ exec P t (xcp, h, f # Frs)" by(auto intro: red_external_imp_red_external_aggr) with check have "P,t \ Normal (xcp, h, (stk, loc, C, M, pc) # Frs) -extTA2JVM P ta''-jvmd\ Normal (extRet2JVM n h'' stk loc C M pc Frs va')" by -(rule exec_1_d.exec_1_d_NormalI, auto simp add: exec_d_def) moreover from ok' ta have "final_thread.actions_ok JVM_final (ls, (ts, h), ws, is) t (extTA2JVM P ta'') \ final_thread.actions_ok' (ls, (ts, h), ws, is) t (extTA2JVM P ta'') \ final_thread.actions_subset (extTA2JVM P ta'') ta" by(auto simp add: final_thread.actions_ok'_convert_extTA elim: final_thread.actions_subset.cases del: subsetI) ultimately show ?thesis by blast next case MEnter with exec sees naok ws have False by(cases xcp)(auto split: if_split_asm simp add: lock_ok_las'_def finfun_upd_apply ta_upd_simps) thus ?thesis .. next case MExit with exec sees False check ws obtain a where [simp]: "hd stk = Addr a" "xcp = None" "ws t = None" and ta: "ta = \Unlock\a, SyncUnlock a\ \ ta = \UnlockFail\a\" by(cases xcp)(fastforce split: if_split_asm simp add: lock_ok_las'_def finfun_upd_apply is_Ref_def check_def)+ from ta show ?thesis proof(rule disjE) assume ta: "ta = \Unlock\a, SyncUnlock a\" let ?ta' = "\UnlockFail\a\" from ta exec sees MExit obtain \' where "(?ta', \') \ exec P t (xcp, h, f # Frs)" by auto with check have "P,t \ Normal (xcp, h, (stk, loc, C, M, pc) # Frs) -?ta'-jvmd\ Normal \'" by -(rule exec_1_d.exec_1_d_NormalI, auto simp add: exec_d_def) moreover from False ta have "has_locks (ls $ a) t = 0" by(auto simp add: lock_ok_las'_def finfun_upd_apply ta_upd_simps) hence "final_thread.actions_ok' (ls, (ts, h), ws, is) t ?ta'" by(auto simp add: lock_ok_las'_def finfun_upd_apply ta_upd_simps) moreover from ta have "final_thread.actions_subset ?ta' ta" by(auto simp add: final_thread.actions_subset_iff collect_locks'_def finfun_upd_apply ta_upd_simps) ultimately show ?thesis by(fastforce simp add: ta_upd_simps) next assume ta: "ta = \UnlockFail\a\" let ?ta' = "\Unlock\a, SyncUnlock a\" from ta exec sees MExit obtain \' where "(?ta', \') \ exec P t (xcp, h, f # Frs)" by auto with check have "P,t \ Normal (xcp, h, (stk, loc, C, M, pc) # Frs) -?ta'-jvmd\ Normal \'" by -(rule exec_1_d.exec_1_d_NormalI, auto simp add: exec_d_def) moreover from False ta have "has_lock (ls $ a) t" by(auto simp add: lock_ok_las'_def finfun_upd_apply ta_upd_simps) hence "final_thread.actions_ok' (ls, (ts, h), ws, is) t ?ta'" by(auto simp add: lock_ok_las'_def finfun_upd_apply ta_upd_simps) moreover from ta have "final_thread.actions_subset ?ta' ta" by(auto simp add: final_thread.actions_subset_iff collect_locks'_def finfun_upd_apply ta_upd_simps) ultimately show ?thesis by(fastforce simp add: ta_upd_simps) qed qed(case_tac [!] xcp, auto simp add: split_beta lock_ok_las'_def split: if_split_asm) qed thus "\ta' x' m'. mexecd P t (x, shr s) ta' (x', m') \ (final_thread.actions_ok JVM_final s t ta' \ final_thread.actions_ok' s t ta' \ final_thread.actions_subset ta' ta)" by fastforce next fix s t x assume wfs: "s \ ?wf_state" and tst: "thr s t = \(x, no_wait_locks)\" and "\ JVM_final x" from wfs have correct: "correct_state_ts \ (thr s) (shr s)" by(auto dest: execd_mthr.wset_Suspend_okD1 simp add: correct_jvm_state_def) obtain xcp frs where x: "x = (xcp, frs)" by (cases x, auto) with \\ JVM_final x\ obtain f Frs where "frs = f # Frs" by(fastforce simp add: neq_Nil_conv) with tst correct x have "\ \ t: (xcp, shr s, f # Frs) \" by(auto dest: ts_okD) with \wf_jvm_prog\<^bsub>\\<^esub> P\ have "exec_d P t (xcp, shr s, f # Frs) \ TypeError" by(auto dest: no_type_error) then obtain \ where "exec_d P t (xcp, shr s, f # Frs) = Normal \" by(auto) hence "exec P t (xcp, shr s, f # Frs) = \" by(auto simp add: exec_d_def check_def split: if_split_asm) with progress[OF wf \\ \ t: (xcp, shr s, f # Frs) \\] obtain ta \ where "(ta, \) \ \" unfolding exec_1_iff by blast with \x = (xcp, frs)\ \frs = f # Frs\ \\ \ t: (xcp, shr s, f # Frs) \\ \wf_jvm_prog\<^bsub>\\<^esub> P\ \exec_d P t (xcp, shr s, f # Frs) = Normal \\ show "\ta x' m'. mexecd P t (x, shr s) ta (x', m')" by(cases ta, cases \)(fastforce simp add: split_paired_Ex intro: exec_1_d_NormalI) qed(fastforce dest: defensive_imp_aggressive_1 mexec_instr_Wakeup_no_Join exec_ta_satisfiable)+ end context JVM_conf_read begin lemma mexecT_eq_mexecdT: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and cs: "correct_state_ts \ (thr s) (shr s)" shows "P \ s -t\ta\\<^bsub>jvm\<^esub> s' = P \ s -t\ta\\<^bsub>jvmd\<^esub> s'" proof(rule iffI) assume "P \ s -t\ta\\<^bsub>jvm\<^esub> s'" thus "P \ s -t\ta\\<^bsub>jvmd\<^esub> s'" proof(cases rule: exec_mthr.redT_elims[consumes 1, case_names normal acquire]) case (normal x x' m') obtain xcp frs where x [simp]: "x = (xcp, frs)" by(cases x, auto) from \thr s t = \(x, no_wait_locks)\\ cs have "\ \ t: (xcp, shr s, frs) \" by(auto dest: ts_okD) from mexec_eq_mexecd[OF wf \\ \ t: (xcp, shr s, frs) \\] \mexec P t (x, shr s) ta (x', m')\ have *: "mexecd P t (x, shr s) ta (x', m')" by simp with lifting_wf.redT_updTs_preserves[OF lifting_wf_correct_state_d[OF wf] cs, OF this \thr s t = \(x, no_wait_locks)\\] \thread_oks (thr s) \ta\\<^bsub>t\<^esub>\ - have "correct_state_ts \ (redT_updTs (thr s) \ta\\<^bsub>t\<^esub>(t \ (x', redT_updLns (locks s) t no_wait_locks \ta\\<^bsub>l\<^esub>))) m'" by simp + have "correct_state_ts \ ((redT_updTs (thr s) \ta\\<^bsub>t\<^esub>)(t \ (x', redT_updLns (locks s) t no_wait_locks \ta\\<^bsub>l\<^esub>))) m'" by simp with * show ?thesis using normal by(cases s')(erule execd_mthr.redT_normal, auto) next case acquire thus ?thesis apply(cases s', clarify) apply(rule execd_mthr.redT_acquire, assumption+) by(auto) qed next assume "P \ s -t\ta\\<^bsub>jvmd\<^esub> s'" thus "P \ s -t\ta\\<^bsub>jvm\<^esub> s'" proof(cases rule: execd_mthr.redT_elims[consumes 1, case_names normal acquire]) case (normal x x' m') obtain xcp frs where x [simp]: "x = (xcp, frs)" by(cases x, auto) from \thr s t = \(x, no_wait_locks)\\ cs have "\ \ t: (xcp, shr s, frs) \" by(auto dest: ts_okD) from mexec_eq_mexecd[OF wf \\ \ t: (xcp, shr s, frs) \\] \mexecd P t (x, shr s) ta (x', m')\ have "mexec P t (x, shr s) ta (x', m')" by simp moreover from lifting_wf.redT_updTs_preserves[OF lifting_wf_correct_state_d[OF wf] cs, OF \mexecd P t (x, shr s) ta (x', m')\ \thr s t = \(x, no_wait_locks)\\] \thread_oks (thr s) \ta\\<^bsub>t\<^esub>\ - have "correct_state_ts \ (redT_updTs (thr s) \ta\\<^bsub>t\<^esub>(t \ (x', redT_updLns (locks s) t no_wait_locks \ta\\<^bsub>l\<^esub>))) m'" by simp + have "correct_state_ts \ ((redT_updTs (thr s) \ta\\<^bsub>t\<^esub>)(t \ (x', redT_updLns (locks s) t no_wait_locks \ta\\<^bsub>l\<^esub>))) m'" by simp ultimately show ?thesis using normal by(cases s')(erule exec_mthr.redT_normal, auto) next case acquire thus ?thesis apply(cases s', clarify) apply(rule exec_mthr.redT_acquire, assumption+) by(auto) qed qed lemma mExecT_eq_mExecdT: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and ct: "correct_state_ts \ (thr s) (shr s)" shows "P \ s -\ttas\\<^bsub>jvm\<^esub>* s' = P \ s -\ttas\\<^bsub>jvmd\<^esub>* s'" proof assume Red: "P \ s -\ttas\\<^bsub>jvm\<^esub>* s'" thus "P \ s -\ttas\\<^bsub>jvmd\<^esub>* s'" using ct proof(induct rule: exec_mthr.RedT_induct[consumes 1, case_names refl step]) case refl thus ?case by auto next case (step s ttas s' t ta s'') hence "P \ s -\ttas\\<^bsub>jvmd\<^esub>* s'" by blast moreover from \correct_state_ts \ (thr s) (shr s)\ \P \ s -\ttas\\<^bsub>jvm\<^esub>* s'\ have "correct_state_ts \ (thr s') (shr s')" by(auto dest: preserves_correct_state[OF wf]) with \P \ s' -t\ta\\<^bsub>jvm\<^esub> s''\ have "P \ s' -t\ta\\<^bsub>jvmd\<^esub> s''" by(unfold mexecT_eq_mexecdT[OF wf]) ultimately show ?case by(blast intro: execd_mthr.RedTI rtrancl3p_step elim: execd_mthr.RedTE) qed next assume Red: "P \ s -\ttas\\<^bsub>jvmd\<^esub>* s'" thus "P \ s -\ttas\\<^bsub>jvm\<^esub>* s'" using ct proof(induct rule: execd_mthr.RedT_induct[consumes 1, case_names refl step]) case refl thus ?case by auto next case (step s ttas s' t ta s'') hence "P \ s -\ttas\\<^bsub>jvm\<^esub>* s'" by blast moreover from \correct_state_ts \ (thr s) (shr s)\ \P \ s -\ttas\\<^bsub>jvmd\<^esub>* s'\ have "correct_state_ts \ (thr s') (shr s')" by(auto dest: preserves_correct_state_d[OF wf]) with \P \ s' -t\ta\\<^bsub>jvmd\<^esub> s''\ have "P \ s' -t\ta\\<^bsub>jvm\<^esub> s''" by(unfold mexecT_eq_mexecdT[OF wf]) ultimately show ?case by(blast intro: exec_mthr.RedTI rtrancl3p_step elim: exec_mthr.RedTE) qed qed lemma mexecT_preserves_thread_conf: "\ wf_jvm_prog\<^bsub>\\<^esub> P; correct_state_ts \ (thr s) (shr s); P \ s -t'\ta\\<^bsub>jvm\<^esub> s'; thread_conf P (thr s) (shr s) \ \ thread_conf P (thr s') (shr s')" by(simp only: mexecT_eq_mexecdT)(rule execd_tconf.redT_preserves) lemma mExecT_preserves_thread_conf: "\ wf_jvm_prog\<^bsub>\\<^esub> P; correct_state_ts \ (thr s) (shr s); P \ s -\tta\\<^bsub>jvm\<^esub>* s'; thread_conf P (thr s) (shr s) \ \ thread_conf P (thr s') (shr s')" by(simp only: mExecT_eq_mExecdT)(rule execd_tconf.RedT_preserves) lemma wset_Suspend_ok_mexecd_mexec: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "exec_mthr.wset_Suspend_ok P (correct_jvm_state \) = execd_mthr.wset_Suspend_ok P (correct_jvm_state \)" apply(safe) apply(rule execd_mthr.wset_Suspend_okI) apply(erule exec_mthr.wset_Suspend_okD1) apply(drule (1) exec_mthr.wset_Suspend_okD2) apply(subst (asm) (2) split_paired_Ex) apply(elim bexE exE conjE) apply(subst (asm) mexec_eq_mexecd[OF wf]) apply(simp add: correct_jvm_state_def) apply(blast dest: ts_okD) apply(subst (asm) mexecT_eq_mexecdT[OF wf]) apply(simp add: correct_jvm_state_def) apply(subst (asm) mExecT_eq_mExecdT[OF wf]) apply(simp add: correct_jvm_state_def) apply(rule bexI exI|erule conjI|assumption)+ apply(rule exec_mthr.wset_Suspend_okI) apply(erule execd_mthr.wset_Suspend_okD1) apply(drule (1) execd_mthr.wset_Suspend_okD2) apply(subst (asm) (2) split_paired_Ex) apply(elim bexE exE conjE) apply(subst (asm) mexec_eq_mexecd[OF wf, symmetric]) apply(simp add: correct_jvm_state_def) apply(blast dest: ts_okD) apply(subst (asm) mexecT_eq_mexecdT[OF wf, symmetric]) apply(simp add: correct_jvm_state_def) apply(subst (asm) mExecT_eq_mExecdT[OF wf, symmetric]) apply(simp add: correct_jvm_state_def) apply(rule bexI exI|erule conjI|assumption)+ done end context JVM_typesafe begin lemma exec_wf_progress: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "progress JVM_final (mexec P) (exec_mthr.wset_Suspend_ok P (correct_jvm_state \))" (is "progress _ _ ?wf_state") proof - interpret progress: progress JVM_final "mexecd P" convert_RA ?wf_state using assms unfolding wset_Suspend_ok_mexecd_mexec[OF wf] by(rule execd_wf_progress) show ?thesis proof(unfold_locales) fix s assume "s \ ?wf_state" thus "lock_thread_ok (locks s) (thr s) \ exec_mthr.wset_final_ok (wset s) (thr s)" by(rule progress.wf_stateD) next fix s t x ta x' m' assume wfs: "s \ ?wf_state" and tst: "thr s t = \(x, no_wait_locks)\" and exec: "mexec P t (x, shr s) ta (x', m')" and wait: "\ waiting (wset s t)" from wfs tst have correct: "\ \ t: (fst x, shr s, snd x) \" by(auto dest!: exec_mthr.wset_Suspend_okD1 ts_okD simp add: correct_jvm_state_def) with exec have "mexecd P t (x, shr s) ta (x', m')" by(cases x)(simp only: mexec_eq_mexecd[OF wf] fst_conv snd_conv) from progress.wf_red[OF wfs tst this wait] correct show "\ta' x' m'. mexec P t (x, shr s) ta' (x', m') \ (final_thread.actions_ok JVM_final s t ta' \ final_thread.actions_ok' s t ta' \ final_thread.actions_subset ta' ta)" by(cases x)(simp only: fst_conv snd_conv mexec_eq_mexecd[OF wf]) next fix s t x ta x' m' w assume wfs: "s \ ?wf_state" and tst: "thr s t = \(x, no_wait_locks)\" and exec: "mexec P t (x, shr s) ta (x', m')" and wait: "\ waiting (wset s t)" and Suspend: "Suspend w \ set \ta\\<^bsub>w\<^esub>" from wfs tst have correct: "\ \ t: (fst x, shr s, snd x) \" by(auto dest!: exec_mthr.wset_Suspend_okD1 ts_okD simp add: correct_jvm_state_def) with exec have "mexecd P t (x, shr s) ta (x', m')" by(cases x)(simp only: mexec_eq_mexecd[OF wf] fst_conv snd_conv) with wfs tst show "\ JVM_final x'" using wait Suspend by(rule progress.red_wait_set_not_final) next fix s t x assume wfs: "s \ ?wf_state" and tst: "thr s t = \(x, no_wait_locks)\" and "\ JVM_final x" from progress.wf_progress[OF this] show "\ta x' m'. mexec P t (x, shr s) ta (x', m')" by(auto dest: defensive_imp_aggressive_1 simp add: split_beta) qed(fastforce dest: mexec_instr_Wakeup_no_Join exec_ta_satisfiable)+ qed theorem mexecd_TypeSafety: fixes ln :: "'addr \f nat" assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and s: "s \ execd_mthr.wset_Suspend_ok P (correct_jvm_state \)" and Exec: "P \ s -\ttas\\<^bsub>jvmd\<^esub>* s'" and "\ (\t ta s''. P \ s' -t\ta\\<^bsub>jvmd\<^esub> s'')" and ts't: "thr s' t = \((xcp, frs), ln)\" shows "frs \ [] \ ln \ no_wait_locks \ t \ execd_mthr.deadlocked P s'" and "\ \ t: (xcp, shr s', frs) \" proof - interpret progress JVM_final "mexecd P" convert_RA "execd_mthr.wset_Suspend_ok P (correct_jvm_state \)" by(rule execd_wf_progress) fact+ from Exec s have wfs': "s' \ execd_mthr.wset_Suspend_ok P (correct_jvm_state \)" unfolding execd_mthr.RedT_def by(blast intro: invariant3p_rtrancl3p execd_mthr.invariant3p_wset_Suspend_ok invariant3p_correct_jvm_state_mexecdT[OF wf]) with ts't show cst: "\ \ t: (xcp, shr s', frs) \" by(auto dest: ts_okD execd_mthr.wset_Suspend_okD1 simp add: correct_jvm_state_def) assume nfin: "frs \ [] \ ln \ no_wait_locks" from nfin \thr s' t = \((xcp, frs), ln)\\ have "exec_mthr.not_final_thread s' t" by(auto simp: exec_mthr.not_final_thread_iff) from \\ (\t ta s''. P \ s' -t\ta\\<^bsub>jvmd\<^esub> s'')\ show "t \ execd_mthr.deadlocked P s'" proof(rule contrapos_np) assume "t \ execd_mthr.deadlocked P s'" with \exec_mthr.not_final_thread s' t\ have "\ execd_mthr.deadlocked' P s'" by(auto simp add: execd_mthr.deadlocked'_def) hence "\ execd_mthr.deadlock P s'" unfolding execd_mthr.deadlock_eq_deadlocked' . thus "\t ta s''. P \ s' -t\ta\\<^bsub>jvmd\<^esub> s''" by(rule redT_progress[OF wfs']) qed qed theorem mexec_TypeSafety: fixes ln :: "'addr \f nat" assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and s: "s \ exec_mthr.wset_Suspend_ok P (correct_jvm_state \)" and Exec: "P \ s -\ttas\\<^bsub>jvm\<^esub>* s'" and "\ (\t ta s''. P \ s' -t\ta\\<^bsub>jvm\<^esub> s'')" and ts't: "thr s' t = \((xcp, frs), ln)\" shows "frs \ [] \ ln \ no_wait_locks \ t \ multithreaded_base.deadlocked JVM_final (mexec P) s'" and "\ \ t: (xcp, shr s', frs) \" proof - interpret progress JVM_final "mexec P" convert_RA "exec_mthr.wset_Suspend_ok P (correct_jvm_state \)" by(rule exec_wf_progress) fact+ from Exec s have wfs': "s' \ exec_mthr.wset_Suspend_ok P (correct_jvm_state \)" unfolding exec_mthr.RedT_def by(blast intro: invariant3p_rtrancl3p exec_mthr.invariant3p_wset_Suspend_ok invariant3p_correct_jvm_state_mexecT[OF wf]) with ts't show cst: "\ \ t: (xcp, shr s', frs) \" by(auto dest: ts_okD exec_mthr.wset_Suspend_okD1 simp add: correct_jvm_state_def) assume nfin: "frs \ [] \ ln \ no_wait_locks" from nfin \thr s' t = \((xcp, frs), ln)\\ have "exec_mthr.not_final_thread s' t" by(auto simp: exec_mthr.not_final_thread_iff) from \\ (\t ta s''. P \ s' -t\ta\\<^bsub>jvm\<^esub> s'')\ show "t \ exec_mthr.deadlocked P s'" proof(rule contrapos_np) assume "t \ exec_mthr.deadlocked P s'" with \exec_mthr.not_final_thread s' t\ have "\ exec_mthr.deadlocked' P s'" by(auto simp add: exec_mthr.deadlocked'_def) hence "\ exec_mthr.deadlock P s'" unfolding exec_mthr.deadlock_eq_deadlocked' . thus "\t ta s''. P \ s' -t\ta\\<^bsub>jvm\<^esub> s''" by(rule redT_progress[OF wfs']) qed qed lemma start_mexec_mexecd_commute: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and start: "wf_start_state P C M vs" shows "P \ JVM_start_state P C M vs -\ttas\\<^bsub>jvmd\<^esub>* s \ P \ JVM_start_state P C M vs -\ttas\\<^bsub>jvm\<^esub>* s" using correct_jvm_state_initial[OF assms] by(clarsimp simp add: correct_jvm_state_def)(rule mExecT_eq_mExecdT[symmetric, OF wf]) theorem mRtrancl_eq_mRtrancld: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and ct: "correct_state_ts \ (thr s) (shr s)" shows "exec_mthr.mthr.Rtrancl3p P s ttas \ execd_mthr.mthr.Rtrancl3p P s ttas" (is "?lhs \ ?rhs") proof show ?lhs if ?rhs using that ct proof(coinduction arbitrary: s ttas) case Rtrancl3p interpret lifting_wf "JVM_final" "mexecd P" convert_RA "\t (xcp, frs) h. \ \ t: (xcp, h, frs) \" using wf by(rule lifting_wf_correct_state_d) from Rtrancl3p(1) show ?case proof cases case stop: Rtrancl3p_stop then show ?thesis using mexecT_eq_mexecdT[OF wf Rtrancl3p(2)] by clarsimp next case (Rtrancl3p_into_Rtrancl3p s' ttas' tta) then show ?thesis using mexecT_eq_mexecdT[OF wf Rtrancl3p(2)] Rtrancl3p(2) by(cases tta; cases s')(fastforce simp add: split_paired_Ex dest: redT_preserves) qed qed show ?rhs if ?lhs using that ct proof(coinduction arbitrary: s ttas) case Rtrancl3p interpret lifting_wf "JVM_final" "mexec P" convert_RA "\t (xcp, frs) h. \ \ t: (xcp, h, frs) \" using wf by(rule lifting_wf_correct_state) from Rtrancl3p(1) show ?case proof cases case stop: Rtrancl3p_stop then show ?thesis using mexecT_eq_mexecdT[OF wf Rtrancl3p(2)] by clarsimp next case (Rtrancl3p_into_Rtrancl3p s' ttas' tta) then show ?thesis using mexecT_eq_mexecdT[OF wf Rtrancl3p(2)] Rtrancl3p(2) by(cases tta; cases s')(fastforce simp add: split_paired_Ex dest: redT_preserves) qed qed qed lemma start_mRtrancl_mRtrancld_commute: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and start: "wf_start_state P C M vs" shows "exec_mthr.mthr.Rtrancl3p P (JVM_start_state P C M vs) ttas \ execd_mthr.mthr.Rtrancl3p P (JVM_start_state P C M vs) ttas" using correct_jvm_state_initial[OF assms] by(clarsimp simp add: correct_jvm_state_def mRtrancl_eq_mRtrancld[OF wf]) end subsection \Determinism\ context JVM_heap_conf begin lemma exec_instr_deterministic: assumes wf: "wf_prog wf_md P" and det: "deterministic_heap_ops" and exec1: "(ta', \') \ exec_instr i P t (shr s) stk loc C M pc frs" and exec2: "(ta'', \'') \ exec_instr i P t (shr s) stk loc C M pc frs" and check: "check_instr i P (shr s) stk loc C M pc frs" and aok1: "final_thread.actions_ok final s t ta'" and aok2: "final_thread.actions_ok final s t ta''" and tconf: "P,shr s \ t \t" shows "ta' = ta'' \ \' = \''" using exec1 exec2 aok1 aok2 proof(cases i) case (Invoke M' n) { fix T ta''' ta'''' va' va'' h' h'' assume T: "typeof_addr (shr s) (the_Addr (stk ! n)) = \T\" and "method": "snd (snd (snd (method P (class_type_of T) M'))) = None" "P \ class_type_of T has M'" and params: "P,shr s \ rev (take n stk) [:\] fst (snd (method P (class_type_of T) M'))" and red1: "(ta''', va', h') \ red_external_aggr P t (the_Addr (stk ! n)) M' (rev (take n stk)) (shr s)" and red2: "(ta'''', va'', h'') \ red_external_aggr P t (the_Addr (stk ! n)) M' (rev (take n stk)) (shr s)" and ta': "ta' = extTA2JVM P ta'''" and ta'': "ta'' = extTA2JVM P ta''''" from T "method" params obtain T' where "P,shr s \ the_Addr (stk ! n)\M'(rev (take n stk)) : T'" by(fastforce simp add: has_method_def confs_conv_map external_WT'_iff) hence "P,t \ \the_Addr (stk ! n)\M'(rev (take n stk)), shr s\ -ta'''\ext \va', h'\" and "P,t \ \the_Addr (stk ! n)\M'(rev (take n stk)), shr s\ -ta''''\ext \va'', h''\" using red1 red2 tconf by-(rule WT_red_external_aggr_imp_red_external[OF wf], assumption+)+ moreover from aok1 aok2 ta' ta'' have "final_thread.actions_ok final s t ta'''" and "final_thread.actions_ok final s t ta''''" by(auto simp add: final_thread.actions_ok_iff) ultimately have "ta''' = ta'''' \ va' = va'' \ h' = h''" by(rule red_external_deterministic[OF det]) } with assms Invoke show ?thesis by(clarsimp simp add: split_beta split: if_split_asm) blast next case MExit { assume "final_thread.actions_ok final s t \UnlockFail\the_Addr (hd stk)\" and "final_thread.actions_ok final s t \Unlock\the_Addr (hd stk), SyncUnlock (the_Addr (hd stk))\" hence False by(auto simp add: final_thread.actions_ok_iff lock_ok_las_def finfun_upd_apply elim!: allE[where x="the_Addr (hd stk)"]) } with assms MExit show ?thesis by(auto split: if_split_asm) qed(auto simp add: split_beta split: if_split_asm dest: deterministic_heap_ops_readD[OF det] deterministic_heap_ops_writeD[OF det] deterministic_heap_ops_allocateD[OF det]) lemma exec_1_deterministic: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and det: "deterministic_heap_ops" and exec1: "P,t \ (xcp, shr s, frs) -ta'-jvm\ \'" and exec2: "P,t \ (xcp, shr s, frs) -ta''-jvm\ \''" and aok1: "final_thread.actions_ok final s t ta'" and aok2: "final_thread.actions_ok final s t ta''" and conf: "\ \ t:(xcp, shr s, frs) \" shows "ta' = ta'' \ \' = \''" proof - from wf obtain wf_md where wf': "wf_prog wf_md P" by(blast dest: wt_jvm_progD) from conf have tconf: "P,shr s \ t \t" by(simp add: correct_state_def) from exec1 conf have "P,t \ Normal (xcp, shr s, frs) -ta'-jvmd\ Normal \'" by(simp add: welltyped_commute[OF wf]) hence "check P (xcp, shr s, frs)" by(rule jvmd_NormalE) with exec1 exec2 aok1 aok2 tconf show ?thesis by(cases xcp)(case_tac [!] frs, auto elim!: exec_1.cases dest: exec_instr_deterministic[OF wf' det] simp add: check_def split_beta) qed end context JVM_conf_read begin lemma invariant3p_correct_state_ts: assumes "wf_jvm_prog\<^bsub>\\<^esub> P" shows "invariant3p (mexecT P) {s. correct_state_ts \ (thr s) (shr s)}" using assms by(rule lifting_wf.invariant3p_ts_ok[OF lifting_wf_correct_state]) lemma mexec_deterministic: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and det: "deterministic_heap_ops" shows "exec_mthr.deterministic P {s. correct_state_ts \ (thr s) (shr s)}" proof(rule exec_mthr.determisticI) fix s t x ta' x' m' ta'' x'' m'' assume tst: "thr s t = \(x, no_wait_locks)\" and red: "mexec P t (x, shr s) ta' (x', m')" "mexec P t (x, shr s) ta'' (x'', m'')" and aok: "exec_mthr.actions_ok s t ta'" "exec_mthr.actions_ok s t ta''" and correct [simplified]: "s \ {s. correct_state_ts \ (thr s) (shr s)}" moreover obtain xcp frs where [simp]: "x = (xcp, frs)" by(cases x) moreover obtain xcp' frs' where [simp]: "x' = (xcp', frs')" by(cases x') moreover obtain xcp'' frs'' where [simp]: "x'' = (xcp'', frs'')" by(cases x'') ultimately have exec1: "P,t \ (xcp, shr s, frs) -ta'-jvm\ (xcp', m', frs')" and exec1: "P,t \ (xcp, shr s, frs) -ta''-jvm\ (xcp'', m'', frs'')" by simp_all moreover note aok moreover from correct tst have "\ \ t:(xcp, shr s, frs)\" by(auto dest: ts_okD) ultimately have "ta' = ta'' \ (xcp', m', frs') = (xcp'', m'', frs'')" by(rule exec_1_deterministic[OF wf det]) thus "ta' = ta'' \ x' = x'' \ m' = m''" by simp qed(rule invariant3p_correct_state_ts[OF wf]) end end diff --git a/thys/JinjaThreads/BV/JVMDeadlocked.thy b/thys/JinjaThreads/BV/JVMDeadlocked.thy --- a/thys/JinjaThreads/BV/JVMDeadlocked.thy +++ b/thys/JinjaThreads/BV/JVMDeadlocked.thy @@ -1,342 +1,342 @@ (* Title: JinjaThreads/BV/JVMDeadlocked.thy Author: Andreas Lochbihler *) section \Preservation of deadlock for the JVMs\ theory JVMDeadlocked imports BVProgressThreaded begin context JVM_progress begin lemma must_sync_preserved_d: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and ml: "execd_mthr.must_sync P t (xcp, frs) h" and hext: "hext h h'" and hconf': "hconf h'" and cs: "\ \ t: (xcp, h, frs) \" shows "execd_mthr.must_sync P t (xcp, frs) h'" proof(rule execd_mthr.must_syncI) from ml obtain ta xcp' frs' m' where red: "P,t \ Normal (xcp, h, frs) -ta-jvmd\ Normal (xcp', m', frs')" by(auto elim: execd_mthr.must_syncE) then obtain f Frs where check: "check P (xcp, h, frs)" and exec: "(ta, xcp', m', frs') \ exec P t (xcp, h, frs)" and [simp]: "frs = f # Frs" by(auto elim: jvmd_NormalE) from cs hext hconf' have cs': "\ \ t: (xcp, h', frs) \" by(rule correct_state_hext_mono) then obtain ta \' where exec: "P,t \ (xcp, h', frs) -ta-jvm\ \'" by(auto dest: progress[OF wf]) hence "P,t \ Normal (xcp, h', frs) -ta-jvmd\ Normal \'" unfolding welltyped_commute[OF wf cs'] . moreover from exec have "\s. exec_mthr.actions_ok s t ta" by(rule exec_ta_satisfiable) ultimately show "\ta x' m' s. mexecd P t ((xcp, frs), h') ta (x', m') \ exec_mthr.actions_ok s t ta" by(cases \')(fastforce simp del: split_paired_Ex) qed lemma can_sync_devreserp_d: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and cl': "execd_mthr.can_sync P t (xcp, frs) h' L" and cs: "\ \ t: (xcp, h, frs) \" and hext: "hext h h'" and hconf': "hconf h'" shows "\L'\L. execd_mthr.can_sync P t (xcp, frs) h L'" proof - from cl' obtain ta xcp' frs' m' where red: "P,t \ Normal (xcp, h', frs) -ta-jvmd\ Normal (xcp', m', frs')" and L: "L = collect_locks \ta\\<^bsub>l\<^esub> <+> collect_cond_actions \ta\\<^bsub>c\<^esub> <+> collect_interrupts \ta\\<^bsub>i\<^esub>" by -(erule execd_mthr.can_syncE, auto) then obtain f Frs where check: "check P (xcp, h', frs)" and exec: "(ta, xcp', m', frs') \ exec P t (xcp, h', frs)" and [simp]: "frs = f # Frs" by(auto elim: jvmd_NormalE simp add: finfun_upd_apply) obtain stk loc C M pc where [simp]: "f = (stk, loc, C, M, pc)" by (cases f, blast) from cs obtain ST LT Ts T mxs mxl ins xt where hconf: "hconf h" and tconf: "P,h \ t \t" and meth: "P \ C sees M:Ts\T = \(mxs, mxl, ins, xt)\ in C" and \: "\ C M ! pc = Some (ST,LT)" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T Frs" by (fastforce simp add: correct_state_def dest: sees_method_fun) from cs have "exec P t (xcp, h, f # Frs) \ {}" by(auto dest!: progress[OF wf] simp add: exec_1_iff) with no_type_error[OF wf cs] have check': "check P (xcp, h, frs)" by(auto simp add: exec_d_def split: if_split_asm) from wf obtain wfmd where wfp: "wf_prog wfmd P" by(auto dest: wt_jvm_progD) from tconf hext have tconf': "P,h' \ t \t" by(rule tconf_hext_mono) show ?thesis proof(cases xcp) case [simp]: (Some a) with exec have [simp]: "m' = h'" by(auto) from \\ \ t: (xcp, h, frs) \\ obtain D where D: "typeof_addr h a = \Class_type D\" by(auto simp add: correct_state_def) with hext have "cname_of h a = cname_of h' a" by(auto dest: hext_objD simp add: cname_of_def) with exec have "(ta, xcp', h, frs') \ exec P t (xcp, h, frs)" by auto moreover from check D hext have "check P (xcp, h, frs)" by(auto simp add: check_def check_xcpt_def dest: hext_objD) ultimately have "P,t \ Normal (xcp, h, frs) -ta-jvmd\ Normal (xcp', h, frs')" by -(rule exec_1_d_NormalI, simp only: exec_d_def if_True) with L have "execd_mthr.can_sync P t (xcp, frs) h L" by(auto intro: execd_mthr.can_syncI) thus ?thesis by auto next case [simp]: None note [simp] = defs1 list_all2_Cons2 from frame have ST: "P,h \ stk [:\] ST" and LT: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < length ins" by simp_all from wf meth pc have wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" by(rule wt_jvm_prog_impl_wt_instr) from \\ \ t: (xcp, h, frs) \\ have "\ta \'. P,t \ (xcp, h, f # Frs) -ta-jvm\ \'" by(auto dest: progress[OF wf] simp del: correct_state_def split_paired_Ex) with exec meth have "\ta' \'. (ta', \') \ exec P t (xcp, h, frs) \ 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>" proof(cases "ins ! pc") case (Invoke M' n) show ?thesis proof(cases "stk ! n = Null") case True with Invoke exec meth show ?thesis by simp next case False with check meth obtain a where a: "stk ! n = Addr a" and n: "n < length stk" by(auto simp add: check_def is_Ref_def Invoke) from frame have stk: "P,h \ stk [:\] ST" by(auto simp add: conf_f_def) hence "P,h \ stk ! n :\ ST ! n" using n by(rule list_all2_nthD) with a obtain ao Ta where Ta: "typeof_addr h a = \Ta\" by(auto simp add: conf_def) from hext Ta have Ta': "typeof_addr h' a = \Ta\" by(rule typeof_addr_hext_mono) with check a meth Invoke False obtain D Ts' T' meth D' where C: "D = class_type_of Ta" and sees': "P \ D sees M':Ts'\T' = meth in D'" and params: "P,h' \ rev (take n stk) [:\] Ts'" by(auto simp add: check_def has_method_def) show ?thesis proof(cases "meth") case Some with exec meth a Ta Ta' Invoke n sees' C show ?thesis by(simp add: split_beta) next case None with exec meth a Ta Ta' Invoke n sees' C obtain ta' va h'' where ta': "ta = extTA2JVM P ta'" and va: "(xcp', m', frs') = extRet2JVM n h'' stk loc C M pc Frs va" and exec': "(ta', va, h'') \ red_external_aggr P t a M' (rev (take n stk)) h'" by(fastforce) from va have [simp]: "h'' = m'" by(cases va) simp_all note Ta moreover from None sees' wfp have "D'\M'(Ts') :: T'" by(auto intro: sees_wf_native) with C sees' params Ta' None have "P,h' \ a\M'(rev (take n stk)) : T'" by(auto simp add: external_WT'_iff confs_conv_map) with wfp exec' tconf' have red: "P,t \ \a\M'(rev (take n stk)), h'\ -ta'\ext \va, m'\" by(simp add: WT_red_external_list_conv) from stk have "P,h \ take n stk [:\] take n ST" by(rule list_all2_takeI) then obtain Ts where "map typeof\<^bsub>h\<^esub> (take n stk) = map Some Ts" by(auto simp add: confs_conv_map) hence "map typeof\<^bsub>h\<^esub> (rev (take n stk)) = map Some (rev Ts)" by(simp only: rev_map[symmetric]) moreover hence "map typeof\<^bsub>h'\<^esub> (rev (take n stk)) = map Some (rev Ts)" using hext by(rule map_typeof_hext_mono) with \P,h' \ a\M'(rev (take n stk)) : T'\ \D'\M'(Ts') :: T'\ sees' C Ta' Ta have "P \ rev Ts [\] Ts'" by cases (auto dest: sees_method_fun) ultimately have "P,h \ a\M'(rev (take n stk)) : T'" using Ta C sees' params None \D'\M'(Ts') :: T'\ by(auto simp add: external_WT'_iff confs_conv_map) from red_external_wt_hconf_hext[OF wfp red hext this tconf hconf] obtain ta'' va' h''' where "P,t \ \a\M'(rev (take n stk)),h\ -ta''\ext \va',h'''\" and ta'': "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>" by auto with None a Ta Invoke meth Ta' n C sees' have "(extTA2JVM P ta'', extRet2JVM n h''' stk loc C M pc Frs va') \ exec P t (xcp, h, frs)" by(force intro: red_external_imp_red_external_aggr simp del: split_paired_Ex) with ta'' ta' show ?thesis by(fastforce simp del: split_paired_Ex) qed qed qed(auto 4 4 split: if_split_asm simp add: split_beta ta_upd_simps exec_1_iff intro: rev_image_eqI simp del: split_paired_Ex) with check' have "\ta' \'. P,t \ Normal (xcp, h, frs) -ta'-jvmd\ Normal \' \ 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>" apply clarify apply(rule exI conjI)+ apply(rule exec_1_d.exec_1_d_NormalI, auto simp add: exec_d_def) done with L show ?thesis apply - apply(erule exE conjE|rule exI conjI)+ prefer 2 apply(rule_tac x'="(fst \', snd (snd \'))" and m'="fst (snd \')" in execd_mthr.can_syncI) apply auto done qed qed end context JVM_typesafe begin lemma execd_preserve_deadlocked: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "preserve_deadlocked JVM_final (mexecd P) convert_RA (correct_jvm_state \)" proof(unfold_locales) show "invariant3p (mexecdT P) (correct_jvm_state \)" by(rule invariant3p_correct_jvm_state_mexecdT[OF wf]) next fix s t' ta' s' t x ln assume s: "s \ correct_jvm_state \" and red: "P \ s -t'\ta'\\<^bsub>jvmd\<^esub> s'" and tst: "thr s t = \(x, ln)\" and "execd_mthr.must_sync P t x (shr s)" moreover obtain xcp frs where x [simp]: "x = (xcp, frs)" by(cases x, auto) ultimately have ml: "execd_mthr.must_sync P t (xcp, frs) (shr s)" by simp moreover from s have cs': "correct_state_ts \ (thr s) (shr s)" by(simp add: correct_jvm_state_def) with tst have "\ \ t: (xcp, shr s, frs) \" by(auto dest: ts_okD) moreover from red have "hext (shr s) (shr s')" by(rule execd_hext) moreover from wf red cs' have "correct_state_ts \ (thr s') (shr s')" by(rule lifting_wf.redT_preserves[OF lifting_wf_correct_state_d]) from red tst have "thr s' t \ None" by(cases s)(cases s', rule notI, auto dest: execd_mthr.redT_thread_not_disappear) with \correct_state_ts \ (thr s') (shr s')\ have "hconf (shr s')" by(auto dest: ts_okD simp add: correct_state_def) ultimately have "execd_mthr.must_sync P t (xcp, frs) (shr s')" by-(rule must_sync_preserved_d[OF wf]) thus "execd_mthr.must_sync P t x (shr s')" by simp next fix s t' ta' s' t x ln L assume s: "s \ correct_jvm_state \" and red: "P \ s -t'\ta'\\<^bsub>jvmd\<^esub> s'" and tst: "thr s t = \(x, ln)\" and "execd_mthr.can_sync P t x (shr s') L" moreover obtain xcp frs where x [simp]: "x = (xcp, frs)" by(cases x, auto) ultimately have ml: "execd_mthr.can_sync P t (xcp, frs) (shr s') L" by simp moreover from s have cs': "correct_state_ts \ (thr s) (shr s)" by(simp add: correct_jvm_state_def) with tst have "\ \ t: (xcp, shr s, frs) \" by(auto dest: ts_okD) moreover from red have "hext (shr s) (shr s')" by(rule execd_hext) moreover from red tst have "thr s' t \ None" by(cases s)(cases s', rule notI, auto dest: execd_mthr.redT_thread_not_disappear) from red cs' have "correct_state_ts \ (thr s') (shr s')" by(rule lifting_wf.redT_preserves[OF lifting_wf_correct_state_d[OF wf]]) with \thr s' t \ None\ have "hconf (shr s')" by(auto dest: ts_okD simp add: correct_state_def) ultimately have "\L' \ L. execd_mthr.can_sync P t (xcp, frs) (shr s) L'" by-(rule can_sync_devreserp_d[OF wf]) thus "\L' \ L. execd_mthr.can_sync P t x (shr s) L'" by simp qed end text \and now everything again for the aggresive VM\ context JVM_heap_conf_base' begin lemma must_lock_d_eq_must_lock: "\ wf_jvm_prog\<^bsub>\\<^esub> P; \ \ t: (xcp, h, frs) \ \ \ execd_mthr.must_sync P t (xcp, frs) h = exec_mthr.must_sync P t (xcp, frs) h" apply(rule iffI) apply(rule exec_mthr.must_syncI) apply(erule execd_mthr.must_syncE) apply(simp only: mexec_eq_mexecd) apply(blast) apply(rule execd_mthr.must_syncI) apply(erule exec_mthr.must_syncE) apply(simp only: mexec_eq_mexecd[symmetric]) apply(blast) done lemma can_lock_d_eq_can_lock: "\ wf_jvm_prog\<^bsub>\\<^esub> P; \ \ t: (xcp, h, frs) \ \ \ execd_mthr.can_sync P t (xcp, frs) h L = exec_mthr.can_sync P t (xcp, frs) h L" apply(rule iffI) apply(erule execd_mthr.can_syncE) apply(rule exec_mthr.can_syncI) apply(simp only: mexec_eq_mexecd) apply(assumption)+ apply(erule exec_mthr.can_syncE) apply(rule execd_mthr.can_syncI) by(simp only: mexec_eq_mexecd) end context JVM_typesafe begin lemma exec_preserve_deadlocked: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "preserve_deadlocked JVM_final (mexec P) convert_RA (correct_jvm_state \)" proof - interpret preserve_deadlocked JVM_final "mexecd P" convert_RA "correct_jvm_state \" by(rule execd_preserve_deadlocked) fact+ { fix s t' ta' s' t x assume s: "s \ correct_jvm_state \" and red: "P \ s -t'\ta'\\<^bsub>jvm\<^esub> s'" and tst: "thr s t = \(x, no_wait_locks)\" obtain xcp frs where x [simp]: "x = (xcp, frs)" by(cases x, auto) from s have css: "correct_state_ts \ (thr s) (shr s)" by(simp add: correct_jvm_state_def) with red have redd: "P \ s -t'\ta'\\<^bsub>jvmd\<^esub> s'" by(simp add: mexecT_eq_mexecdT[OF wf]) from css tst have cst: "\ \ t: (xcp, shr s, frs) \" by(auto dest: ts_okD) from redd have cst': "\ \ t: (xcp, shr s', frs) \" proof(cases rule: execd_mthr.redT_elims) case acquire with cst show ?thesis by simp next case (normal X X' M' ws') obtain XCP FRS where X [simp]: "X = (XCP, FRS)" by(cases X, auto) obtain XCP' FRS' where X' [simp]: "X' = (XCP', FRS')" by(cases X', auto) from \mexecd P t' (X, shr s) ta' (X', M')\ have "P,t' \ Normal (XCP, shr s, FRS) -ta'-jvmd\ Normal (XCP', M', FRS')" by simp moreover from \thr s t' = \(X, no_wait_locks)\\ css have "\ \ t': (XCP, shr s, FRS) \" by(auto dest: ts_okD) ultimately have "\ \ t': (XCP, M', FRS) \" by -(rule correct_state_heap_change[OF wf]) moreover from lifting_wf.redT_updTs_preserves[OF lifting_wf_correct_state_d[OF wf] css, OF \mexecd P t' (X, shr s) ta' (X', M')\ \thr s t' = \(X, no_wait_locks)\\, of no_wait_locks] \thread_oks (thr s) \ta'\\<^bsub>t\<^esub>\ - have "correct_state_ts \ (redT_updTs (thr s) \ta'\\<^bsub>t\<^esub>(t' \ (X', no_wait_locks))) M'" by simp + have "correct_state_ts \ ((redT_updTs (thr s) \ta'\\<^bsub>t\<^esub>)(t' \ (X', no_wait_locks))) M'" by simp ultimately have "correct_state_ts \ (redT_updTs (thr s) \ta'\\<^bsub>t\<^esub>) M'" using \thr s t' = \(X, no_wait_locks)\\ \thread_oks (thr s) \ta'\\<^bsub>t\<^esub>\ apply(auto intro!: ts_okI dest: ts_okD) apply(case_tac "t=t'") apply(fastforce dest: redT_updTs_Some) apply(drule_tac t=t in ts_okD, fastforce+) done hence "correct_state_ts \ (redT_updTs (thr s) \ta'\\<^bsub>t\<^esub>) (shr s')" - using \s' = (redT_updLs (locks s) t' \ta'\\<^bsub>l\<^esub>, (redT_updTs (thr s) \ta'\\<^bsub>t\<^esub>(t' \ (X', redT_updLns (locks s) t' no_wait_locks \ta'\\<^bsub>l\<^esub>)), M'), ws', redT_updIs (interrupts s) \ta'\\<^bsub>i\<^esub>)\ + using \s' = (redT_updLs (locks s) t' \ta'\\<^bsub>l\<^esub>, ((redT_updTs (thr s) \ta'\\<^bsub>t\<^esub>)(t' \ (X', redT_updLns (locks s) t' no_wait_locks \ta'\\<^bsub>l\<^esub>)), M'), ws', redT_updIs (interrupts s) \ta'\\<^bsub>i\<^esub>)\ by simp moreover from tst \thread_oks (thr s) \ta'\\<^bsub>t\<^esub>\ have "redT_updTs (thr s) \ta'\\<^bsub>t\<^esub> t = \(x, no_wait_locks)\" by(auto intro: redT_updTs_Some) ultimately show ?thesis by(auto dest: ts_okD) qed { assume "exec_mthr.must_sync P t x (shr s)" hence ml: "exec_mthr.must_sync P t (xcp, frs) (shr s)" by simp with cst have "execd_mthr.must_sync P t (xcp, frs) (shr s)" by(auto dest: must_lock_d_eq_must_lock[OF wf]) with s redd tst have "execd_mthr.must_sync P t x (shr s')" unfolding x by(rule can_lock_preserved) with cst' have "exec_mthr.must_sync P t x (shr s')" by(auto dest: must_lock_d_eq_must_lock[OF wf]) } note ml = this { fix L assume "exec_mthr.can_sync P t x (shr s') L" hence cl: "exec_mthr.can_sync P t (xcp, frs) (shr s') L" by simp with cst' have "execd_mthr.can_sync P t (xcp, frs) (shr s') L" by(auto dest: can_lock_d_eq_can_lock[OF wf]) with s redd tst have "\L' \ L. execd_mthr.can_sync P t x (shr s) L'" unfolding x by(rule can_lock_devreserp) then obtain L' where "execd_mthr.can_sync P t x (shr s) L'" and L': "L'\ L" by blast with cst have "exec_mthr.can_sync P t x (shr s) L'" by(auto dest: can_lock_d_eq_can_lock[OF wf]) with L' have "\L' \ L. exec_mthr.can_sync P t x (shr s) L'" by(blast) } note this ml } moreover have "invariant3p (mexecT P) (correct_jvm_state \)" by(rule invariant3p_correct_jvm_state_mexecT[OF wf]) ultimately show ?thesis by(unfold_locales) qed end end diff --git a/thys/JinjaThreads/Compiler/Correctness1Threaded.thy b/thys/JinjaThreads/Compiler/Correctness1Threaded.thy --- a/thys/JinjaThreads/Compiler/Correctness1Threaded.thy +++ b/thys/JinjaThreads/Compiler/Correctness1Threaded.thy @@ -1,746 +1,746 @@ (* Title: JinjaThreads/Compiler/Correctness1Threaded.thy Author: Andreas Lochbihler *) section \Unlocking a sync block never fails\ theory Correctness1Threaded imports J0J1Bisim "../Framework/FWInitFinLift" begin definition lock_oks1 :: "('addr,'thread_id) locks \ ('addr,'thread_id,(('a,'b,'addr) exp \ 'c) \ (('a,'b,'addr) exp \ 'c) list) thread_info \ bool" where "\ln. lock_oks1 ls ts \ \t. (case (ts t) of None \ (\l. has_locks (ls $ l) t = 0) | \((ex, exs), ln)\ \ (\l. has_locks (ls $ l) t + ln $ l = expr_lockss (map fst (ex # exs)) l))" primrec el_loc_ok :: "'addr expr1 \ 'addr locals1 \ bool" and els_loc_ok :: "'addr expr1 list \ 'addr locals1 \ bool" where "el_loc_ok (new C) xs \ True" | "el_loc_ok (newA T\e\) xs \ el_loc_ok e xs" | "el_loc_ok (Cast T e) xs \ el_loc_ok e xs" | "el_loc_ok (e instanceof T) xs \ el_loc_ok e xs" | "el_loc_ok (e\bop\e') xs \ el_loc_ok e xs \ el_loc_ok e' xs" | "el_loc_ok (Var V) xs \ True" | "el_loc_ok (Val v) xs \ True" | "el_loc_ok (V := e) xs \ el_loc_ok e xs" | "el_loc_ok (a\i\) xs \ el_loc_ok a xs \ el_loc_ok i xs" | "el_loc_ok (a\i\ := e) xs \ el_loc_ok a xs \ el_loc_ok i xs \ el_loc_ok e xs" | "el_loc_ok (a\length) xs \ el_loc_ok a xs" | "el_loc_ok (e\F{D}) xs \ el_loc_ok e xs" | "el_loc_ok (e\F{D} := e') xs \ el_loc_ok e xs \ el_loc_ok e' xs" | "el_loc_ok (e\compareAndSwap(D\F, e', e'')) xs \ el_loc_ok e xs \ el_loc_ok e' xs \ el_loc_ok e'' xs" | "el_loc_ok (e\M(ps)) xs \ el_loc_ok e xs \ els_loc_ok ps xs" | "el_loc_ok {V:T=vo; e} xs \ (case vo of None \ el_loc_ok e xs | \v\ \ el_loc_ok e (xs[V := v]))" | "el_loc_ok (sync\<^bsub>V\<^esub>(e) e') xs \ el_loc_ok e xs \ el_loc_ok e' xs \ unmod e' V" | "el_loc_ok (insync\<^bsub>V\<^esub>(a) e) xs \ xs ! V = Addr a \ el_loc_ok e xs \ unmod e V" | "el_loc_ok (e;;e') xs \ el_loc_ok e xs \ el_loc_ok e' xs" | "el_loc_ok (if (b) e else e') xs \ el_loc_ok b xs \ el_loc_ok e xs \ el_loc_ok e' xs" | "el_loc_ok (while (b) c) xs \ el_loc_ok b xs \ el_loc_ok c xs" | "el_loc_ok (throw e) xs \ el_loc_ok e xs" | "el_loc_ok (try e catch(C V) e') xs \ el_loc_ok e xs \ el_loc_ok e' xs" | "els_loc_ok [] xs \ True" | "els_loc_ok (e # es) xs \ el_loc_ok e xs \ els_loc_ok es xs" lemma el_loc_okI: "\ \ contains_insync e; syncvars e; \ e n \ \ el_loc_ok e xs" and els_loc_okI: "\ \ contains_insyncs es; syncvarss es; \s es n \ \ els_loc_ok es xs" by(induct e and es arbitrary: xs n and xs n rule: el_loc_ok.induct els_loc_ok.induct)(auto intro: fv_B_unmod) lemma el_loc_ok_compE1: "\ \ contains_insync e; fv e \ set Vs \ \ el_loc_ok (compE1 Vs e) xs" and els_loc_ok_compEs1: "\ \ contains_insyncs es; fvs es \ set Vs \ \ els_loc_ok (compEs1 Vs es) xs" by(auto intro: el_loc_okI els_loc_okI syncvars_compE1 syncvarss_compEs1 \ \s simp del: compEs1_conv_map) lemma shows el_loc_ok_not_contains_insync_local_change: "\ \ contains_insync e; el_loc_ok e xs \ \ el_loc_ok e xs'" and els_loc_ok_not_contains_insyncs_local_change: "\ \ contains_insyncs es; els_loc_ok es xs \ \ els_loc_ok es xs'" by(induct e and es arbitrary: xs xs' and xs xs' rule: el_loc_ok.induct els_loc_ok.induct)(fastforce)+ lemma el_loc_ok_update: "\ \ e n; V < n \ \ el_loc_ok e (xs[V := v]) = el_loc_ok e xs" and els_loc_ok_update: "\ \s es n; V < n \ \ els_loc_ok es (xs[V := v]) = els_loc_ok es xs" apply(induct e and es arbitrary: n xs and n xs rule: el_loc_ok.induct els_loc_ok.induct) apply(auto simp add: list_update_swap) done lemma els_loc_ok_map_Val [simp]: "els_loc_ok (map Val vs) xs" by(induct vs) auto lemma els_loc_ok_map_Val_append [simp]: "els_loc_ok (map Val vs @ es) xs = els_loc_ok es xs" by(induct vs) auto lemma el_loc_ok_extRet2J [simp]: "el_loc_ok e xs \ el_loc_ok (extRet2J e va) xs" by(cases va) auto definition el_loc_ok1 :: "((nat, nat, 'addr) exp \ 'addr locals1) \ ((nat, nat, 'addr) exp \ 'addr locals1) list \ bool" where "el_loc_ok1 = (\((e, xs), exs). el_loc_ok e xs \ sync_ok e \ (\(e,xs)\set exs. el_loc_ok e xs \ sync_ok e))" lemma el_loc_ok1_simps: "el_loc_ok1 ((e, xs), exs) = (el_loc_ok e xs \ sync_ok e \ (\(e,xs)\set exs. el_loc_ok e xs \ sync_ok e))" by(simp add: el_loc_ok1_def) lemma el_loc_ok_blocks1 [simp]: "el_loc_ok (blocks1 n Ts body) xs = el_loc_ok body xs" by(induct n Ts body rule: blocks1.induct) auto lemma sync_oks_blocks1 [simp]: "sync_ok (blocks1 n Ts e) = sync_ok e" by(induct n Ts e rule: blocks1.induct) auto lemma assumes fin: "final e'" shows el_loc_ok_inline_call: "el_loc_ok e xs \ el_loc_ok (inline_call e' e) xs" and els_loc_ok_inline_calls: "els_loc_ok es xs \ els_loc_ok (inline_calls e' es) xs" apply(induct e and es arbitrary: xs and xs rule: el_loc_ok.induct els_loc_ok.induct) apply(insert fin) apply(auto simp add: unmod_inline_call) done lemma assumes "sync_ok e'" shows sync_ok_inline_call: "sync_ok e \ sync_ok (inline_call e' e)" and sync_oks_inline_calls: "sync_oks es \ sync_oks (inline_calls e' es)" apply(induct e and es rule: sync_ok.induct sync_oks.induct) apply(insert \sync_ok e'\) apply auto done lemma bisim_sync_ok: "bisim Vs e e' xs \ sync_ok e" "bisim Vs e e' xs \ sync_ok e'" and bisims_sync_oks: "bisims Vs es es' xs \ sync_oks es" "bisims Vs es es' xs \ sync_oks es'" apply(induct rule: bisim_bisims.inducts) apply(auto intro: not_contains_insync_sync_ok not_contains_insyncs_sync_oks simp del: compEs1_conv_map) done lemma assumes "final e'" shows expr_locks_inline_call_final: "expr_locks (inline_call e' e) = expr_locks e" and expr_lockss_inline_calls_final: "expr_lockss (inline_calls e' es) = expr_lockss es" apply(induct e and es rule: expr_locks.induct expr_lockss.induct) apply(insert \final e'\) apply(auto simp add: is_vals_conv intro: ext) done lemma lock_oks1I: "\ \t l. ts t = None \ has_locks (ls $ l) t = 0; \t e x exs ln l. ts t = \(((e, x), exs), ln)\ \ has_locks (ls $ l) t + ln $ l= expr_locks e l + expr_lockss (map fst exs) l \ \ lock_oks1 ls ts" apply(fastforce simp add: lock_oks1_def) done lemma lock_oks1E: "\ lock_oks1 ls ts; \t. ts t = None \ (\l. has_locks (ls $ l) t = 0) \ Q; \t e x exs ln. ts t = \(((e, x), exs), ln)\ \ (\l. has_locks (ls $ l) t + ln $ l = expr_locks e l + expr_lockss (map fst exs) l) \ Q \ \ Q" by(fastforce simp add: lock_oks1_def) lemma lock_oks1D1: "\ lock_oks1 ls ts; ts t = None \ \ \l. has_locks (ls $ l) t = 0" apply(simp add: lock_oks1_def) apply(erule_tac x="t" in allE) apply(auto) done lemma lock_oks1D2: "\ln. \ lock_oks1 ls ts; ts t = \(((e, x), exs), ln)\ \ \ \l. has_locks (ls $ l) t + ln $ l = expr_locks e l + expr_lockss (map fst exs) l" apply(fastforce simp add: lock_oks1_def) done lemma lock_oks1_thr_updI: "\ln. \ lock_oks1 ls ts; ts t = \(((e, xs), exs), ln)\; \l. expr_locks e l + expr_lockss (map fst exs) l = expr_locks e' l + expr_lockss (map fst exs') l \ \ lock_oks1 ls (ts(t \ (((e', xs'), exs'), ln)))" by(rule lock_oks1I)(auto split: if_split_asm dest: lock_oks1D2 lock_oks1D1) definition mbisim_Red1'_Red1 :: "(('addr,'thread_id,('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list,'heap,'addr) state, ('addr,'thread_id,('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list,'heap,'addr) state) bisim" where "mbisim_Red1'_Red1 s1 s2 = (s1 = s2 \ lock_oks1 (locks s1) (thr s1) \ ts_ok (\t exexs h. el_loc_ok1 exexs) (thr s1) (shr s1))" lemma sync_ok_blocks: "\ length vs = length pns; length Ts = length pns\ \ sync_ok (blocks pns Ts vs body) = sync_ok body" by(induct pns Ts vs body rule: blocks.induct) auto context J1_heap_base begin lemma red1_True_into_red1_False: "\ True,P,t \1 \e, s\ -ta\ \e', s'\; el_loc_ok e (lcl s) \ \ False,P,t \1 \e, s\ -ta\ \e', s'\ \ (\l. ta = \UnlockFail\l\ \ expr_locks e l > 0)" and reds1_True_into_reds1_False: "\ True,P,t \1 \es, s\ [-ta\] \es', s'\; els_loc_ok es (lcl s) \ \ False,P,t \1 \es, s\ [-ta\] \es', s'\ \ (\l. ta = \UnlockFail\l\ \ expr_lockss es l > 0)" apply(induct rule: red1_reds1.inducts) apply(auto intro: red1_reds1.intros split: if_split_asm) done lemma Red1_True_into_Red1_False: assumes "True,P,t \1 \ex/exs,shr s\ -ta\ \ex'/exs',m'\" and "el_loc_ok1 (ex, exs)" shows "False,P,t \1 \ex/exs,shr s\ -ta\ \ex'/exs',m'\ \ (\l. ta = \UnlockFail\l\ \ expr_lockss (fst ex # map fst exs) l > 0)" using assms by(cases)(auto dest: Red1.intros red1_True_into_red1_False simp add: el_loc_ok1_def ta_upd_simps) lemma shows red1_preserves_el_loc_ok: "\ uf,P,t \1 \e, s\ -ta\ \e', s'\; sync_ok e; el_loc_ok e (lcl s) \ \ el_loc_ok e' (lcl s')" and reds1_preserves_els_loc_ok: "\ uf,P,t \1 \es, s\ [-ta\] \es', s'\; sync_oks es; els_loc_ok es (lcl s) \ \ els_loc_ok es' (lcl s')" proof(induct rule: red1_reds1.inducts) case (Synchronized1Red2 e s ta e' s' V a) from \el_loc_ok (insync\<^bsub>V\<^esub> (a) e) (lcl s)\ have "el_loc_ok e (lcl s)" "unmod e V" "lcl s ! V = Addr a" by auto from \sync_ok (insync\<^bsub>V\<^esub> (a) e)\ have "sync_ok e" by simp hence "el_loc_ok e' (lcl s')" using \el_loc_ok e (lcl s)\ by(rule Synchronized1Red2) moreover from \uf,P,t \1 \e,s\ -ta\ \e',s'\\ \unmod e V\ have "unmod e' V" by(rule red1_unmod_preserved) moreover from red1_preserves_unmod[OF \uf,P,t \1 \e,s\ -ta\ \e',s'\\ \unmod e V\] \lcl s ! V = Addr a\ have "lcl s' ! V = Addr a" by simp ultimately show ?case by auto qed(auto elim: el_loc_ok_not_contains_insync_local_change els_loc_ok_not_contains_insyncs_local_change) lemma red1_preserves_sync_ok: "\ uf,P,t \1 \e, s\ -ta\ \e', s'\; sync_ok e \ \ sync_ok e'" and reds1_preserves_sync_oks: "\ uf,P,t \1 \es, s\ [-ta\] \es', s'\; sync_oks es \ \ sync_oks es'" by(induct rule: red1_reds1.inducts)(auto elim: not_contains_insync_sync_ok) lemma Red1_preserves_el_loc_ok1: assumes wf: "wf_J1_prog P" shows "\ uf,P,t \1 \ex/exs,m\ -ta\ \ex'/exs',m'\; el_loc_ok1 (ex, exs) \ \ el_loc_ok1 (ex', exs')" apply(erule Red1.cases) apply(auto simp add: el_loc_ok1_def dest: red1_preserves_el_loc_ok red1_preserves_sync_ok intro: el_loc_ok_inline_call sync_ok_inline_call) apply(fastforce dest!: sees_wf_mdecl[OF wf] simp add: wf_mdecl_def intro!: el_loc_okI dest: WT1_not_contains_insync intro: not_contains_insync_sync_ok)+ done lemma assumes wf: "wf_J1_prog P" shows red1_el_loc_ok1_new_thread: "\ uf,P,t \1 \e, s\ -ta\ \e', s'\; NewThread t' (C, M, a) h \ set \ta\\<^bsub>t\<^esub> \ \ el_loc_ok1 (({0:Class (fst (method P C M))=None; the (snd (snd (snd (method P C M))))}, xs), [])" and reds1_el_loc_ok1_new_thread: "\ uf,P,t \1 \es, s\ [-ta\] \es', s'\; NewThread t' (C, M, a) h \ set \ta\\<^bsub>t\<^esub> \ \ el_loc_ok1 (({0:Class (fst (method P C M))=None; the (snd (snd (snd (method P C M))))}, xs), [])" proof(induct rule: red1_reds1.inducts) case Red1CallExternal thus ?case apply(auto dest!: red_external_new_thread_sees[OF wf] simp add: el_loc_ok1_simps) apply(auto dest!: sees_wf_mdecl[OF wf] WT1_not_contains_insync simp add: wf_mdecl_def intro!: el_loc_okI not_contains_insync_sync_ok) done qed auto lemma Red1_el_loc_ok1_new_thread: assumes wf: "wf_J1_prog P" shows "\ uf,P,t \1 \ex/exs,m\ -ta\ \ex'/exs',m'\; NewThread t' exexs m' \ set \ta\\<^bsub>t\<^esub> \ \ el_loc_ok1 exexs" by(erule Red1.cases)(fastforce elim: red1_el_loc_ok1_new_thread[OF wf] simp add: ta_upd_simps)+ lemma Red1_el_loc_ok: assumes wf: "wf_J1_prog P" shows "lifting_wf final_expr1 (mred1g uf P) (\t exexs h. el_loc_ok1 exexs)" by(unfold_locales)(auto elim: Red1_preserves_el_loc_ok1[OF wf] Red1_el_loc_ok1_new_thread[OF wf]) lemma mred1_eq_mred1': assumes lok: "lock_oks1 (locks s) (thr s)" and elo: "ts_ok (\t exexs h. el_loc_ok1 exexs) (thr s) (shr s)" and tst: "thr s t = \(exexs, no_wait_locks)\" and aoe: "Red1_mthr.actions_ok s t ta" shows "mred1 P t (exexs, shr s) ta = mred1' P t (exexs, shr s) ta" proof(intro ext iffI) fix xm' assume "mred1 P t (exexs, shr s) ta xm'" moreover obtain ex exs where exexs [simp]: "exexs = (ex, exs)" by(cases exexs) moreover obtain ex' exs' m' where xm' [simp]: "xm' = ((ex', exs'), m')" by(cases xm') auto ultimately have red: "True,P,t \1 \ex/exs,shr s\ -ta\ \ex'/exs',m'\" by simp from elo tst have "el_loc_ok1 (ex, exs)" by(auto dest: ts_okD) from Red1_True_into_Red1_False[OF red this] have "False,P,t \1 \ex/exs,shr s\ -ta\ \ex'/exs',m'\" proof assume "\l. ta = \UnlockFail\l\ \ 0 < expr_lockss (fst ex # map fst exs) l" then obtain l where ta: "ta = \UnlockFail\l\" and el: "expr_lockss (fst ex # map fst exs) l > 0" by blast from aoe have "lock_actions_ok (locks s $ l) t (\ta\\<^bsub>l\<^esub> $ l)" by(auto simp add: lock_ok_las_def) with ta have "has_locks (locks s $ l) t = 0" by simp with lok tst have "expr_lockss (map fst (ex # exs)) l = 0" by(cases ex)(auto 4 6 simp add: lock_oks1_def) with el have False by simp thus ?thesis .. qed thus "mred1' P t (exexs, shr s) ta xm'" by simp next fix xm' assume "mred1' P t (exexs, shr s) ta xm'" thus "mred1 P t (exexs, shr s) ta xm'" by(cases xm')(auto simp add: split_beta intro: Red1_False_into_Red1_True) qed lemma Red1_mthr_eq_Red1_mthr': assumes lok: "lock_oks1 (locks s) (thr s)" and elo: "ts_ok (\t exexs h. el_loc_ok1 exexs) (thr s) (shr s)" shows "Red1_mthr.redT True P s = Red1_mthr.redT False P s" proof(intro ext) fix tta s' show "Red1_mthr.redT True P s tta s' = Red1_mthr.redT False P s tta s'" (is "?lhs = ?rhs") proof assume "?lhs" thus ?rhs proof cases case (redT_normal t x ta x' m') from \mred1 P t (x, shr s) ta (x', m')\ have "mred1' P t (x, shr s) ta (x', m')" unfolding mred1_eq_mred1'[OF lok elo \thr s t = \(x, no_wait_locks)\\ \Red1_mthr.actions_ok s t ta\] . thus ?thesis using redT_normal(3-) unfolding \tta = (t, ta)\ .. next case (redT_acquire t x ln n) from this(2-) show ?thesis unfolding redT_acquire(1) .. qed next assume ?rhs thus ?lhs proof(cases) case (redT_normal t x ta x' m') from \mred1' P t (x, shr s) ta (x', m')\ have "mred1 P t (x, shr s) ta (x', m')" unfolding mred1_eq_mred1'[OF lok elo \thr s t = \(x, no_wait_locks)\\ \Red1_mthr.actions_ok s t ta\] . thus ?thesis using redT_normal(3-) unfolding \tta = (t, ta)\ .. next case (redT_acquire t x ln n) from this(2-) show ?thesis unfolding redT_acquire(1) .. qed qed qed lemma assumes wf: "wf_J1_prog P" shows expr_locks_new_thread1: "\ uf,P,t \1 \e,s\ -TA\ \e',s'\; NewThread t' (ex, exs) h \ set (map (convert_new_thread_action (extNTA2J1 P)) \TA\\<^bsub>t\<^esub>) \ \ expr_lockss (map fst (ex # exs)) = (\ad. 0)" and expr_lockss_new_thread1: "\ uf,P,t \1 \es,s\ [-TA\] \es',s'\; NewThread t' (ex, exs) h \ set (map (convert_new_thread_action (extNTA2J1 P)) \TA\\<^bsub>t\<^esub>) \ \ expr_lockss (map fst (ex # exs)) = (\ad. 0)" proof(induct rule: red1_reds1.inducts) case (Red1CallExternal s a T M vs ta va h' e' s') then obtain C fs ad where subThread: "P \ C \\<^sup>* Thread" and ext: "extNTA2J1 P (C, run, ad) = (ex, exs)" by(fastforce dest: red_external_new_thread_sub_thread) from sub_Thread_sees_run[OF wf subThread] obtain D body where sees: "P \ C sees run: []\Void = \body\ in D" by auto from sees_wf_mdecl[OF wf this] obtain T where "P,[Class D] \1 body :: T" by(auto simp add: wf_mdecl_def) hence "\ contains_insync body" by(rule WT1_not_contains_insync) hence "expr_locks body = (\ad. 0)" by(auto simp add: contains_insync_conv fun_eq_iff) with sees ext show ?case by(auto) qed auto lemma assumes wf: "wf_J1_prog P" shows red1_update_expr_locks: "\ False,P,t \1 \e, s\ -ta\ \e', s'\; sync_ok e; el_loc_ok e (lcl s) \ \ upd_expr_locks (int o expr_locks e) \ta\\<^bsub>l\<^esub> = int o expr_locks e'" and reds1_update_expr_lockss: "\ False,P,t \1 \es, s\ [-ta\] \es', s'\; sync_oks es; els_loc_ok es (lcl s) \ \ upd_expr_locks (int o expr_lockss es) \ta\\<^bsub>l\<^esub> = int o expr_lockss es'" proof - have "\ False,P,t \1 \e, s\ -ta\ \e', s'\; sync_ok e; el_loc_ok e (lcl s) \ \ upd_expr_locks (\ad. 0) \ta\\<^bsub>l\<^esub> = (\ad. (int o expr_locks e') ad - (int o expr_locks e) ad)" and "\ False,P,t \1 \es, s\ [-ta\] \es', s'\; sync_oks es; els_loc_ok es (lcl s) \ \ upd_expr_locks (\ad. 0) \ta\\<^bsub>l\<^esub> = (\ad. (int o expr_lockss es') ad - (int o expr_lockss es) ad)" proof(induct rule: red1_reds1.inducts) case Red1CallExternal thus ?case by(auto simp add: fun_eq_iff contains_insync_conv contains_insyncs_conv finfun_upd_apply elim!: red_external.cases) qed(fastforce simp add: fun_eq_iff contains_insync_conv contains_insyncs_conv finfun_upd_apply)+ hence "\ False,P,t \1 \e, s\ -ta\ \e', s'\; sync_ok e; el_loc_ok e (lcl s) \ \ upd_expr_locks (\ad. 0 + (int \ expr_locks e) ad) \ta\\<^bsub>l\<^esub> = int \ expr_locks e'" and "\ False,P,t \1 \es, s\ [-ta\] \es', s'\; sync_oks es; els_loc_ok es (lcl s) \ \ upd_expr_locks (\ad. 0 + (int \ expr_lockss es) ad) \ta\\<^bsub>l\<^esub> = int \ expr_lockss es'" by(fastforce simp only: upd_expr_locks_add)+ thus "\ False,P,t \1 \e, s\ -ta\ \e', s'\; sync_ok e; el_loc_ok e (lcl s) \ \ upd_expr_locks (int o expr_locks e) \ta\\<^bsub>l\<^esub> = int o expr_locks e'" and "\ False,P,t \1 \es, s\ [-ta\] \es', s'\; sync_oks es; els_loc_ok es (lcl s) \ \ upd_expr_locks (int o expr_lockss es) \ta\\<^bsub>l\<^esub> = int o expr_lockss es'" by(auto simp add: o_def) qed lemma Red1'_preserves_lock_oks: assumes wf: "wf_J1_prog P" and Red: "Red1_mthr.redT False P s1 ta1 s1'" and loks: "lock_oks1 (locks s1) (thr s1)" and sync: "ts_ok (\t exexs h. el_loc_ok1 exexs) (thr s1) (shr s1)" shows "lock_oks1 (locks s1') (thr s1')" using Red proof(cases rule: Red1_mthr.redT.cases) case (redT_normal t x ta x' m') note [simp] = \ta1 = (t, ta)\ obtain ex exs where x: "x = (ex, exs)" by (cases x) obtain ex' exs' where x': "x' = (ex', exs')" by (cases x') note thrst = \thr s1 t = \(x, no_wait_locks)\\ note aoe = \Red1_mthr.actions_ok s1 t ta\ from \mred1' P t (x, shr s1) ta (x', m')\ have red: "False,P,t \1 \ex/exs,shr s1\ -ta\ \ex'/exs',m'\" unfolding x x' by simp_all note s1' = \redT_upd s1 t ta x' m' s1'\ moreover from red have "lock_oks1 (locks s1') (thr s1')" proof cases case (red1Red e x TA e' x') note [simp] = \ex = (e, x)\ \ta = extTA2J1 P TA\ \ex' = (e', x')\ \exs' = exs\ and red = \False,P,t \1 \e,(shr s1, x)\ -TA\ \e',(m', x')\\ { fix t' - assume None: "(redT_updTs (thr s1) (map (convert_new_thread_action (extNTA2J1 P)) \TA\\<^bsub>t\<^esub>)(t \ (((e', x'), exs), redT_updLns (locks s1) t (snd (the (thr s1 t))) \TA\\<^bsub>l\<^esub>))) t' = None" + assume None: "((redT_updTs (thr s1) (map (convert_new_thread_action (extNTA2J1 P)) \TA\\<^bsub>t\<^esub>))(t \ (((e', x'), exs), redT_updLns (locks s1) t (snd (the (thr s1 t))) \TA\\<^bsub>l\<^esub>))) t' = None" { fix l from aoe have "lock_actions_ok (locks s1 $ l) t (\ta\\<^bsub>l\<^esub> $ l)" by(auto simp add: lock_ok_las_def) with None have "has_locks ((redT_updLs (locks s1) t \ta\\<^bsub>l\<^esub>) $ l) t' = has_locks (locks s1 $ l) t'" by(auto split: if_split_asm) also from loks None have "has_locks (locks s1 $ l) t' = 0" unfolding lock_oks1_def by(force split: if_split_asm dest!: redT_updTs_None) finally have "has_locks (upd_locks (locks s1 $ l) t (\TA\\<^bsub>l\<^esub> $ l)) t' = 0" by simp } hence "\l. has_locks (upd_locks (locks s1 $ l) t (\TA\\<^bsub>l\<^esub> $ l)) t' = 0" .. } moreover { fix t' eX eXS LN - assume Some: "(redT_updTs (thr s1) (map (convert_new_thread_action (extNTA2J1 P)) \TA\\<^bsub>t\<^esub>)(t \ (((e', x'), exs), redT_updLns (locks s1) t (snd (the (thr s1 t))) \TA\\<^bsub>l\<^esub>))) t' = \((eX, eXS), LN)\" + assume Some: "((redT_updTs (thr s1) (map (convert_new_thread_action (extNTA2J1 P)) \TA\\<^bsub>t\<^esub>))(t \ (((e', x'), exs), redT_updLns (locks s1) t (snd (the (thr s1 t))) \TA\\<^bsub>l\<^esub>))) t' = \((eX, eXS), LN)\" { fix l from aoe have lao: "lock_actions_ok (locks s1 $ l) t (\ta\\<^bsub>l\<^esub> $ l)" by(auto simp add: lock_ok_las_def) have "has_locks ((redT_updLs (locks s1) t \ta\\<^bsub>l\<^esub>) $ l) t' + LN $ l = expr_lockss (map fst (eX # eXS)) l" proof(cases "t = t'") case True from loks thrst x have "has_locks (locks s1 $ l) t = expr_locks e l + expr_lockss (map fst exs) l" by(force simp add: lock_oks1_def) hence "lock_expr_locks_ok (locks s1 $ l) t 0 (int (expr_locks e l + expr_lockss (map fst exs) l))" by(simp add: lock_expr_locks_ok_def) with lao have "lock_expr_locks_ok (upd_locks (locks s1 $ l) t (\ta\\<^bsub>l\<^esub> $ l)) t (upd_threadRs 0 (locks s1 $ l) t (\ta\\<^bsub>l\<^esub> $ l)) (upd_expr_lock_actions (int (expr_locks e l + expr_lockss (map fst exs) l)) (\ta\\<^bsub>l\<^esub> $ l))" by(rule upd_locks_upd_expr_lock_preserve_lock_expr_locks_ok) moreover from sync thrst x have "sync_ok e" "el_loc_ok e x" unfolding el_loc_ok1_def by(auto dest: ts_okD) with red1_update_expr_locks[OF wf red] have "upd_expr_locks (int \ expr_locks e) \TA\\<^bsub>l\<^esub> = int \ expr_locks e'" by(simp) hence "upd_expr_lock_actions (int (expr_locks e l)) (\TA\\<^bsub>l\<^esub> $ l) = int (expr_locks e' l)" by(simp add: upd_expr_locks_def fun_eq_iff) ultimately show ?thesis using lao Some thrst x True by(auto simp add: lock_expr_locks_ok_def upd_expr_locks_def) next case False from aoe have tok: "thread_oks (thr s1) \ta\\<^bsub>t\<^esub>" by auto show ?thesis proof(cases "thr s1 t' = None") case True with Some tok False obtain m where nt: "NewThread t' (eX, eXS) m \ set (map (convert_new_thread_action (extNTA2J1 P)) \TA\\<^bsub>t\<^esub>)" and [simp]: "LN = no_wait_locks" by(auto dest: redT_updTs_new_thread) note expr_locks_new_thread1[OF wf red nt] moreover from loks True have "has_locks (locks s1 $ l) t' = 0" by(force simp add: lock_oks1_def) ultimately show ?thesis using lao False by simp next case False with Some \t \ t'\ tok have "thr s1 t' = \((eX, eXS), LN)\" by(fastforce dest: redT_updTs_Some[OF _ tok]) with loks tok lao \t \ t'\ show ?thesis by(cases eX)(auto simp add: lock_oks1_def) qed qed } hence "\l. has_locks ((redT_updLs (locks s1) t \ta\\<^bsub>l\<^esub>) $ l) t' + LN $ l = expr_lockss (map fst (eX # eXS)) l" .. } ultimately show ?thesis using s1' unfolding lock_oks1_def x' by(clarsimp simp del: fun_upd_apply) next case (red1Call e a M vs U Ts T body D x) from wf \P \ class_type_of U sees M: Ts\T = \body\ in D\ obtain T' where "P,Class D # Ts \1 body :: T'" by(auto simp add: wf_mdecl_def dest!: sees_wf_mdecl) hence "expr_locks (blocks1 0 (Class D#Ts) body) = (\l. 0)" by(auto simp add: expr_locks_blocks1 contains_insync_conv fun_eq_iff dest!: WT1_not_contains_insync) thus ?thesis using red1Call thrst loks s1' unfolding lock_oks1_def x' x by auto force+ next case (red1Return e' x' e x) thus ?thesis using thrst loks s1' unfolding lock_oks1_def x' x apply(auto simp add: redT_updWs_def elim!: rtrancl3p_cases) apply(erule_tac x=t in allE) apply(erule conjE) apply(erule disjE) apply(force simp add: expr_locks_inline_call_final ac_simps) apply(fastforce simp add: expr_locks_inline_call_final) apply hypsubst_thin apply(erule_tac x=ta in allE) apply fastforce done qed moreover from sync \mred1' P t (x, shr s1) ta (x', m')\ thrst aoe s1' have "ts_ok (\t exexs h. el_loc_ok1 exexs) (thr s1') (shr s1')" by(auto intro: lifting_wf.redT_updTs_preserves[OF Red1_el_loc_ok[OF wf]]) ultimately show ?thesis by simp next case (redT_acquire t x n ln) thus ?thesis using loks unfolding lock_oks1_def apply auto apply force apply(case_tac "ln $ l::nat") apply simp apply(erule allE) apply(erule conjE) apply(erule allE)+ apply(erule (1) impE) apply(erule_tac x=l in allE) apply fastforce apply(erule may_acquire_allE) apply(erule allE) apply(erule_tac x=l in allE) apply(erule impE) apply simp apply(simp only: has_locks_acquire_locks_conv) apply(erule conjE) apply(erule allE)+ apply(erule (1) impE) apply(erule_tac x=l in allE) apply simp done qed lemma Red1'_Red1_bisimulation: assumes wf: "wf_J1_prog P" shows "bisimulation (Red1_mthr.redT False P) (Red1_mthr.redT True P) mbisim_Red1'_Red1 (=)" proof fix s1 s2 tl1 s1' assume "mbisim_Red1'_Red1 s1 s2" and "Red1_mthr.redT False P s1 tl1 s1'" thus "\s2' tl2. Red1_mthr.redT True P s2 tl2 s2' \ mbisim_Red1'_Red1 s1' s2' \ tl1 = tl2" by(cases tl1)(auto simp add: mbisim_Red1'_Red1_def Red1_mthr_eq_Red1_mthr' simp del: split_paired_Ex elim: Red1'_preserves_lock_oks[OF wf] lifting_wf.redT_preserves[OF Red1_el_loc_ok, OF wf]) next fix s1 s2 tl2 s2' assume "mbisim_Red1'_Red1 s1 s2" "Red1_mthr.redT True P s2 tl2 s2'" thus "\s1' tl1. Red1_mthr.redT False P s1 tl1 s1' \ mbisim_Red1'_Red1 s1' s2' \ tl1 = tl2" by(cases tl2)(auto simp add: mbisim_Red1'_Red1_def Red1_mthr_eq_Red1_mthr' simp del: split_paired_Ex elim: Red1'_preserves_lock_oks[OF wf] lifting_wf.redT_preserves[OF Red1_el_loc_ok, OF wf]) qed lemma Red1'_Red1_bisimulation_final: "wf_J1_prog P \ bisimulation_final (Red1_mthr.redT False P) (Red1_mthr.redT True P) mbisim_Red1'_Red1 (=) Red1_mthr.mfinal Red1_mthr.mfinal" apply(intro_locales) apply(erule Red1'_Red1_bisimulation) apply(unfold_locales) apply(auto simp add: mbisim_Red1'_Red1_def) done lemma bisim_J1_J1_start: assumes wf: "wf_J1_prog P" and wf_start: "wf_start_state P C M vs" shows "mbisim_Red1'_Red1 (J1_start_state P C M vs) (J1_start_state P C M vs)" proof - from wf_start obtain Ts T body D where sees: "P \ C sees M:Ts\T=\body\ in D" and conf: "P,start_heap \ vs [:\] Ts" by cases let ?e = "blocks1 0 (Class C#Ts) body" let ?xs = "Null # vs @ replicate (max_vars body) undefined_value" from sees_wf_mdecl[OF wf sees] obtain T' where B: "\ body (Suc (length Ts))" and wt: "P,Class D # Ts \1 body :: T'" and da: "\ body \{..length Ts}\" and sv: "syncvars body" by(auto simp add: wf_mdecl_def) from wt have "expr_locks ?e = (\_. 0)" by(auto intro: WT1_expr_locks) thus ?thesis using da sees sv B unfolding start_state_def by(fastforce simp add: mbisim_Red1'_Red1_def lock_oks1_def el_loc_ok1_def contains_insync_conv intro!: ts_okI expr_locks_sync_ok split: if_split_asm intro: el_loc_okI) qed lemma Red1'_Red1_bisim_into_weak: assumes wf: "wf_J1_prog P" shows "bisimulation_into_delay (Red1_mthr.redT False P) (Red1_mthr.redT True P) mbisim_Red1'_Red1 (=) (Red1_mthr.m\move P) (Red1_mthr.m\move P)" proof - interpret b: bisimulation "Red1_mthr.redT False P" "Red1_mthr.redT True P" "mbisim_Red1'_Red1" "(=)" by(rule Red1'_Red1_bisimulation[OF wf]) show ?thesis by(unfold_locales)(simp add: mbisim_Red1'_Red1_def) qed end sublocale J1_heap_base < Red1_mthr: if_\multithreaded_wf final_expr1 "mred1g uf P" convert_RA "\MOVE1 P" for uf P by(unfold_locales) context J1_heap_base begin abbreviation if_lock_oks1 :: "('addr,'thread_id) locks \ ('addr,'thread_id,(status \ (('a,'b,'addr) exp \ 'c) \ (('a,'b,'addr) exp \ 'c) list)) thread_info \ bool" where "if_lock_oks1 ls ts \ lock_oks1 ls (init_fin_descend_thr ts)" definition if_mbisim_Red1'_Red1 :: "(('addr,'thread_id,status \ (('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list),'heap,'addr) state, ('addr,'thread_id,status \ (('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list),'heap,'addr) state) bisim" where "if_mbisim_Red1'_Red1 s1 s2 \ s1 = s2 \ if_lock_oks1 (locks s1) (thr s1) \ ts_ok (init_fin_lift (\t exexs h. el_loc_ok1 exexs)) (thr s1) (shr s1)" lemma if_mbisim_Red1'_Red1_imp_mbisim_Red1'_Red1: "if_mbisim_Red1'_Red1 s1 s2 \ mbisim_Red1'_Red1 (init_fin_descend_state s1) (init_fin_descend_state s2)" by(auto simp add: mbisim_Red1'_Red1_def if_mbisim_Red1'_Red1_def ts_ok_init_fin_descend_state) lemma if_Red1_mthr_imp_if_Red1_mthr': assumes lok: "if_lock_oks1 (locks s) (thr s)" and elo: "ts_ok (init_fin_lift (\t exexs h. el_loc_ok1 exexs)) (thr s) (shr s)" and Red: "Red1_mthr.if.redT uf P s tta s'" shows "Red1_mthr.if.redT (\ uf) P s tta s'" using Red proof(cases) case (redT_acquire t x ln n) from this(2-) show ?thesis unfolding redT_acquire(1) .. next case (redT_normal t x ta x' m') note aok = \Red1_mthr.if.actions_ok s t ta\ and tst = \thr s t = \(x, no_wait_locks)\\ from \Red1_mthr.init_fin uf P t (x, shr s) ta (x', m')\ have "Red1_mthr.init_fin (\ uf) P t (x, shr s) ta (x', m')" proof(cases) case InitialThreadAction show ?thesis unfolding InitialThreadAction .. next case (ThreadFinishAction exexs) from \final_expr1 exexs\ show ?thesis unfolding ThreadFinishAction .. next case (NormalAction exexs ta' exexs') let ?s = "init_fin_descend_state s" from lok have "lock_oks1 (locks ?s) (thr ?s)" by(simp) moreover from elo have elo: "ts_ok (\t exexs h. el_loc_ok1 exexs) (thr ?s) (shr ?s)" by(simp add: ts_ok_init_fin_descend_state) moreover from tst \x = (Running, exexs)\ have "thr ?s t = \(exexs, no_wait_locks)\" by simp moreover from aok have "Red1_mthr.actions_ok ?s t ta'" using \ta = convert_TA_initial (convert_obs_initial ta')\ by auto ultimately have "mred1 P t (exexs, shr ?s) ta' = mred1' P t (exexs, shr ?s) ta'" by(rule mred1_eq_mred1') with \mred1g uf P t (exexs, shr s) ta' (exexs', m')\ have "mred1g (\ uf) P t (exexs, shr s) ta' (exexs', m')" by(cases uf) simp_all thus ?thesis unfolding NormalAction(1-3) by(rule Red1_mthr.init_fin.NormalAction) qed thus ?thesis using tst aok \redT_upd s t ta x' m' s'\ unfolding \tta = (t, ta)\ .. qed lemma if_Red1_mthr_eq_if_Red1_mthr': assumes lok: "if_lock_oks1 (locks s) (thr s)" and elo: "ts_ok (init_fin_lift (\t exexs h. el_loc_ok1 exexs)) (thr s) (shr s)" shows "Red1_mthr.if.redT True P s = Red1_mthr.if.redT False P s" using if_Red1_mthr_imp_if_Red1_mthr'[OF assms, of True P, simplified] if_Red1_mthr_imp_if_Red1_mthr'[OF assms, of False P, simplified] by(blast del: equalityI) lemma if_Red1_el_loc_ok: assumes wf: "wf_J1_prog P" shows "lifting_wf Red1_mthr.init_fin_final (Red1_mthr.init_fin uf P) (init_fin_lift (\t exexs h. el_loc_ok1 exexs))" by(rule lifting_wf.lifting_wf_init_fin_lift)(rule Red1_el_loc_ok[OF wf]) lemma if_Red1'_preserves_if_lock_oks: assumes wf: "wf_J1_prog P" and Red: "Red1_mthr.if.redT False P s1 ta1 s1'" and loks: "if_lock_oks1 (locks s1) (thr s1)" and sync: "ts_ok (init_fin_lift (\t exexs h. el_loc_ok1 exexs)) (thr s1) (shr s1)" shows "if_lock_oks1 (locks s1') (thr s1')" proof - let ?s1 = "init_fin_descend_state s1" let ?s1' = "init_fin_descend_state s1'" from loks have loks': "lock_oks1 (locks ?s1) (thr ?s1)" by simp from sync have sync': "ts_ok (\t exexs h. el_loc_ok1 exexs) (thr ?s1) (shr ?s1)" by(simp add: ts_ok_init_fin_descend_state) from Red show ?thesis proof(cases) case (redT_acquire t x n ln) hence "Red1_mthr.redT False P ?s1 (t, K$ [], [], [], [], [], convert_RA ln) ?s1'" by(cases x)(auto intro!: Red1_mthr.redT.redT_acquire simp add: init_fin_descend_thr_def) with wf have "lock_oks1 (locks ?s1') (thr ?s1')" using loks' sync' by(rule Red1'_preserves_lock_oks) thus ?thesis by simp next case (redT_normal t sx ta sx' m') note tst = \thr s1 t = \(sx, no_wait_locks)\\ from \Red1_mthr.init_fin False P t (sx, shr s1) ta (sx', m')\ show ?thesis proof(cases) case (InitialThreadAction x) thus ?thesis using redT_normal loks by(cases x)(auto 4 3 simp add: init_fin_descend_thr_def redT_updLns_def expand_finfun_eq fun_eq_iff intro: lock_oks1_thr_updI) next case (ThreadFinishAction x) thus ?thesis using redT_normal loks by(cases x)(auto 4 3 simp add: init_fin_descend_thr_def redT_updLns_def expand_finfun_eq fun_eq_iff intro: lock_oks1_thr_updI) next case (NormalAction x ta' x') note ta = \ta = convert_TA_initial (convert_obs_initial ta')\ from \mred1' P t (x, shr s1) ta' (x', m')\ have "mred1' P t (x, shr ?s1) ta' (x', m')" by simp moreover have tst': "thr ?s1 t = \(x, no_wait_locks)\" using tst \sx = (Running, x)\ by simp moreover have "Red1_mthr.actions_ok ?s1 t ta'" using ta \Red1_mthr.if.actions_ok s1 t ta\ by simp moreover from \redT_upd s1 t ta sx' m' s1'\ tst tst' ta \sx' = (Running, x')\ have "redT_upd ?s1 t ta' x' m' ?s1'" by auto ultimately have "Red1_mthr.redT False P ?s1 (t, ta') ?s1'" .. with wf have "lock_oks1 (locks ?s1') (thr ?s1')" using loks' sync' by(rule Red1'_preserves_lock_oks) thus ?thesis by simp qed qed qed lemma Red1'_Red1_if_bisimulation: assumes wf: "wf_J1_prog P" shows "bisimulation (Red1_mthr.if.redT False P) (Red1_mthr.if.redT True P) if_mbisim_Red1'_Red1 (=)" proof fix s1 s2 tl1 s1' assume "if_mbisim_Red1'_Red1 s1 s2" and "Red1_mthr.if.redT False P s1 tl1 s1'" thus "\s2' tl2. Red1_mthr.if.redT True P s2 tl2 s2' \ if_mbisim_Red1'_Red1 s1' s2' \ tl1 = tl2" by(cases tl1)(auto simp add: if_mbisim_Red1'_Red1_def if_Red1_mthr_eq_if_Red1_mthr' simp del: split_paired_Ex elim: if_Red1'_preserves_if_lock_oks[OF wf] lifting_wf.redT_preserves[OF if_Red1_el_loc_ok, OF wf]) next fix s1 s2 tl2 s2' assume "if_mbisim_Red1'_Red1 s1 s2" "Red1_mthr.if.redT True P s2 tl2 s2'" thus "\s1' tl1. Red1_mthr.if.redT False P s1 tl1 s1' \ if_mbisim_Red1'_Red1 s1' s2' \ tl1 = tl2" by(cases tl2)(auto simp add: if_mbisim_Red1'_Red1_def if_Red1_mthr_eq_if_Red1_mthr' simp del: split_paired_Ex elim: if_Red1'_preserves_if_lock_oks[OF wf] lifting_wf.redT_preserves[OF if_Red1_el_loc_ok, OF wf]) qed lemma if_bisim_J1_J1_start: assumes wf: "wf_J1_prog P" and wf_start: "wf_start_state P C M vs" shows "if_mbisim_Red1'_Red1 (init_fin_lift_state status (J1_start_state P C M vs)) (init_fin_lift_state status (J1_start_state P C M vs))" proof - from assms have "mbisim_Red1'_Red1 (J1_start_state P C M vs) (J1_start_state P C M vs)" by(rule bisim_J1_J1_start) thus ?thesis by(simp add: if_mbisim_Red1'_Red1_def mbisim_Red1'_Red1_def)(simp add: init_fin_lift_state_conv_simps init_fin_descend_thr_def thr_init_fin_list_state' o_def map_option.compositionality map_option.identity split_beta) qed lemma if_Red1'_Red1_bisim_into_weak: assumes wf: "wf_J1_prog P" shows "bisimulation_into_delay (Red1_mthr.if.redT False P) (Red1_mthr.if.redT True P) if_mbisim_Red1'_Red1 (=) (Red1_mthr.if.m\move P) (Red1_mthr.if.m\move P)" proof - interpret b: bisimulation "Red1_mthr.if.redT False P" "Red1_mthr.if.redT True P" "if_mbisim_Red1'_Red1" "(=)" by(rule Red1'_Red1_if_bisimulation[OF wf]) show ?thesis by(unfold_locales)(simp add: if_mbisim_Red1'_Red1_def) qed lemma if_Red1'_Red1_bisimulation_final: "wf_J1_prog P \ bisimulation_final (Red1_mthr.if.redT False P) (Red1_mthr.if.redT True P) if_mbisim_Red1'_Red1 (=) Red1_mthr.if.mfinal Red1_mthr.if.mfinal" apply(intro_locales) apply(erule Red1'_Red1_if_bisimulation) apply(unfold_locales) apply(auto simp add: if_mbisim_Red1'_Red1_def) done end end diff --git a/thys/JinjaThreads/Compiler/ListIndex.thy b/thys/JinjaThreads/Compiler/ListIndex.thy --- a/thys/JinjaThreads/Compiler/ListIndex.thy +++ b/thys/JinjaThreads/Compiler/ListIndex.thy @@ -1,223 +1,223 @@ (* Title: JinjaThreads/Compiler/ListIndex.thy Author: Tobias Nipkow, Andreas Lochbihler *) section \Indexing variables in variable lists\ theory ListIndex imports Main begin text\In order to support local variables and arbitrarily nested blocks, the local variables are arranged as an indexed list. The outermost local variable (``this'') is the first element in the list, the most recently created local variable the last element. When descending into a block structure, a corresponding list @{term Vs} of variable names is maintained. To find the index of some variable @{term V}, we have to find the index of the \emph{last} occurrence of @{term V} in @{term Vs}. This is what @{term index} does:\ primrec index :: "'a list \ 'a \ nat" where "index [] y = 0" | "index (x#xs) y = (if x=y then if x \ set xs then index xs y + 1 else 0 else index xs y + 1)" definition hidden :: "'a list \ nat \ bool" where "hidden xs i \ i < size xs \ xs!i \ set(drop (i+1) xs)" subsection \@{term index}\ lemma [simp]: "index (xs @ [x]) x = size xs" (*<*)by(induct xs) simp_all(*>*) lemma [simp]: "(index (xs @ [x]) y = size xs) = (x = y)" (*<*)by(induct xs) auto(*>*) lemma [simp]: "x \ set xs \ xs ! index xs x = x" (*<*)by(induct xs) auto(*>*) lemma [simp]: "x \ set xs \ index xs x = size xs" (*<*)by(induct xs) auto(*>*) lemma index_size_conv[simp]: "(index xs x = size xs) = (x \ set xs)" (*<*)by(induct xs) auto(*>*) lemma size_index_conv[simp]: "(size xs = index xs x) = (x \ set xs)" (*<*)by(induct xs) auto(*>*) lemma "(index xs x < size xs) = (x \ set xs)" (*<*)by(induct xs) auto(*>*) lemma [simp]: "\ y \ set xs; x \ y \ \ index (xs @ [x]) y = index xs y" (*<*)by(induct xs) auto(*>*) lemma index_less_size[simp]: "x \ set xs \ index xs x < size xs" (*<*) apply(induct xs) apply simp apply(fastforce) done (*>*) lemma index_less_aux: "\x \ set xs; size xs \ n\ \ index xs x < n" (*<*) apply(subgoal_tac "index xs x < size xs") apply(simp (no_asm_simp)) apply simp done (*>*) lemma [simp]: "x \ set xs \ y \ set xs \ (index xs x = index xs y) = (x = y)" (*<*)by (induct xs) auto(*>*) lemma inj_on_index: "inj_on (index xs) (set xs)" (*<*)by(simp add:inj_on_def)(*>*) lemma index_drop: "\x i. \ x \ set xs; index xs x < i \ \ x \ set(drop i xs)" (*<*) apply(induct xs) apply (auto simp:drop_Cons split:if_split_asm nat.splits dest:in_set_dropD) done (*>*) subsection \@{term hidden}\ lemma hidden_index: "x \ set xs \ hidden (xs @ [x]) (index xs x)" (*<*) apply(auto simp add:hidden_def index_less_aux nth_append) apply(drule index_less_size) apply(simp del:index_less_size) done (*>*) lemma hidden_inacc: "hidden xs i \ index xs x \ i" (*<*) apply(case_tac "x \ set xs") apply(auto simp add:hidden_def index_less_aux nth_append index_drop) done (*>*) lemma [simp]: "hidden xs i \ hidden (xs@[x]) i" (*<*)by(auto simp add:hidden_def nth_append)(*>*) lemma fun_upds_apply: "\m ys. (m(xs[\]ys)) x = (let xs' = take (size ys) xs in if x \ set xs' then Some(ys ! index xs' x) else m x)" (*<*) apply(induct xs) apply (simp add:Let_def) apply(case_tac ys) apply (simp add:Let_def) apply (simp add:Let_def) done (*>*) lemma map_upds_apply_eq_Some: "((m(xs[\]ys)) x = Some y) = (let xs' = take (size ys) xs in if x \ set xs' then ys ! index xs' x = y else m x = Some y)" (*<*)by(simp add:fun_upds_apply Let_def)(*>*) lemma map_upds_upd_conv_index: "\x \ set xs; size xs \ size ys \ - \ m(xs[\]ys)(x\y) = m(xs[\]ys[index xs x := y])" + \ m(xs[\]ys, x\y) = m(xs[\]ys[index xs x := y])" (*<*) apply(rule ext) apply(simp add:fun_upds_apply index_less_aux eq_sym_conv Let_def) done (*>*) lemma image_index: "A \ set(xs@[x]) \ index (xs @ [x]) ` A = (if x \ A then insert (size xs) (index xs ` (A-{x})) else index xs ` A)" (*<*) apply(auto simp:image_def) apply(rule bexI) prefer 2 apply blast apply simp apply(rule ccontr) apply(erule_tac x=xa in ballE) prefer 2 apply blast apply(fastforce simp add:neq_commute) apply(subgoal_tac "x \ xa") prefer 2 apply blast apply(fastforce simp add:neq_commute) apply(subgoal_tac "x \ xa") prefer 2 apply blast apply(force) done (*>*) lemma index_le_lengthD: "index xs x < length xs \ x \ set xs" by(erule contrapos_pp)(simp) lemma not_hidden_index_nth: "\ i < length Vs; \ hidden Vs i \ \ index Vs (Vs ! i) = i" by(induct Vs arbitrary: i)(auto split: if_split_asm nat.split_asm simp add: nth_Cons hidden_def) lemma hidden_snoc_nth: assumes len: "i < length Vs" shows "hidden (Vs @ [Vs ! i]) i" proof(cases "hidden Vs i") case True thus ?thesis by simp next case False with len have "index Vs (Vs ! i) = i" by(rule not_hidden_index_nth) moreover from len have "hidden (Vs @ [Vs ! i]) (index Vs (Vs ! i))" by(auto intro: hidden_index) ultimately show ?thesis by simp qed lemma map_upds_Some_eq_nth_index: assumes "[Vs [\] vs] V = Some v" "length Vs \ length vs" shows "vs ! index Vs V = v" proof - from \[Vs [\] vs] V = Some v\ have "V \ set Vs" by -(rule classical, auto) with \[Vs [\] vs] V = Some v\ \length Vs \ length vs\ show ?thesis proof(induct Vs arbitrary: vs) case Nil thus ?case by simp next case (Cons x xs ys) note IH = \\vs. \ [xs [\] vs] V = Some v; length xs \ length vs; V \ set xs \ \ vs ! index xs V = v\ from \[x # xs [\] ys] V = Some v\ obtain y Ys where "ys = y # Ys" by(cases ys, auto) with \length (x # xs) \ length ys\ have "length xs \ length Ys" by simp show ?case proof(cases "V \ set xs") case True with \[x # xs [\] ys] V = Some v\ \length xs \ length Ys\ \ys = y # Ys\ have "[xs [\] Ys] V = Some v" apply(auto simp add: map_upds_def map_of_eq_None_iff set_zip image_Collect split: if_split_asm) apply(clarsimp simp add: in_set_conv_decomp) apply(hypsubst_thin) apply(erule_tac x="length ys" in allE) by(simp) with IH[OF this \length xs \ length Ys\ True] \ys = y # Ys\ True show ?thesis by(simp) next case False with \V \ set (x # xs)\ have "x = V" by auto with False \[x # xs [\] ys] V = Some v\ \ys = y # Ys\ have "y = v" by(auto) with False \x = V\ \ys = y # Ys\ show ?thesis by(simp) qed qed qed end diff --git a/thys/JinjaThreads/Framework/FWBisimulation.thy b/thys/JinjaThreads/Framework/FWBisimulation.thy --- a/thys/JinjaThreads/Framework/FWBisimulation.thy +++ b/thys/JinjaThreads/Framework/FWBisimulation.thy @@ -1,1759 +1,1759 @@ (* Title: JinjaThreads/Framework/FWBisimulation.thy Author: Andreas Lochbihler *) section \Bisimulation relations for the multithreaded semantics\ theory FWBisimulation imports FWLTS Bisimulation begin subsection \Definitions for lifting bisimulation relations\ primrec nta_bisim :: "('t \ ('x1 \ 'm1, 'x2 \ 'm2) bisim) \ (('t,'x1,'m1) new_thread_action, ('t,'x2,'m2) new_thread_action) bisim" where [code del]: "nta_bisim bisim (NewThread t x m) ta = (\x' m'. ta = NewThread t x' m' \ bisim t (x, m) (x', m'))" | "nta_bisim bisim (ThreadExists t b) ta = (ta = ThreadExists t b)" lemma nta_bisim_1_code [code]: "nta_bisim bisim (NewThread t x m) ta = (case ta of NewThread t' x' m' \ t = t' \ bisim t (x, m) (x', m') | _ \ False)" by(auto split: new_thread_action.split) lemma nta_bisim_simps_sym [simp]: "nta_bisim bisim ta (NewThread t x m) = (\x' m'. ta = NewThread t x' m' \ bisim t (x', m') (x, m))" "nta_bisim bisim ta (ThreadExists t b) = (ta = ThreadExists t b)" by(cases ta, auto)+ definition ta_bisim :: "('t \ ('x1 \ 'm1, 'x2 \ 'm2) bisim) \ (('l,'t,'x1,'m1,'w,'o) thread_action, ('l,'t,'x2,'m2,'w,'o) thread_action) bisim" where "ta_bisim bisim ta1 ta2 \ \ ta1 \\<^bsub>l\<^esub> = \ ta2 \\<^bsub>l\<^esub> \ \ ta1 \\<^bsub>w\<^esub> = \ ta2 \\<^bsub>w\<^esub> \ \ ta1 \\<^bsub>c\<^esub> = \ ta2 \\<^bsub>c\<^esub> \ \ ta1 \\<^bsub>o\<^esub> = \ ta2 \\<^bsub>o\<^esub> \ \ ta1 \\<^bsub>i\<^esub> = \ ta2 \\<^bsub>i\<^esub> \ list_all2 (nta_bisim bisim) \ ta1 \\<^bsub>t\<^esub> \ ta2 \\<^bsub>t\<^esub>" lemma ta_bisim_empty [iff]: "ta_bisim bisim \ \" by(auto simp add: ta_bisim_def) lemma ta_bisim_\ [simp]: "ta_bisim b \ ta' \ ta' = \" "ta_bisim b ta \ \ ta = \" apply(cases ta', fastforce simp add: ta_bisim_def) apply(cases ta, fastforce simp add: ta_bisim_def) done lemma nta_bisim_mono: assumes major: "nta_bisim bisim ta ta'" and mono: "\t s1 s2. bisim t s1 s2 \ bisim' t s1 s2" shows "nta_bisim bisim' ta ta'" using major by(cases ta)(auto intro: mono) lemma ta_bisim_mono: assumes major: "ta_bisim bisim ta1 ta2" and mono: "\t s1 s2. bisim t s1 s2 \ bisim' t s1 s2" shows "ta_bisim bisim' ta1 ta2" using major by(auto simp add: ta_bisim_def elim!: List.list_all2_mono nta_bisim_mono intro: mono) lemma nta_bisim_flip [flip_simps]: "nta_bisim (\t. flip (bisim t)) = flip (nta_bisim bisim)" by(rule ext)(case_tac x, auto simp add: flip_simps) lemma ta_bisim_flip [flip_simps]: "ta_bisim (\t. flip (bisim t)) = flip (ta_bisim bisim)" by(auto simp add: fun_eq_iff flip_simps ta_bisim_def) locale FWbisimulation_base = r1: multithreaded_base final1 r1 convert_RA + r2: multithreaded_base final2 r2 convert_RA for final1 :: "'x1 \ bool" and r1 :: "('l,'t,'x1,'m1,'w,'o) semantics" ("_ \ _ -1-_\ _" [50, 0, 0, 50] 80) and final2 :: "'x2 \ bool" and r2 :: "('l,'t,'x2,'m2,'w,'o) semantics" ("_ \ _ -2-_\ _" [50, 0, 0, 50] 80) and convert_RA :: "'l released_locks \ 'o list" + fixes bisim :: "'t \ ('x1 \ 'm1, 'x2 \ 'm2) bisim" ("_ \ _/ \ _" [50, 50, 50] 60) and bisim_wait :: "('x1, 'x2) bisim" ("_/ \w _" [50, 50] 60) begin notation r1.redT_syntax1 ("_ -1-_\_\ _" [50,0,0,50] 80) notation r2.redT_syntax1 ("_ -2-_\_\ _" [50,0,0,50] 80) notation r1.RedT ("_ -1-\_\* _" [50,0,50] 80) notation r2.RedT ("_ -2-\_\* _" [50,0,50] 80) notation r1.must_sync ("_ \ \_,/ _\/ \1" [50,0,0] 81) notation r2.must_sync ("_ \ \_,/ _\/ \2" [50,0,0] 81) notation r1.can_sync ("_ \ \_,/ _\/ _/ \1" [50,0,0,0] 81) notation r2.can_sync ("_ \ \_,/ _\/ _/ \2" [50,0,0,0] 81) abbreviation ta_bisim_bisim_syntax ("_/ \m _" [50, 50] 60) where "ta1 \m ta2 \ ta_bisim bisim ta1 ta2" definition tbisim :: "bool \ 't \ ('x1 \ 'l released_locks) option \ 'm1 \ ('x2 \ 'l released_locks) option \ 'm2 \ bool" where "\ln. tbisim nw t ts1 m1 ts2 m2 \ (case ts1 of None \ ts2 = None | \(x1, ln)\ \ (\x2. ts2 = \(x2, ln)\ \ t \ (x1, m1) \ (x2, m2) \ (nw \ x1 \w x2)))" lemma tbisim_NoneI: "tbisim w t None m None m'" by(simp add: tbisim_def) lemma tbisim_SomeI: "\ln. \ t \ (x, m) \ (x', m'); nw \ x \w x' \ \ tbisim nw t (Some (x, ln)) m (Some (x', ln)) m'" by(simp add: tbisim_def) lemma tbisim_cases[consumes 1, case_names None Some]: assumes major: "tbisim nw t ts1 m1 ts2 m2" and "\ ts1 = None; ts2 = None \ \ thesis" and "\x ln x'. \ ts1 = \(x, ln)\; ts2 = \(x', ln)\; t \ (x, m1) \ (x', m2); nw \ x \w x' \ \ thesis" shows thesis using assms by(auto simp add: tbisim_def) definition mbisim :: "(('l,'t,'x1,'m1,'w) state, ('l,'t,'x2,'m2,'w) state) bisim" ("_ \m _" [50, 50] 60) where "s1 \m s2 \ finite (dom (thr s1)) \ locks s1 = locks s2 \ wset s1 = wset s2 \ wset_thread_ok (wset s1) (thr s1) \ interrupts s1 = interrupts s2 \ (\t. tbisim (wset s2 t = None) t (thr s1 t) (shr s1) (thr s2 t) (shr s2))" lemma mbisim_thrNone_eq: "s1 \m s2 \ thr s1 t = None \ thr s2 t = None" unfolding mbisim_def tbisim_def apply(clarify) apply(erule allE[where x=t]) apply(clarsimp) done lemma mbisim_thrD1: "\ln. \ s1 \m s2; thr s1 t = \(x, ln)\ \ \ \x'. thr s2 t = \(x', ln)\ \ t \ (x, shr s1) \ (x', shr s2) \ (wset s1 t = None \ x \w x')" by(fastforce simp add: mbisim_def tbisim_def) lemma mbisim_thrD2: "\ln. \ s1 \m s2; thr s2 t = \(x, ln)\ \ \ \x'. thr s1 t = \(x', ln)\ \ t \ (x', shr s1) \ (x, shr s2) \ (wset s2 t = None \ x' \w x)" by(frule mbisim_thrNone_eq[where t=t])(cases "thr s1 t",(fastforce simp add: mbisim_def tbisim_def)+) lemma mbisim_dom_eq: "s1 \m s2 \ dom (thr s1) = dom (thr s2)" apply(clarsimp simp add: dom_def fun_eq_iff simp del: not_None_eq) apply(rule Collect_cong) apply(drule mbisim_thrNone_eq) apply(simp del: not_None_eq) done lemma mbisim_wset_thread_ok1: "s1 \m s2 \ wset_thread_ok (wset s1) (thr s1)" by(clarsimp simp add: mbisim_def) lemma mbisim_wset_thread_ok2: assumes "s1 \m s2" shows "wset_thread_ok (wset s2) (thr s2)" using assms apply(clarsimp simp add: mbisim_def) apply(auto intro!: wset_thread_okI simp add: mbisim_thrNone_eq[OF assms, THEN sym] dest: wset_thread_okD) done lemma mbisimI: "\ finite (dom (thr s1)); locks s1 = locks s2; wset s1 = wset s2; interrupts s1 = interrupts s2; wset_thread_ok (wset s1) (thr s1); \t. thr s1 t = None \ thr s2 t = None; \t x1 ln. thr s1 t = \(x1, ln)\ \ \x2. thr s2 t = \(x2, ln)\ \ t \ (x1, shr s1) \ (x2, shr s2) \ (wset s2 t = None \ x1 \w x2) \ \ s1 \m s2" by(fastforce simp add: mbisim_def tbisim_def) lemma mbisimI2: "\ finite (dom (thr s2)); locks s1 = locks s2; wset s1 = wset s2; interrupts s1 = interrupts s2; wset_thread_ok (wset s2) (thr s2); \t. thr s2 t = None \ thr s1 t = None; \t x2 ln. thr s2 t = \(x2, ln)\ \ \x1. thr s1 t = \(x1, ln)\ \ t \ (x1, shr s1) \ (x2, shr s2) \ (wset s2 t = None \ x1 \w x2) \ \ s1 \m s2" apply(auto simp add: mbisim_def tbisim_def) prefer 2 apply(rule wset_thread_okI) apply(case_tac "thr s2 t") apply(auto dest!: wset_thread_okD)[1] apply fastforce apply(erule back_subst[where P=finite]) apply(clarsimp simp add: dom_def fun_eq_iff simp del: not_None_eq) defer apply(rename_tac t) apply(case_tac [!] "thr s2 t") by fastforce+ lemma mbisim_finite1: "s1 \m s2 \ finite (dom (thr s1))" by(simp add: mbisim_def) lemma mbisim_finite2: "s1 \m s2 \ finite (dom (thr s2))" by(frule mbisim_finite1)(simp add: mbisim_dom_eq) definition mta_bisim :: "('t \ ('l,'t,'x1,'m1,'w,'o) thread_action, 't \ ('l,'t,'x2,'m2,'w,'o) thread_action) bisim" ("_/ \T _" [50, 50] 60) where "tta1 \T tta2 \ fst tta1 = fst tta2 \ snd tta1 \m snd tta2" lemma mta_bisim_conv [simp]: "(t, ta1) \T (t', ta2) \ t = t' \ ta1 \m ta2" by(simp add: mta_bisim_def) definition bisim_inv :: "bool" where "bisim_inv \ (\s1 ta1 s1' s2 t. t \ s1 \ s2 \ t \ s1 -1-ta1\ s1' \ (\s2'. t \ s1' \ s2')) \ (\s2 ta2 s2' s1 t. t \ s1 \ s2 \ t \ s2 -2-ta2\ s2' \ (\s1'. t \ s1' \ s2'))" lemma bisim_invI: "\ \s1 ta1 s1' s2 t. \ t \ s1 \ s2; t \ s1 -1-ta1\ s1' \ \ \s2'. t \ s1' \ s2'; \s2 ta2 s2' s1 t. \ t \ s1 \ s2; t \ s2 -2-ta2\ s2' \ \ \s1'. t \ s1' \ s2' \ \ bisim_inv" by(auto simp add: bisim_inv_def) lemma bisim_invD1: "\ bisim_inv; t \ s1 \ s2; t \ s1 -1-ta1\ s1' \ \ \s2'. t \ s1' \ s2'" unfolding bisim_inv_def by blast lemma bisim_invD2: "\ bisim_inv; t \ s1 \ s2; t \ s2 -2-ta2\ s2' \ \ \s1'. t \ s1' \ s2'" unfolding bisim_inv_def by blast lemma thread_oks_bisim_inv: "\ \t. ts1 t = None \ ts2 t = None; list_all2 (nta_bisim bisim) tas1 tas2 \ \ thread_oks ts1 tas1 \ thread_oks ts2 tas2" proof(induct tas2 arbitrary: tas1 ts1 ts2) case Nil thus ?case by(simp) next case (Cons ta2 TAS2 tas1 TS1 TS2) note IH = \\ts1 tas1 ts2. \ \t. ts1 t = None \ ts2 t = None; list_all2 (nta_bisim bisim) tas1 TAS2 \ \ thread_oks ts1 tas1 \ thread_oks ts2 TAS2\ note eqNone = \\t. TS1 t = None \ TS2 t = None\[rule_format] hence fti: "free_thread_id TS1 = free_thread_id TS2" by(auto simp add: free_thread_id_def) from \list_all2 (nta_bisim bisim) tas1 (ta2 # TAS2)\ obtain ta1 TAS1 where "tas1 = ta1 # TAS1" "nta_bisim bisim ta1 ta2" "list_all2 (nta_bisim bisim) TAS1 TAS2" by(auto simp add: list_all2_Cons2) moreover { fix t from \nta_bisim bisim ta1 ta2\ have "redT_updT' TS1 ta1 t = None \ redT_updT' TS2 ta2 t = None" by(cases ta1, auto split: if_split_asm simp add: eqNone) } ultimately have "thread_oks (redT_updT' TS1 ta1) TAS1 \ thread_oks (redT_updT' TS2 ta2) TAS2" by -(rule IH, auto) moreover from \nta_bisim bisim ta1 ta2\ fti have "thread_ok TS1 ta1 = thread_ok TS2 ta2" by(cases ta1, auto) ultimately show ?case using \tas1 = ta1 # TAS1\ by auto qed lemma redT_updT_nta_bisim_inv: "\ nta_bisim bisim ta1 ta2; ts1 T = None \ ts2 T = None \ \ redT_updT ts1 ta1 T = None \ redT_updT ts2 ta2 T = None" by(cases ta1, auto) lemma redT_updTs_nta_bisim_inv: "\ list_all2 (nta_bisim bisim) tas1 tas2; ts1 T = None \ ts2 T = None \ \ redT_updTs ts1 tas1 T = None \ redT_updTs ts2 tas2 T = None" proof(induct tas1 arbitrary: tas2 ts1 ts2) case Nil thus ?case by(simp) next case (Cons TA1 TAS1 tas2 TS1 TS2) note IH = \\tas2 ts1 ts2. \list_all2 (nta_bisim bisim) TAS1 tas2; (ts1 T = None) = (ts2 T = None)\ \ (redT_updTs ts1 TAS1 T = None) = (redT_updTs ts2 tas2 T = None)\ from \list_all2 (nta_bisim bisim) (TA1 # TAS1) tas2\ obtain TA2 TAS2 where "tas2 = TA2 # TAS2" "nta_bisim bisim TA1 TA2" "list_all2 (nta_bisim bisim) TAS1 TAS2" by(auto simp add: list_all2_Cons1) from \nta_bisim bisim TA1 TA2\ \(TS1 T = None) = (TS2 T = None)\ have "redT_updT TS1 TA1 T = None \ redT_updT TS2 TA2 T = None" by(rule redT_updT_nta_bisim_inv) with IH[OF \list_all2 (nta_bisim bisim) TAS1 TAS2\, of "redT_updT TS1 TA1" "redT_updT TS2 TA2"] \tas2 = TA2 # TAS2\ show ?case by simp qed end lemma tbisim_flip [flip_simps]: "FWbisimulation_base.tbisim (\t. flip (bisim t)) (flip bisim_wait) w t ts2 m2 ts1 m1 = FWbisimulation_base.tbisim bisim bisim_wait w t ts1 m1 ts2 m2" unfolding FWbisimulation_base.tbisim_def flip_simps by auto lemma mbisim_flip [flip_simps]: "FWbisimulation_base.mbisim (\t. flip (bisim t)) (flip bisim_wait) s2 s1 = FWbisimulation_base.mbisim bisim bisim_wait s1 s2" apply(rule iffI) apply(frule FWbisimulation_base.mbisim_dom_eq) apply(frule FWbisimulation_base.mbisim_wset_thread_ok2) apply(fastforce simp add: FWbisimulation_base.mbisim_def flip_simps) apply(frule FWbisimulation_base.mbisim_dom_eq) apply(frule FWbisimulation_base.mbisim_wset_thread_ok2) apply(fastforce simp add: FWbisimulation_base.mbisim_def flip_simps) done lemma mta_bisim_flip [flip_simps]: "FWbisimulation_base.mta_bisim (\t. flip (bisim t)) = flip (FWbisimulation_base.mta_bisim bisim)" by(auto simp add: fun_eq_iff flip_simps FWbisimulation_base.mta_bisim_def) lemma flip_const [simp]: "flip (\a b. c) = (\a b. c)" by(rule flip_def) lemma mbisim_K_flip [flip_simps]: "FWbisimulation_base.mbisim (\t. flip (bisim t)) (\x1 x2. c) s1 s2 = FWbisimulation_base.mbisim bisim (\x1 x2. c) s2 s1" using mbisim_flip[of bisim "\x1 x2. c" s1 s2] unfolding flip_const . context FWbisimulation_base begin lemma mbisim_actions_ok_bisim_no_join_12: assumes mbisim: "mbisim s1 s2" and "collect_cond_actions \ta1\\<^bsub>c\<^esub> = {}" and "ta_bisim bisim ta1 ta2" and "r1.actions_ok s1 t ta1" shows "r2.actions_ok s2 t ta2" using assms mbisim_thrNone_eq[OF mbisim] by(auto simp add: ta_bisim_def mbisim_def intro: thread_oks_bisim_inv[THEN iffD1] r2.may_join_cond_action_oks) lemma mbisim_actions_ok_bisim_no_join_21: "\ mbisim s1 s2; collect_cond_actions \ta2\\<^bsub>c\<^esub> = {}; ta_bisim bisim ta1 ta2; r2.actions_ok s2 t ta2 \ \ r1.actions_ok s1 t ta1" using FWbisimulation_base.mbisim_actions_ok_bisim_no_join_12[where bisim="\t. flip (bisim t)" and bisim_wait="flip bisim_wait"] unfolding flip_simps . lemma mbisim_actions_ok_bisim_no_join: "\ mbisim s1 s2; collect_cond_actions \ta1\\<^bsub>c\<^esub> = {}; ta_bisim bisim ta1 ta2 \ \ r1.actions_ok s1 t ta1 = r2.actions_ok s2 t ta2" apply(rule iffI) apply(erule (3) mbisim_actions_ok_bisim_no_join_12) apply(erule mbisim_actions_ok_bisim_no_join_21[where ?ta2.0 = ta2]) apply(simp add: ta_bisim_def) apply assumption+ done end locale FWbisimulation_base_aux = FWbisimulation_base + r1: multithreaded final1 r1 convert_RA + r2: multithreaded final2 r2 convert_RA + constrains final1 :: "'x1 \ bool" and r1 :: "('l,'t,'x1,'m1,'w, 'o) semantics" and final2 :: "'x2 \ bool" and r2 :: "('l,'t,'x2,'m2,'w, 'o) semantics" and convert_RA :: "'l released_locks \ 'o list" and bisim :: "'t \ ('x1 \ 'm1, 'x2 \ 'm2) bisim" and bisim_wait :: "('x1, 'x2) bisim" begin lemma FWbisimulation_base_aux_flip: "FWbisimulation_base_aux final2 r2 final1 r1" by(unfold_locales) end lemma FWbisimulation_base_aux_flip_simps [flip_simps]: "FWbisimulation_base_aux final2 r2 final1 r1 = FWbisimulation_base_aux final1 r1 final2 r2" by(blast intro: FWbisimulation_base_aux.FWbisimulation_base_aux_flip) sublocale FWbisimulation_base_aux < mthr: bisimulation_final_base r1.redT r2.redT mbisim mta_bisim r1.mfinal r2.mfinal . declare split_paired_Ex [simp del] subsection \Lifting for delay bisimulations\ locale FWdelay_bisimulation_base = FWbisimulation_base _ _ _ r2 convert_RA bisim bisim_wait + r1: \multithreaded final1 r1 convert_RA \move1 + r2: \multithreaded final2 r2 convert_RA \move2 for r2 :: "('l,'t,'x2,'m2,'w,'o) semantics" ("_ \ _ -2-_\ _" [50,0,0,50] 80) and convert_RA :: "'l released_locks \ 'o list" and bisim :: "'t \ ('x1 \ 'm1, 'x2 \ 'm2) bisim" ("_ \ _/ \ _" [50, 50, 50] 60) and bisim_wait :: "('x1, 'x2) bisim" ("_/ \w _" [50, 50] 60) and \move1 :: "('l,'t,'x1,'m1,'w,'o) \moves" and \move2 :: "('l,'t,'x2,'m2,'w,'o) \moves" begin abbreviation \mred1 :: "('l,'t,'x1,'m1,'w) state \ ('l,'t,'x1,'m1,'w) state \ bool" where "\mred1 \ r1.\mredT" abbreviation \mred2 :: "('l,'t,'x2,'m2,'w) state \ ('l,'t,'x2,'m2,'w) state \ bool" where "\mred2 \ r2.\mredT" abbreviation m\move1 :: "(('l,'t,'x1,'m1,'w) state, 't \ ('l,'t,'x1,'m1,'w,'o) thread_action) trsys" where "m\move1 \ r1.m\move" abbreviation m\move2 :: "(('l,'t,'x2,'m2,'w) state, 't \ ('l,'t,'x2,'m2,'w,'o) thread_action) trsys" where "m\move2 \ r2.m\move" abbreviation \mRed1 :: "('l,'t,'x1,'m1,'w) state \ ('l,'t,'x1,'m1,'w) state \ bool" where "\mRed1 \ \mred1^**" abbreviation \mRed2 :: "('l,'t,'x2,'m2,'w) state \ ('l,'t,'x2,'m2,'w) state \ bool" where "\mRed2 \ \mred2^**" abbreviation \mtRed1 :: "('l,'t,'x1,'m1,'w) state \ ('l,'t,'x1,'m1,'w) state \ bool" where "\mtRed1 \ \mred1^++" abbreviation \mtRed2 :: "('l,'t,'x2,'m2,'w) state \ ('l,'t,'x2,'m2,'w) state \ bool" where "\mtRed2 \ \mred2^++" lemma bisim_inv_\s1_inv: assumes inv: "bisim_inv" and bisim: "t \ s1 \ s2" and red: "r1.silent_moves t s1 s1'" obtains s2' where "t \ s1' \ s2'" proof(atomize_elim) from red bisim show "\s2'. t \ s1' \ s2'" by(induct rule: rtranclp_induct)(fastforce elim: bisim_invD1[OF inv])+ qed lemma bisim_inv_\s2_inv: assumes inv: "bisim_inv" and bisim: "t \ s1 \ s2" and red: "r2.silent_moves t s2 s2'" obtains s1' where "t \ s1' \ s2'" proof(atomize_elim) from red bisim show "\s1'. t \ s1' \ s2'" by(induct rule: rtranclp_induct)(fastforce elim: bisim_invD2[OF inv])+ qed primrec activate_cond_action1 :: "('l,'t,'x1,'m1,'w) state \ ('l,'t,'x2,'m2,'w) state \ 't conditional_action \ ('l,'t,'x1,'m1,'w) state" where "activate_cond_action1 s1 s2 (Join t) = (case thr s1 t of None \ s1 | \(x1, ln1)\ \ (case thr s2 t of None \ s1 | \(x2, ln2)\ \ if final2 x2 \ ln2 = no_wait_locks then redT_upd_\ s1 t (SOME x1'. r1.silent_moves t (x1, shr s1) (x1', shr s1) \ final1 x1' \ t \ (x1', shr s1) \ (x2, shr s2)) (shr s1) else s1))" | "activate_cond_action1 s1 s2 Yield = s1" primrec activate_cond_actions1 :: "('l,'t,'x1,'m1,'w) state \ ('l,'t,'x2,'m2,'w) state \ ('t conditional_action) list \ ('l,'t,'x1,'m1,'w) state" where "activate_cond_actions1 s1 s2 [] = s1" | "activate_cond_actions1 s1 s2 (ct # cts) = activate_cond_actions1 (activate_cond_action1 s1 s2 ct) s2 cts" primrec activate_cond_action2 :: "('l,'t,'x1,'m1,'w) state \ ('l,'t,'x2,'m2,'w) state \ 't conditional_action \ ('l,'t,'x2,'m2,'w) state" where "activate_cond_action2 s1 s2 (Join t) = (case thr s2 t of None \ s2 | \(x2, ln2)\ \ (case thr s1 t of None \ s2 | \(x1, ln1)\ \ if final1 x1 \ ln1 = no_wait_locks then redT_upd_\ s2 t (SOME x2'. r2.silent_moves t (x2, shr s2) (x2', shr s2) \ final2 x2' \ t \ (x1, shr s1) \ (x2', shr s2)) (shr s2) else s2))" | "activate_cond_action2 s1 s2 Yield = s2" primrec activate_cond_actions2 :: "('l,'t,'x1,'m1,'w) state \ ('l,'t,'x2,'m2,'w) state \ ('t conditional_action) list \ ('l,'t,'x2,'m2,'w) state" where "activate_cond_actions2 s1 s2 [] = s2" | "activate_cond_actions2 s1 s2 (ct # cts) = activate_cond_actions2 s1 (activate_cond_action2 s1 s2 ct) cts" end lemma activate_cond_action1_flip [flip_simps]: "FWdelay_bisimulation_base.activate_cond_action1 final2 r2 final1 (\t. flip (bisim t)) \move2 s2 s1 = FWdelay_bisimulation_base.activate_cond_action2 final1 final2 r2 bisim \move2 s1 s2" apply(rule ext) apply(case_tac x) apply(simp_all only: FWdelay_bisimulation_base.activate_cond_action1.simps FWdelay_bisimulation_base.activate_cond_action2.simps flip_simps) done lemma activate_cond_actions1_flip [flip_simps]: "FWdelay_bisimulation_base.activate_cond_actions1 final2 r2 final1 (\t. flip (bisim t)) \move2 s2 s1 = FWdelay_bisimulation_base.activate_cond_actions2 final1 final2 r2 bisim \move2 s1 s2" (is "?lhs = ?rhs") proof(rule ext) fix xs show "?lhs xs = ?rhs xs" by(induct xs arbitrary: s2) (simp_all only: FWdelay_bisimulation_base.activate_cond_actions1.simps FWdelay_bisimulation_base.activate_cond_actions2.simps flip_simps) qed lemma activate_cond_action2_flip [flip_simps]: "FWdelay_bisimulation_base.activate_cond_action2 final2 final1 r1 (\t. flip (bisim t)) \move1 s2 s1 = FWdelay_bisimulation_base.activate_cond_action1 final1 r1 final2 bisim \move1 s1 s2" apply(rule ext) apply(case_tac x) apply(simp_all only: FWdelay_bisimulation_base.activate_cond_action1.simps FWdelay_bisimulation_base.activate_cond_action2.simps flip_simps) done lemma activate_cond_actions2_flip [flip_simps]: "FWdelay_bisimulation_base.activate_cond_actions2 final2 final1 r1 (\t. flip (bisim t)) \move1 s2 s1 = FWdelay_bisimulation_base.activate_cond_actions1 final1 r1 final2 bisim \move1 s1 s2" (is "?lhs = ?rhs") proof(rule ext) fix xs show "?lhs xs = ?rhs xs" by(induct xs arbitrary: s1) (simp_all only: FWdelay_bisimulation_base.activate_cond_actions1.simps FWdelay_bisimulation_base.activate_cond_actions2.simps flip_simps) qed context FWdelay_bisimulation_base begin lemma shr_activate_cond_action1 [simp]: "shr (activate_cond_action1 s1 s2 ct) = shr s1" by(cases ct) simp_all lemma shr_activate_cond_actions1 [simp]: "shr (activate_cond_actions1 s1 s2 cts) = shr s1" by(induct cts arbitrary: s1) auto lemma shr_activate_cond_action2 [simp]: "shr (activate_cond_action2 s1 s2 ct) = shr s2" by(cases ct) simp_all lemma shr_activate_cond_actions2 [simp]: "shr (activate_cond_actions2 s1 s2 cts) = shr s2" by(induct cts arbitrary: s2) auto lemma locks_activate_cond_action1 [simp]: "locks (activate_cond_action1 s1 s2 ct) = locks s1" by(cases ct) simp_all lemma locks_activate_cond_actions1 [simp]: "locks (activate_cond_actions1 s1 s2 cts) = locks s1" by(induct cts arbitrary: s1) auto lemma locks_activate_cond_action2 [simp]: "locks (activate_cond_action2 s1 s2 ct) = locks s2" by(cases ct) simp_all lemma locks_activate_cond_actions2 [simp]: "locks (activate_cond_actions2 s1 s2 cts) = locks s2" by(induct cts arbitrary: s2) auto lemma wset_activate_cond_action1 [simp]: "wset (activate_cond_action1 s1 s2 ct) = wset s1" by(cases ct) simp_all lemma wset_activate_cond_actions1 [simp]: "wset (activate_cond_actions1 s1 s2 cts) = wset s1" by(induct cts arbitrary: s1) auto lemma wset_activate_cond_action2 [simp]: "wset (activate_cond_action2 s1 s2 ct) = wset s2" by(cases ct) simp_all lemma wset_activate_cond_actions2 [simp]: "wset (activate_cond_actions2 s1 s2 cts) = wset s2" by(induct cts arbitrary: s2) auto lemma interrupts_activate_cond_action1 [simp]: "interrupts (activate_cond_action1 s1 s2 ct) = interrupts s1" by(cases ct) simp_all lemma interrupts_activate_cond_actions1 [simp]: "interrupts (activate_cond_actions1 s1 s2 cts) = interrupts s1" by(induct cts arbitrary: s1) auto lemma interrupts_activate_cond_action2 [simp]: "interrupts (activate_cond_action2 s1 s2 ct) = interrupts s2" by(cases ct) simp_all lemma interrupts_activate_cond_actions2 [simp]: "interrupts (activate_cond_actions2 s1 s2 cts) = interrupts s2" by(induct cts arbitrary: s2) auto end locale FWdelay_bisimulation_lift_aux = FWdelay_bisimulation_base _ _ _ _ _ _ _ \move1 \move2 + r1: \multithreaded_wf final1 r1 convert_RA \move1 + r2: \multithreaded_wf final2 r2 convert_RA \move2 for \move1 :: "('l,'t,'x1,'m1,'w,'o) \moves" and \move2 :: "('l,'t,'x2,'m2,'w,'o) \moves" begin lemma FWdelay_bisimulation_lift_aux_flip: "FWdelay_bisimulation_lift_aux final2 r2 final1 r1 \move2 \move1" by unfold_locales end lemma FWdelay_bisimulation_lift_aux_flip_simps [flip_simps]: "FWdelay_bisimulation_lift_aux final2 r2 final1 r1 \move2 \move1 = FWdelay_bisimulation_lift_aux final1 r1 final2 r2 \move1 \move2" by(auto dest: FWdelay_bisimulation_lift_aux.FWdelay_bisimulation_lift_aux_flip simp only: flip_flip) context FWdelay_bisimulation_lift_aux begin lemma cond_actions_ok_\mred1_inv: assumes red: "\mred1 s1 s1'" and ct: "r1.cond_action_ok s1 t ct" shows "r1.cond_action_ok s1' t ct" using ct proof(cases ct) case (Join t') show ?thesis using red ct proof(cases "thr s1 t'") case None with red ct Join show ?thesis by(fastforce elim!: r1.mthr.silent_move.cases r1.redT.cases r1.m\move.cases rtrancl3p_cases dest: r1.silent_tl split: if_split_asm) next case (Some a) with red ct Join show ?thesis by(fastforce elim!: r1.mthr.silent_move.cases r1.redT.cases r1.m\move.cases rtrancl3p_cases dest: r1.silent_tl r1.final_no_red split: if_split_asm simp add: redT_updWs_def) qed next case Yield thus ?thesis by simp qed lemma cond_actions_ok_\mred2_inv: "\ \mred2 s2 s2'; r2.cond_action_ok s2 t ct \ \ r2.cond_action_ok s2' t ct" using FWdelay_bisimulation_lift_aux.cond_actions_ok_\mred1_inv[OF FWdelay_bisimulation_lift_aux_flip] . lemma cond_actions_ok_\mRed1_inv: "\ \mRed1 s1 s1'; r1.cond_action_ok s1 t ct \ \ r1.cond_action_ok s1' t ct" by(induct rule: rtranclp_induct)(blast intro: cond_actions_ok_\mred1_inv)+ lemma cond_actions_ok_\mRed2_inv: "\ \mRed2 s2 s2'; r2.cond_action_ok s2 t ct \ \ r2.cond_action_ok s2' t ct" by(rule FWdelay_bisimulation_lift_aux.cond_actions_ok_\mRed1_inv[OF FWdelay_bisimulation_lift_aux_flip]) end locale FWdelay_bisimulation_lift = FWdelay_bisimulation_lift_aux + constrains final1 :: "'x1 \ bool" and r1 :: "('l, 't, 'x1, 'm1, 'w, 'o) semantics" and final2 :: "'x2 \ bool" and r2 :: "('l, 't, 'x2, 'm2, 'w, 'o) semantics" and convert_RA :: "'l released_locks \ 'o list" and bisim :: "'t \ ('x1 \ 'm1, 'x2 \ 'm2) bisim" and bisim_wait :: "('x1, 'x2) bisim" and \move1 :: "('l, 't, 'x1, 'm1, 'w, 'o) \moves" and \move2 :: "('l, 't, 'x2, 'm2, 'w, 'o) \moves" assumes \inv_locale: "\inv (r1 t) (r2 t) (bisim t) (ta_bisim bisim) \move1 \move2" sublocale FWdelay_bisimulation_lift < \inv "r1 t" "r2 t" "bisim t" "ta_bisim bisim" \move1 \move2 for t by(rule \inv_locale) context FWdelay_bisimulation_lift begin lemma FWdelay_bisimulation_lift_flip: "FWdelay_bisimulation_lift final2 r2 final1 r1 (\t. flip (bisim t)) \move2 \move1" apply(rule FWdelay_bisimulation_lift.intro) apply(rule FWdelay_bisimulation_lift_aux_flip) apply(rule FWdelay_bisimulation_lift_axioms.intro) apply(unfold flip_simps) apply(unfold_locales) done end lemma FWdelay_bisimulation_lift_flip_simps [flip_simps]: "FWdelay_bisimulation_lift final2 r2 final1 r1 (\t. flip (bisim t)) \move2 \move1 = FWdelay_bisimulation_lift final1 r1 final2 r2 bisim \move1 \move2" by(auto dest: FWdelay_bisimulation_lift.FWdelay_bisimulation_lift_flip simp only: flip_flip) context FWdelay_bisimulation_lift begin lemma \inv_lift: "\inv r1.redT r2.redT mbisim mta_bisim m\move1 m\move2" proof fix s1 s2 tl1 s1' tl2 s2' assume "s1 \m s2" "s1' \m s2'" "tl1 \T tl2" "r1.redT s1 tl1 s1'" "r2.redT s2 tl2 s2'" moreover obtain t ta1 where tl1: "tl1 = (t, ta1)" by(cases tl1) moreover obtain t' ta2 where tl2: "tl2 = (t', ta2)" by(cases tl2) moreover obtain ls1 ts1 ws1 m1 is1 where s1: "s1 = (ls1, (ts1, m1), ws1, is1)" by(cases s1) fastforce moreover obtain ls2 ts2 ws2 m2 is2 where s2: "s2 = (ls2, (ts2, m2), ws2, is2)" by(cases s2) fastforce moreover obtain ls1' ts1' ws1' m1' is1' where s1': "s1' = (ls1', (ts1', m1'), ws1', is1')" by(cases s1') fastforce moreover obtain ls2' ts2' ws2' m2' is2' where s2': "s2' = (ls2', (ts2', m2'), ws2', is2')" by(cases s2') fastforce ultimately have mbisim: "(ls1, (ts1, m1), ws1, is1) \m (ls2, (ts2, m2), ws2, is2)" and mbisim': "(ls1', (ts1', m1'), ws1', is1') \m (ls2', (ts2', m2'), ws2', is2')" and mred1: "(ls1, (ts1, m1), ws1, is1) -1-t\ta1\ (ls1', (ts1', m1'), ws1', is1')" and mred2: "(ls2, (ts2, m2), ws2, is2) -2-t\ta2\ (ls2', (ts2', m2'), ws2', is2')" and tasim: "ta1 \m ta2" and tt': "t' = t" by simp_all from mbisim have ls: "ls1 = ls2" and ws: "ws1 = ws2" and "is": "is1 = is2" and tbisim: "\t. tbisim (ws2 t = None) t (ts1 t) m1 (ts2 t) m2" by(simp_all add: mbisim_def) from mbisim' have ls': "ls1' = ls2'" and ws': "ws1' = ws2'" and is': "is1' = is2'" and tbisim': "\t. tbisim (ws2' t = None) t (ts1' t) m1' (ts2' t) m2'" by(simp_all add: mbisim_def) from mred1 r1.redT_thread_not_disappear[OF mred1] obtain x1 ln1 x1' ln1' where tst1: "ts1 t = \(x1, ln1)\" and tst1': "ts1' t = \(x1', ln1')\" by(fastforce elim!: r1.redT.cases) from mred2 r2.redT_thread_not_disappear[OF mred2] obtain x2 ln2 x2' ln2' where tst2: "ts2 t = \(x2, ln2)\" and tst2': "ts2' t = \(x2', ln2')\" by(fastforce elim!: r2.redT.cases) from tbisim[of t] tst1 tst2 ws have bisim: "t \ (x1, m1) \ (x2, m2)" and ln: "ln1 = ln2" by(auto simp add: tbisim_def) from tbisim'[of t] tst1' tst2' have bisim': "t \ (x1', m1') \ (x2', m2')" and ln': "ln1' = ln2'" by(auto simp add: tbisim_def) show "m\move1 s1 tl1 s1' = m\move2 s2 tl2 s2'" unfolding s1 s2 s1' s2' tt' tl1 tl2 proof - show "m\move1 (ls1, (ts1, m1), ws1, is1) (t, ta1) (ls1', (ts1', m1'), ws1', is1') = m\move2 (ls2, (ts2, m2), ws2, is2) (t, ta2) (ls2', (ts2', m2'), ws2', is2')" (is "?lhs = ?rhs") proof assume m\: ?lhs with tst1 tst1' obtain \1: "\move1 (x1, m1) ta1 (x1', m1')" and ln1: "ln1 = no_wait_locks" by(fastforce elim!: r1.m\move.cases) from \1 have "ta1 = \" by(rule r1.silent_tl) with mred1 \1 tst1 tst1' ln1 have red1: "t \ (x1, m1) -1-ta1\ (x1', m1')" by(auto elim!: r1.redT.cases rtrancl3p_cases) from tasim \ta1 = \\ have [simp]: "ta2 = \" by(simp) with mred2 ln1 ln tst2 tst2' have red2: "t \ (x2, m2) -2-\\ (x2', m2')" by(fastforce elim!: r2.redT.cases rtrancl3p_cases) from \1 \inv[OF bisim red1 red2] bisim' tasim have \2: "\move2 (x2, m2) \ (x2', m2')" by simp with tst2 tst2' ln ln1 show ?rhs by -(rule r2.m\move.intros, auto) next assume m\: ?rhs with tst2 tst2' obtain \2: "\move2 (x2, m2) ta2 (x2', m2')" and ln2: "ln2 = no_wait_locks" by(fastforce elim!: r2.m\move.cases) from \2 have "ta2 = \" by(rule r2.silent_tl) with mred2 \2 tst2 tst2' ln2 have red2: "t \ (x2, m2) -2-ta2\ (x2', m2')" by(auto elim!: r2.redT.cases rtrancl3p_cases) from tasim \ta2 = \\ have [simp]: "ta1 = \" by simp with mred1 ln2 ln tst1 tst1' have red1: "t \ (x1, m1) -1-\\ (x1', m1')" by(fastforce elim!: r1.redT.cases rtrancl3p_cases) from \2 \inv[OF bisim red1 red2] bisim' tasim have \1: "\move1 (x1, m1) \ (x1', m1')" by auto with tst1 tst1' ln ln2 show ?lhs unfolding \ta1 = \\ by-(rule r1.m\move.intros, auto) qed qed qed end sublocale FWdelay_bisimulation_lift < mthr: \inv r1.redT r2.redT mbisim mta_bisim m\move1 m\move2 by(rule \inv_lift) locale FWdelay_bisimulation_final_base = FWdelay_bisimulation_lift_aux + constrains final1 :: "'x1 \ bool" and r1 :: "('l,'t,'x1,'m1,'w, 'o) semantics" and final2 :: "'x2 \ bool" and r2 :: "('l,'t,'x2,'m2,'w, 'o) semantics" and convert_RA :: "'l released_locks \ 'o list" and bisim :: "'t \ ('x1 \ 'm1, 'x2 \ 'm2) bisim" and bisim_wait :: "('x1, 'x2) bisim" and \move1 :: "('l,'t,'x1,'m1,'w, 'o) \moves" and \move2 :: "('l,'t,'x2,'m2,'w, 'o) \moves" assumes delay_bisim_locale: "delay_bisimulation_final_base (r1 t) (r2 t) (bisim t) \move1 \move2 (\(x1, m). final1 x1) (\(x2, m). final2 x2)" sublocale FWdelay_bisimulation_final_base < delay_bisimulation_final_base "r1 t" "r2 t" "bisim t" "ta_bisim bisim" \move1 \move2 "\(x1, m). final1 x1" "\(x2, m). final2 x2" for t by(rule delay_bisim_locale) context FWdelay_bisimulation_final_base begin lemma FWdelay_bisimulation_final_base_flip: "FWdelay_bisimulation_final_base final2 r2 final1 r1 (\t. flip (bisim t)) \move2 \move1" apply(rule FWdelay_bisimulation_final_base.intro) apply(rule FWdelay_bisimulation_lift_aux_flip) apply(rule FWdelay_bisimulation_final_base_axioms.intro) apply(rule delay_bisimulation_final_base_flip) done end lemma FWdelay_bisimulation_final_base_flip_simps [flip_simps]: "FWdelay_bisimulation_final_base final2 r2 final1 r1 (\t. flip (bisim t)) \move2 \move1 = FWdelay_bisimulation_final_base final1 r1 final2 r2 bisim \move1 \move2" by(auto dest: FWdelay_bisimulation_final_base.FWdelay_bisimulation_final_base_flip simp only: flip_flip) context FWdelay_bisimulation_final_base begin lemma cond_actions_ok_bisim_ex_\1_inv: fixes ls ts1 m1 ws "is" ts2 m2 ct defines "s1' \ activate_cond_action1 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) ct" assumes mbisim: "\t'. t' \ t \ tbisim (ws t' = None) t' (ts1 t') m1 (ts2 t') m2" and ts1t: "ts1 t = Some xln" and ts2t: "ts2 t = Some xln'" and ct: "r2.cond_action_ok (ls, (ts2, m2), ws, is) t ct" shows "\mRed1 (ls, (ts1, m1), ws, is) s1'" and "\t'. t' \ t \ tbisim (ws t' = None) t' (thr s1' t') m1 (ts2 t') m2" and "r1.cond_action_ok s1' t ct" and "thr s1' t = Some xln" proof - have "\mRed1 (ls, (ts1, m1), ws, is) s1' \ (\t'. t' \ t \ tbisim (ws t' = None) t' (thr s1' t') m1 (ts2 t') m2) \ r1.cond_action_ok s1' t ct \ thr s1' t = \xln\" using ct proof(cases ct) case (Join t') show ?thesis proof(cases "ts1 t'") case None with mbisim ts1t have "t \ t'" by auto moreover from None Join have "s1' = (ls, (ts1, m1), ws, is)" by(simp add: s1'_def) ultimately show ?thesis using mbisim Join ct None ts1t by(simp add: tbisim_def) next case (Some xln) moreover obtain x1 ln where "xln = (x1, ln)" by(cases xln) ultimately have ts1t': "ts1 t' = \(x1, ln)\" by simp from Join ct Some ts2t have tt': "t' \ t" by auto from mbisim[OF tt'] ts1t' obtain x2 where ts2t': "ts2 t' = \(x2, ln)\" and bisim: "t' \ (x1, m1) \ (x2, m2)" by(auto simp add: tbisim_def) from ct Join ts2t' have final2: "final2 x2" and ln: "ln = no_wait_locks" and wst': "ws t' = None" by simp_all let ?x1' = "SOME x. r1.silent_moves t' (x1, m1) (x, m1) \ final1 x \ t' \ (x, m1) \ (x2, m2)" { from final2_simulation[OF bisim] final2 obtain x1' m1' where "r1.silent_moves t' (x1, m1) (x1', m1')" and "t' \ (x1', m1') \ (x2, m2)" and "final1 x1'" by auto moreover hence "m1' = m1" using bisim by(auto dest: r1.red_rtrancl_\_heapD_inv) ultimately have "\x. r1.silent_moves t' (x1, m1) (x, m1) \ final1 x \ t' \ (x, m1) \ (x2, m2)" by blast } from someI_ex[OF this] have red1: "r1.silent_moves t' (x1, m1) (?x1', m1)" and final1: "final1 ?x1'" and bisim': "t' \ (?x1', m1) \ (x2, m2)" by blast+ let ?S1' = "redT_upd_\ (ls, (ts1, m1), ws, is) t' ?x1' m1" from r1.silent_moves_into_RedT_\_inv[where ?s="(ls, (ts1, m1), ws, is)" and t=t', simplified, OF red1] bisim ts1t' ln wst' have Red1: "\mRed1 (ls, (ts1, m1), ws, is) ?S1'" by auto moreover from Join ln ts1t' final1 wst' tt' have ct': "r1.cond_action_ok ?S1' t ct" by(auto intro: finfun_ext) { fix t'' assume "t \ t''" with Join mbisim[OF this[symmetric]] bisim' ts1t' ts2t' wst' s1'_def have "tbisim (ws t'' = None) t'' (thr s1' t'') m1 (ts2 t'') m2" by(auto simp add: tbisim_def redT_updLns_def o_def finfun_Diag_const2) } moreover from Join ts1t' ts2t' final2 ln have "s1' = ?S1'" by(simp add: s1'_def) ultimately show ?thesis using Red1 ct' ts1t' tt' ts1t by(auto) qed next case Yield thus ?thesis using mbisim ts1t by(simp add: s1'_def) qed thus "\mRed1 (ls, (ts1, m1), ws, is) s1'" and "\t'. t' \ t \ tbisim (ws t' = None) t' (thr s1' t') m1 (ts2 t') m2" and "r1.cond_action_ok s1' t ct" and "thr s1' t = \xln\" by blast+ qed lemma cond_actions_oks_bisim_ex_\1_inv: fixes ls ts1 m1 ws "is" ts2 m2 cts defines "s1' \ activate_cond_actions1 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) cts" assumes tbisim: "\t'. t' \ t \ tbisim (ws t' = None) t' (ts1 t') m1 (ts2 t') m2" and ts1t: "ts1 t = Some xln" and ts2t: "ts2 t = Some xln'" and ct: "r2.cond_action_oks (ls, (ts2, m2), ws, is) t cts" shows "\mRed1 (ls, (ts1, m1), ws, is) s1'" and "\t'. t' \ t \ tbisim (ws t' = None) t' (thr s1' t') m1 (ts2 t') m2" and "r1.cond_action_oks s1' t cts" and "thr s1' t = Some xln" using tbisim ts1t ct unfolding s1'_def proof(induct cts arbitrary: ts1) case (Cons ct cts) note IH1 = \\ts1. \\t'. t' \ t \ tbisim (ws t' = None) t' (ts1 t') m1 (ts2 t') m2; ts1 t = \xln\; r2.cond_action_oks (ls, (ts2, m2), ws, is) t cts\ \ \mred1\<^sup>*\<^sup>* (ls, (ts1, m1), ws, is) (activate_cond_actions1 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) cts)\ note IH2 = \\t' ts1. \t' \ t; \t'. t' \ t \ tbisim (ws t' = None) t' (ts1 t') m1 (ts2 t') m2; ts1 t = \xln\; r2.cond_action_oks (ls, (ts2, m2), ws, is) t cts\ \ tbisim (ws t' = None) t' (thr (activate_cond_actions1 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) cts) t') m1 (ts2 t') m2\ note IH3 = \\ts1. \\t'. t' \ t \ tbisim (ws t' = None) t' (ts1 t') m1 (ts2 t') m2; ts1 t = \xln\; r2.cond_action_oks (ls, (ts2, m2), ws, is) t cts\ \ r1.cond_action_oks (activate_cond_actions1 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) cts) t cts\ note IH4 = \\ts1. \\t'. t' \ t \ tbisim (ws t' = None) t' (ts1 t') m1 (ts2 t') m2; ts1 t = \xln\; r2.cond_action_oks (ls, (ts2, m2), ws, is) t cts\ \ thr (activate_cond_actions1 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) cts) t = \xln\\ { fix ts1 assume tbisim: "\t'. t' \ t \ tbisim (ws t' = None) t' (ts1 t') m1 (ts2 t') m2" and ts1t: "ts1 t = \xln\" and ct: "r2.cond_action_oks (ls, (ts2, m2), ws, is) t (ct # cts)" from ct have 1: "r2.cond_action_ok (ls, (ts2, m2), ws, is) t ct" and 2: "r2.cond_action_oks (ls, (ts2, m2), ws, is) t cts" by auto let ?s1' = "activate_cond_action1 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) ct" from cond_actions_ok_bisim_ex_\1_inv[OF tbisim, OF _ ts1t ts2t 1] have tbisim': "\t'. t' \ t \ tbisim (ws t' = None) t' (thr ?s1' t') m1 (ts2 t') m2" and red: "\mRed1 (ls, (ts1, m1), ws, is) ?s1'" and ct': "r1.cond_action_ok ?s1' t ct" and ts1't: "thr ?s1' t = \xln\" by blast+ let ?s1'' = "activate_cond_actions1 ?s1' (ls, (ts2, m2), ws, is) cts" have "locks ?s1' = ls" "shr ?s1' = m1" "wset ?s1' = ws" "interrupts ?s1' = is" by simp_all hence s1': "(ls, (thr ?s1', m1), ws, is) = ?s1'" by(cases "?s1'") auto from IH1[OF tbisim', OF _ ts1't 2] s1' have red': "\mRed1 ?s1' ?s1''" by simp with red show "\mRed1 (ls, (ts1, m1), ws, is) (activate_cond_actions1 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) (ct # cts))" by auto { fix t' assume t't: "t' \ t" from IH2[OF t't tbisim', OF _ ts1't 2] s1' show "tbisim (ws t' = None) t' (thr (activate_cond_actions1 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) (ct # cts)) t') m1 (ts2 t') m2" by auto } from red' ct' have "r1.cond_action_ok ?s1'' t ct" by(rule cond_actions_ok_\mRed1_inv) with IH3[OF tbisim', OF _ ts1't 2] s1' show "r1.cond_action_oks (activate_cond_actions1 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) (ct # cts)) t (ct # cts)" by auto from ts1't IH4[OF tbisim', OF _ ts1't 2] s1' show "thr (activate_cond_actions1 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) (ct # cts)) t = \xln\" by auto } qed(auto) lemma cond_actions_ok_bisim_ex_\2_inv: fixes ls ts1 m1 "is" ws ts2 m2 ct defines "s2' \ activate_cond_action2 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) ct" assumes mbisim: "\t'. t' \ t \ tbisim (ws t' = None) t' (ts1 t') m1 (ts2 t') m2" and ts1t: "ts1 t = Some xln" and ts2t: "ts2 t = Some xln'" and ct: "r1.cond_action_ok (ls, (ts1, m1), ws, is) t ct" shows "\mRed2 (ls, (ts2, m2), ws, is) s2'" and "\t'. t' \ t \ tbisim (ws t' = None) t' (ts1 t') m1 (thr s2' t') m2" and "r2.cond_action_ok s2' t ct" and "thr s2' t = Some xln'" unfolding s2'_def by(blast intro: FWdelay_bisimulation_final_base.cond_actions_ok_bisim_ex_\1_inv[OF FWdelay_bisimulation_final_base_flip, where bisim_wait = "flip bisim_wait", unfolded flip_simps, OF mbisim _ _ ct, OF _ ts2t ts1t])+ lemma cond_actions_oks_bisim_ex_\2_inv: fixes ls ts1 m1 ws "is" ts2 m2 cts defines "s2' \ activate_cond_actions2 (ls, (ts1, m1), ws, is) (ls, (ts2, m2), ws, is) cts" assumes tbisim: "\t'. t' \ t \ tbisim (ws t' = None) t' (ts1 t') m1 (ts2 t') m2" and ts1t: "ts1 t = Some xln" and ts2t: "ts2 t = Some xln'" and ct: "r1.cond_action_oks (ls, (ts1, m1), ws, is) t cts" shows "\mRed2 (ls, (ts2, m2), ws, is) s2'" and "\t'. t' \ t \ tbisim (ws t' = None) t' (ts1 t') m1 (thr s2' t') m2" and "r2.cond_action_oks s2' t cts" and "thr s2' t = Some xln'" unfolding s2'_def by(blast intro: FWdelay_bisimulation_final_base.cond_actions_oks_bisim_ex_\1_inv[OF FWdelay_bisimulation_final_base_flip, where bisim_wait = "flip bisim_wait", unfolded flip_simps, OF tbisim _ _ ct, OF _ ts2t ts1t])+ lemma mfinal1_inv_simulation: assumes "s1 \m s2" shows "\s2'. r2.mthr.silent_moves s2 s2' \ s1 \m s2' \ r1.final_threads s1 \ r2.final_threads s2' \ shr s2' = shr s2" proof - from \s1 \m s2\ have "finite (dom (thr s1))" by(auto dest: mbisim_finite1) moreover have "r1.final_threads s1 \ dom (thr s1)" by(auto simp add: r1.final_thread_def) ultimately have "finite (r1.final_threads s1)" by(blast intro: finite_subset) thus ?thesis using \s1 \m s2\ proof(induct A\"r1.final_threads s1" arbitrary: s1 s2 rule: finite_induct) case empty from \{} = r1.final_threads s1\[symmetric] have "\t. \ r1.final_thread s1 t" by(auto) with \s1 \m s2\ show ?case by blast next case (insert t A) define s1' where "s1' = (locks s1, ((thr s1)(t := None), shr s1), wset s1, interrupts s1)" define s2' where "s2' = (locks s2, ((thr s2)(t := None), shr s2), wset s2, interrupts s2)" from \t \ A\ \insert t A = r1.final_threads s1\ have "A = r1.final_threads s1'" unfolding s1'_def by(auto simp add: r1.final_thread_def r1.final_threads_def) moreover from \insert t A = r1.final_threads s1\ have "r1.final_thread s1 t" by auto hence "wset s1 t = None" by(auto simp add: r1.final_thread_def) with \s1 \m s2\ have "s1' \m s2'" unfolding s1'_def s2'_def by(auto simp add: mbisim_def intro: tbisim_NoneI intro!: wset_thread_okI dest: wset_thread_okD split: if_split_asm) ultimately have "\s2''. r2.mthr.silent_moves s2' s2'' \ s1' \m s2'' \ r1.final_threads s1' \ r2.final_threads s2'' \ shr s2'' = shr s2'" by(rule insert) then obtain s2'' where reds: "r2.mthr.silent_moves s2' s2''" and "s1' \m s2''" and fin: "\t. r1.final_thread s1' t \ r2.final_thread s2'' t" and "shr s2'' = shr s2'" by blast have "thr s2' t = None" unfolding s2'_def by simp with \r2.mthr.silent_moves s2' s2''\ - have "r2.mthr.silent_moves (locks s2', (thr s2'(t \ the (thr s2 t)), shr s2'), wset s2', interrupts s2') - (locks s2'', (thr s2''(t \ the (thr s2 t)), shr s2''), wset s2'', interrupts s2'')" + have "r2.mthr.silent_moves (locks s2', ((thr s2')(t \ the (thr s2 t)), shr s2'), wset s2', interrupts s2') + (locks s2'', ((thr s2'')(t \ the (thr s2 t)), shr s2''), wset s2'', interrupts s2'')" by(rule r2.\mRedT_add_thread_inv) - also let ?s2'' = "(locks s2, (thr s2''(t \ the (thr s2 t)), shr s2), wset s2, interrupts s2)" + also let ?s2'' = "(locks s2, ((thr s2'')(t \ the (thr s2 t)), shr s2), wset s2, interrupts s2)" from \shr s2'' = shr s2'\ \s1' \m s2''\ \s1 \m s2\ - have "(locks s2'', (thr s2''(t \ the (thr s2 t)), shr s2''), wset s2'', interrupts s2'') = ?s2''" + have "(locks s2'', ((thr s2'')(t \ the (thr s2 t)), shr s2''), wset s2'', interrupts s2'') = ?s2''" unfolding s2'_def s1'_def by(simp add: mbisim_def) also (back_subst) from \s1 \m s2\ have "dom (thr s1) = dom (thr s2)" by(rule mbisim_dom_eq) with \r1.final_thread s1 t\ have "t \ dom (thr s2)" by(auto simp add: r1.final_thread_def) then obtain x2 ln where tst2: "thr s2 t = \(x2, ln)\" by auto - hence "(locks s2', (thr s2'(t \ the (thr s2 t)), shr s2'), wset s2', interrupts s2') = s2" + hence "(locks s2', ((thr s2')(t \ the (thr s2 t)), shr s2'), wset s2', interrupts s2') = s2" unfolding s2'_def by(cases s2)(auto intro!: ext) also from \s1 \m s2\ tst2 obtain x1 where tst1: "thr s1 t = \(x1, ln)\" and bisim: "t \ (x1, shr s1) \ (x2, shr s2)" by(auto dest: mbisim_thrD2) from \shr s2'' = shr s2'\ have "shr ?s2'' = shr s2" by(simp add: s2'_def) from \r1.final_thread s1 t\ tst1 have final: "final1 x1" "ln = no_wait_locks" "wset s1 t = None" by(auto simp add: r1.final_thread_def) with final1_simulation[OF bisim] \shr ?s2'' = shr s2\ obtain x2' m2' where red: "r2.silent_moves t (x2, shr ?s2'') (x2', m2')" and bisim': "t \ (x1, shr s1) \ (x2', m2')" and "final2 x2'" by auto from \wset s1 t = None\ \s1 \m s2\ have "wset s2 t = None" by(simp add: mbisim_def) with bisim r2.silent_moves_into_RedT_\_inv[OF red] tst2 \ln = no_wait_locks\ have "r2.mthr.silent_moves ?s2'' (redT_upd_\ ?s2'' t x2' m2')" unfolding s2'_def by auto also (rtranclp_trans) from bisim r2.red_rtrancl_\_heapD_inv[OF red] have "m2' = shr s2" by auto hence "s1 \m (redT_upd_\ ?s2'' t x2' m2')" using \s1' \m s2''\ \s1 \m s2\ tst1 tst2 \shr ?s2'' = shr s2\ bisim' \shr s2'' = shr s2'\ \wset s2 t = None\ unfolding s1'_def s2'_def by(auto simp add: mbisim_def redT_updLns_def split: if_split_asm intro: tbisim_SomeI) moreover { fix t' assume "r1.final_thread s1 t'" with fin[of t'] \final2 x2'\ tst2 \ln = no_wait_locks\ \wset s2 t = None\ \s1' \m s2''\ \s1 \m s2\ have "r2.final_thread (redT_upd_\ ?s2'' t x2' m2') t'" unfolding s1'_def by(fastforce split: if_split_asm simp add: r2.final_thread_def r1.final_thread_def redT_updLns_def finfun_Diag_const2 o_def mbisim_def) } moreover have "shr (redT_upd_\ ?s2'' t x2' m2') = shr s2" using \m2' = shr s2\ by simp ultimately show ?case by blast qed qed lemma mfinal2_inv_simulation: "s1 \m s2 \ \s1'. r1.mthr.silent_moves s1 s1' \ s1' \m s2 \ r2.final_threads s2 \ r1.final_threads s1' \ shr s1' = shr s1" using FWdelay_bisimulation_final_base.mfinal1_inv_simulation[OF FWdelay_bisimulation_final_base_flip, where bisim_wait="flip bisim_wait"] by(unfold flip_simps) lemma mfinal1_simulation: assumes "s1 \m s2" and "r1.mfinal s1" shows "\s2'. r2.mthr.silent_moves s2 s2' \ s1 \m s2' \ r2.mfinal s2' \ shr s2' = shr s2" proof - from mfinal1_inv_simulation[OF \s1 \m s2\] obtain s2' where 1: "r2.mthr.silent_moves s2 s2'" "s1 \m s2'" "shr s2' = shr s2" and fin: "\t. r1.final_thread s1 t \ r2.final_thread s2' t" by blast have "r2.mfinal s2'" proof(rule r2.mfinalI) fix t x2 ln assume "thr s2' t = \(x2, ln)\" with \s1 \m s2'\ obtain x1 where "thr s1 t = \(x1, ln)\" "t \ (x1, shr s1) \ (x2, shr s2')" by(auto dest: mbisim_thrD2) from \thr s1 t = \(x1, ln)\\ \r1.mfinal s1\ have "r1.final_thread s1 t" by(auto elim!: r1.mfinalE simp add: r1.final_thread_def) hence "r2.final_thread s2' t" by(rule fin) thus "final2 x2 \ ln = no_wait_locks \ wset s2' t = None" using \thr s2' t = \(x2, ln)\\ by(auto simp add: r2.final_thread_def) qed with 1 show ?thesis by blast qed lemma mfinal2_simulation: "\ s1 \m s2; r2.mfinal s2 \ \ \s1'. r1.mthr.silent_moves s1 s1' \ s1' \m s2 \ r1.mfinal s1' \ shr s1' = shr s1" using FWdelay_bisimulation_final_base.mfinal1_simulation[OF FWdelay_bisimulation_final_base_flip, where bisim_wait = "flip bisim_wait"] by(unfold flip_simps) end locale FWdelay_bisimulation_obs = FWdelay_bisimulation_final_base _ _ _ _ _ _ _ \move1 \move2 for \move1 :: "('l,'t,'x1,'m1,'w, 'o) \moves" and \move2 :: "('l,'t,'x2,'m2,'w, 'o) \moves" + assumes delay_bisimulation_obs_locale: "delay_bisimulation_obs (r1 t) (r2 t) (bisim t) (ta_bisim bisim) \move1 \move2" and bisim_inv_red_other: "\ t' \ (x, m1) \ (xx, m2); t \ (x1, m1) \ (x2, m2); r1.silent_moves t (x1, m1) (x1', m1); t \ (x1', m1) -1-ta1\ (x1'', m1'); \ \move1 (x1', m1) ta1 (x1'', m1'); r2.silent_moves t (x2, m2) (x2', m2); t \ (x2', m2) -2-ta2\ (x2'', m2'); \ \move2 (x2', m2) ta2 (x2'', m2'); t \ (x1'', m1') \ (x2'', m2'); ta_bisim bisim ta1 ta2 \ \ t' \ (x, m1') \ (xx, m2')" and bisim_waitI: "\ t \ (x1, m1) \ (x2, m2); r1.silent_moves t (x1, m1) (x1', m1); t \ (x1', m1) -1-ta1\ (x1'', m1'); \ \move1 (x1', m1) ta1 (x1'', m1'); r2.silent_moves t (x2, m2) (x2', m2); t \ (x2', m2) -2-ta2\ (x2'', m2'); \ \move2 (x2', m2) ta2 (x2'', m2'); t \ (x1'', m1') \ (x2'', m2'); ta_bisim bisim ta1 ta2; Suspend w \ set \ta1\\<^bsub>w\<^esub>; Suspend w \ set \ta2\\<^bsub>w\<^esub> \ \ x1'' \w x2''" and simulation_Wakeup1: "\ t \ (x1, m1) \ (x2, m2); x1 \w x2; t \ (x1, m1) -1-ta1\ (x1', m1'); Notified \ set \ta1\\<^bsub>w\<^esub> \ WokenUp \ set \ta1\\<^bsub>w\<^esub> \ \ \ta2 x2' m2'. t \ (x2, m2) -2-ta2\ (x2', m2') \ t \ (x1', m1') \ (x2', m2') \ ta_bisim bisim ta1 ta2" and simulation_Wakeup2: "\ t \ (x1, m1) \ (x2, m2); x1 \w x2; t \ (x2, m2) -2-ta2\ (x2', m2'); Notified \ set \ta2\\<^bsub>w\<^esub> \ WokenUp \ set \ta2\\<^bsub>w\<^esub> \ \ \ta1 x1' m1'. t \ (x1, m1) -1-ta1\ (x1', m1') \ t \ (x1', m1') \ (x2', m2') \ ta_bisim bisim ta1 ta2" and ex_final1_conv_ex_final2: "(\x1. final1 x1) \ (\x2. final2 x2)" sublocale FWdelay_bisimulation_obs < delay_bisimulation_obs "r1 t" "r2 t" "bisim t" "ta_bisim bisim" \move1 \move2 for t by(rule delay_bisimulation_obs_locale) context FWdelay_bisimulation_obs begin lemma FWdelay_bisimulation_obs_flip: "FWdelay_bisimulation_obs final2 r2 final1 r1 (\t. flip (bisim t)) (flip bisim_wait) \move2 \move1" apply(rule FWdelay_bisimulation_obs.intro) apply(rule FWdelay_bisimulation_final_base_flip) apply(rule FWdelay_bisimulation_obs_axioms.intro) apply(unfold flip_simps) apply(rule delay_bisimulation_obs_axioms) apply(erule (9) bisim_inv_red_other) apply(erule (10) bisim_waitI) apply(erule (3) simulation_Wakeup2) apply(erule (3) simulation_Wakeup1) apply(rule ex_final1_conv_ex_final2[symmetric]) done end lemma FWdelay_bisimulation_obs_flip_simps [flip_simps]: "FWdelay_bisimulation_obs final2 r2 final1 r1 (\t. flip (bisim t)) (flip bisim_wait) \move2 \move1 = FWdelay_bisimulation_obs final1 r1 final2 r2 bisim bisim_wait \move1 \move2" by(auto dest: FWdelay_bisimulation_obs.FWdelay_bisimulation_obs_flip simp only: flip_flip) context FWdelay_bisimulation_obs begin lemma mbisim_redT_upd: fixes s1 t ta1 x1' m1' s2 ta2 x2' m2' ln assumes s1': "redT_upd s1 t ta1 x1' m1' s1'" and s2': "redT_upd s2 t ta2 x2' m2' s2'" and [simp]: "wset s1 = wset s2" "locks s1 = locks s2" and wset: "wset s1' = wset s2'" and interrupts: "interrupts s1' = interrupts s2'" and fin1: "finite (dom (thr s1))" and wsts: "wset_thread_ok (wset s1) (thr s1)" and tst: "thr s1 t = \(x1, ln)\" and tst': "thr s2 t = \(x2, ln)\" and aoe1: "r1.actions_ok s1 t ta1" and aoe2: "r2.actions_ok s2 t ta2" and tasim: "ta_bisim bisim ta1 ta2" and bisim': "t \ (x1', m1') \ (x2', m2')" and bisimw: "wset s1' t = None \ x1' \w x2'" and \red1: "r1.silent_moves t (x1'', shr s1) (x1, shr s1)" and red1: "t \ (x1, shr s1) -1-ta1\ (x1', m1')" and \red2: "r2.silent_moves t (x2'', shr s2) (x2, shr s2)" and red2: "t \ (x2, shr s2) -2-ta2\ (x2', m2')" and bisim: "t \ (x1'', shr s1) \ (x2'', shr s2)" and \1: "\ \move1 (x1, shr s1) ta1 (x1', m1')" and \2: "\ \move2 (x2, shr s2) ta2 (x2', m2')" and tbisim: "\t'. t \ t' \ tbisim (wset s1 t' = None) t' (thr s1 t') (shr s1) (thr s2 t') (shr s2)" shows "s1' \m s2'" proof(rule mbisimI) from fin1 s1' show "finite (dom (thr s1'))" by(auto simp add: redT_updTs_finite_dom_inv) next from tasim s1' s2' show "locks s1' = locks s2'" by(auto simp add: redT_updLs_def o_def ta_bisim_def) next from wset show "wset s1' = wset s2'" . next from interrupts show "interrupts s1' = interrupts s2'" . next from wsts s1' s2' wset show "wset_thread_ok (wset s1') (thr s1')" by(fastforce intro!: wset_thread_okI split: if_split_asm dest: redT_updTs_None wset_thread_okD redT_updWs_None_implies_None) next fix T assume "thr s1' T = None" moreover with tst s1' have [simp]: "t \ T" by auto from tbisim[OF this] have "(thr s1 T = None) = (thr s2 T = None)" by(auto simp add: tbisim_def) hence "(redT_updTs (thr s1) \ta1\\<^bsub>t\<^esub> T = None) = (redT_updTs (thr s2) \ta2\\<^bsub>t\<^esub> T = None)" using tasim by -(rule redT_updTs_nta_bisim_inv, simp_all add: ta_bisim_def) ultimately show "thr s2' T = None" using s2' s1' by(auto split: if_split_asm) next fix T X1 LN assume tsT: "thr s1' T = \(X1, LN)\" show "\x2. thr s2' T = \(x2, LN)\ \ T \ (X1, shr s1') \ (x2, shr s2') \ (wset s2' T = None \ X1 \w x2)" proof(cases "thr s1 T") case None with tst have "t \ T" by auto with tbisim[OF this] None have tsT': "thr s2 T = None" by(simp add: tbisim_def) from None \t \ T\ tsT aoe1 s1' obtain M1 where ntset: "NewThread T X1 M1 \ set \ta1\\<^bsub>t\<^esub>" and [simp]: "LN = no_wait_locks" by(auto dest!: redT_updTs_new_thread) from ntset obtain tas1 tas1' where "\ta1\\<^bsub>t\<^esub> = tas1 @ NewThread T X1 M1 # tas1'" by(auto simp add: in_set_conv_decomp) with tasim obtain tas2 X2 M2 tas2' where "\ta2\\<^bsub>t\<^esub> = tas2 @ NewThread T X2 M2 # tas2'" "length tas2 = length tas2" "length tas1' = length tas2'" and Bisim: "T \ (X1, M1) \ (X2, M2)" by(auto simp add: list_all2_append1 list_all2_Cons1 ta_bisim_def) hence ntset': "NewThread T X2 M2 \ set \ta2\\<^bsub>t\<^esub>" by auto with tsT' \t \ T\ aoe2 s2' have "thr s2' T = \(X2, no_wait_locks)\" by(auto intro: redT_updTs_new_thread_ts) moreover from ntset' red2 have "m2' = M2" by(auto dest: r2.new_thread_memory) moreover from ntset red1 have "m1' = M1" by(auto dest: r1.new_thread_memory) moreover from wsts None have "wset s1 T = None" by(rule wset_thread_okD) ultimately show ?thesis using Bisim \t \ T\ s1' s2' by(auto simp add: redT_updWs_None_implies_None) next case (Some a) show ?thesis proof(cases "t = T") case True with tst tsT s1' have [simp]: "X1 = x1'" "LN = redT_updLns (locks s1) t ln \ta1\\<^bsub>l\<^esub>" by(auto) show ?thesis using True bisim' bisimw tasim tst tst' s1' s2' wset by(auto simp add: redT_updLns_def ta_bisim_def) next case False with Some aoe1 tsT s1' have "thr s1 T = \(X1, LN)\" by(auto dest: redT_updTs_Some) with tbisim[OF False] obtain X2 where tsT': "thr s2 T = \(X2, LN)\" and Bisim: "T \ (X1, shr s1) \ (X2, shr s2)" and bisimw: "wset s1 T = None \ X1 \w X2" by(auto simp add: tbisim_def) with aoe2 False s2' have tsT': "thr s2' T = \(X2, LN)\" by(auto simp add: redT_updTs_Some) moreover from Bisim bisim \red1 red1 \1 \red2 red2 \2 bisim' tasim have "T \ (X1, m1') \ (X2, m2')" by(rule bisim_inv_red_other) ultimately show ?thesis using False bisimw s1' s2' by(auto simp add: redT_updWs_None_implies_None) qed qed qed theorem mbisim_simulation1: assumes mbisim: "mbisim s1 s2" and "\ m\move1 s1 tl1 s1'" "r1.redT s1 tl1 s1'" shows "\s2' s2'' tl2. r2.mthr.silent_moves s2 s2' \ r2.redT s2' tl2 s2'' \ \ m\move2 s2' tl2 s2'' \ mbisim s1' s2'' \ mta_bisim tl1 tl2" proof - from assms obtain t ta1 where tl1 [simp]: "tl1 = (t, ta1)" and redT: "s1 -1-t\ta1\ s1'" and m\: "\ m\move1 s1 (t, ta1) s1'" by(cases tl1) fastforce obtain ls1 ts1 m1 ws1 is1 where [simp]: "s1 = (ls1, (ts1, m1), ws1, is1)" by(cases s1) fastforce obtain ls1' ts1' m1' ws1' is1' where [simp]: "s1' = (ls1', (ts1', m1'), ws1', is1')" by(cases s1') fastforce obtain ls2 ts2 m2 ws2 is2 where [simp]: "s2 = (ls2, (ts2, m2), ws2, is2)" by(cases s2) fastforce from mbisim have [simp]: "ls2 = ls1" "ws2 = ws1" "is2 = is1" "finite (dom ts1)" by(auto simp add: mbisim_def) from redT show ?thesis proof cases case (redT_normal x1 x1' M1') hence red: "t \ (x1, m1) -1-ta1\ (x1', M1')" and tst: "ts1 t = \(x1, no_wait_locks)\" and aoe: "r1.actions_ok s1 t ta1" and s1': "redT_upd s1 t ta1 x1' M1' s1'" by auto from mbisim tst obtain x2 where tst': "ts2 t = \(x2, no_wait_locks)\" and bisim: "t \ (x1, m1) \ (x2, m2)" by(auto dest: mbisim_thrD1) from m\ have \: "\ \move1 (x1, m1) ta1 (x1', M1')" proof(rule contrapos_nn) assume \: "\move1 (x1, m1) ta1 (x1', M1')" moreover hence [simp]: "ta1 = \" by(rule r1.silent_tl) moreover have [simp]: "M1' = m1" by(rule r1.\move_heap[OF red \, symmetric]) ultimately show "m\move1 s1 (t, ta1) s1'" using s1' tst s1' by(auto simp add: redT_updLs_def o_def intro: r1.m\move.intros elim: rtrancl3p_cases) qed show ?thesis proof(cases "ws1 t") case None note wst = this from simulation1[OF bisim red \] obtain x2' M2' x2'' M2'' ta2 where red21: "r2.silent_moves t (x2, m2) (x2', M2')" and red22: "t \ (x2', M2') -2-ta2\ (x2'', M2'')" and \2: "\ \move2 (x2', M2') ta2 (x2'', M2'')" and bisim': "t \ (x1', M1') \ (x2'', M2'')" and tasim: "ta_bisim bisim ta1 ta2" by auto let ?s2' = "redT_upd_\ s2 t x2' M2'" let ?S2' = "activate_cond_actions2 s1 ?s2' \ta2\\<^bsub>c\<^esub>" let ?s2'' = "(redT_updLs (locks ?S2') t \ta2\\<^bsub>l\<^esub>, ((redT_updTs (thr ?S2') \ta2\\<^bsub>t\<^esub>)(t \ (x2'', redT_updLns (locks ?S2') t (snd (the (thr ?S2' t))) \ta2\\<^bsub>l\<^esub>)), M2''), wset s1', interrupts s1')" from red21 tst' wst bisim have "\mRed2 s2 ?s2'" by -(rule r2.silent_moves_into_RedT_\_inv, auto) moreover from red21 bisim have [simp]: "M2' = m2" by(auto dest: r2.red_rtrancl_\_heapD_inv) from tasim have [simp]: "\ ta1 \\<^bsub>l\<^esub> = \ ta2 \\<^bsub>l\<^esub>" "\ ta1 \\<^bsub>w\<^esub> = \ ta2 \\<^bsub>w\<^esub>" "\ ta1 \\<^bsub>c\<^esub> = \ ta2 \\<^bsub>c\<^esub>" "\ ta1 \\<^bsub>i\<^esub> = \ ta2 \\<^bsub>i\<^esub>" and nta: "list_all2 (nta_bisim bisim) \ ta1 \\<^bsub>t\<^esub> \ ta2 \\<^bsub>t\<^esub>" by(auto simp add: ta_bisim_def) from mbisim have tbisim: "\t. tbisim (ws1 t = None) t (ts1 t) m1 (ts2 t) m2" by(simp add: mbisim_def) hence tbisim': "\t'. t' \ t \ tbisim (ws1 t' = None) t' (ts1 t') m1 (thr ?s2' t') m2" by(auto) from aoe have cao1: "r1.cond_action_oks (ls1, (ts1, m1), ws1, is1) t \ta2\\<^bsub>c\<^esub>" by auto from tst' have "thr ?s2' t = \(x2', no_wait_locks)\" by(auto simp add: redT_updLns_def o_def finfun_Diag_const2) from cond_actions_oks_bisim_ex_\2_inv[OF tbisim', OF _ tst this cao1] have red21': "\mRed2 ?s2' ?S2'" and tbisim'': "\t'. t' \ t \ tbisim (ws1 t' = None) t' (ts1 t') m1 (thr ?S2' t') m2" and cao2: "r2.cond_action_oks ?S2' t \ta2\\<^bsub>c\<^esub>" and tst'': "thr ?S2' t = \(x2', no_wait_locks)\" by(auto simp del: fun_upd_apply) note red21' also (rtranclp_trans) from tbisim'' tst'' tst have "\t'. ts1 t' = None \ thr ?S2' t' = None" by(force simp add: tbisim_def) from aoe thread_oks_bisim_inv[OF this nta] have "thread_oks (thr ?S2') \ta2\\<^bsub>t\<^esub>" by simp with cao2 aoe have aoe': "r2.actions_ok ?S2' t ta2" by auto with red22 tst'' s1' have "?S2' -2-t\ta2\ ?s2''" by -(rule r2.redT.redT_normal, auto) moreover from \2 have "\ m\move2 ?S2' (t, ta2) ?s2''" proof(rule contrapos_nn) assume m\: "m\move2 ?S2' (t, ta2) ?s2''" thus "\move2 (x2', M2') ta2 (x2'', M2'')" using tst'' tst' by cases auto qed moreover { note s1' moreover have "redT_upd ?S2' t ta2 x2'' M2'' ?s2''" using s1' by auto moreover have "wset s1 = wset ?S2'" "locks s1 = locks ?S2'" by simp_all moreover have "wset s1' = wset ?s2''" by simp moreover have "interrupts s1' = interrupts ?s2''" by simp moreover have "finite (dom (thr s1))" by simp moreover from mbisim have "wset_thread_ok (wset s1) (thr s1)" by(simp add: mbisim_def) moreover from tst have "thr s1 t = \(x1, no_wait_locks)\" by simp moreover note tst'' aoe aoe' tasim bisim' moreover have "wset s1' t = None \ x1' \w x2''" proof(cases "wset s1' t") case None thus ?thesis .. next case (Some w) with wst s1' obtain w' where Suspend1: "Suspend w' \ set \ta1\\<^bsub>w\<^esub>" by(auto dest: redT_updWs_None_SomeD) with tasim have Suspend2: "Suspend w' \ set \ta2\\<^bsub>w\<^esub>" by(simp add: ta_bisim_def) from bisim_waitI[OF bisim rtranclp.rtrancl_refl red \ _ _ _ bisim' tasim Suspend1 this, of x2'] red21 red22 \2 have "x1' \w x2''" by auto thus ?thesis .. qed moreover note rtranclp.rtrancl_refl moreover from red have "t \ (x1, shr s1) -1-ta1\ (x1', M1')" by simp moreover from red21 have "r2.silent_moves t (x2, shr ?S2') (x2', shr ?S2')" by simp moreover from red22 have "t \ (x2', shr ?S2') -2-ta2\ (x2'', M2'')" by simp moreover from bisim have "t \ (x1, shr s1) \ (x2, shr ?S2')" by simp moreover from \ have "\ \move1 (x1, shr s1) ta1 (x1', M1')" by simp moreover from \2 have "\ \move2 (x2', shr ?S2') ta2 (x2'', M2'')" by simp moreover from tbisim'' have "\t'. t \ t' \ tbisim (wset s1 t' = None) t' (thr s1 t') (shr s1) (thr ?S2' t') (shr ?S2')" by simp ultimately have "mbisim s1' ?s2''" by(rule mbisim_redT_upd) } ultimately show ?thesis using tasim unfolding tl1 s1' by fastforce next case (Some w) with mbisim tst tst' have "x1 \w x2" by(auto dest: mbisim_thrD1) from aoe Some have wakeup: "Notified \ set \ta1\\<^bsub>w\<^esub> \ WokenUp \ set \ta1\\<^bsub>w\<^esub>" by(auto simp add: wset_actions_ok_def split: if_split_asm) from simulation_Wakeup1[OF bisim \x1 \w x2\ red this] obtain ta2 x2' m2' where red2: "t \ (x2, m2) -2-ta2\ (x2', m2')" and bisim': "t \ (x1', M1') \ (x2', m2')" and tasim: "ta1 \m ta2" by auto let ?S2' = "activate_cond_actions2 s1 s2 \ta2\\<^bsub>c\<^esub>" let ?s2' = "(redT_updLs (locks ?S2') t \ta2\\<^bsub>l\<^esub>, ((redT_updTs (thr ?S2') \ta2\\<^bsub>t\<^esub>)(t \ (x2', redT_updLns (locks ?S2') t (snd (the (thr ?S2' t))) \ta2\\<^bsub>l\<^esub>)), m2'), wset s1', interrupts s1')" from tasim have [simp]: "\ ta1 \\<^bsub>l\<^esub> = \ ta2 \\<^bsub>l\<^esub>" "\ ta1 \\<^bsub>w\<^esub> = \ ta2 \\<^bsub>w\<^esub>" "\ ta1 \\<^bsub>c\<^esub> = \ ta2 \\<^bsub>c\<^esub>" "\ ta1 \\<^bsub>i\<^esub> = \ ta2 \\<^bsub>i\<^esub>" and nta: "list_all2 (nta_bisim bisim) \ ta1 \\<^bsub>t\<^esub> \ ta2 \\<^bsub>t\<^esub>" by(auto simp add: ta_bisim_def) from mbisim have tbisim: "\t. tbisim (ws1 t = None) t (ts1 t) m1 (ts2 t) m2" by(simp add: mbisim_def) hence tbisim': "\t'. t' \ t \ tbisim (ws1 t' = None) t' (ts1 t') m1 (thr s2 t') m2" by(auto) from aoe have cao1: "r1.cond_action_oks (ls1, (ts1, m1), ws1, is1) t \ta2\\<^bsub>c\<^esub>" by auto from tst' have "thr s2 t = \(x2, no_wait_locks)\" by(auto simp add: redT_updLns_def o_def finfun_Diag_const2) from cond_actions_oks_bisim_ex_\2_inv[OF tbisim', OF _ tst this cao1] have red21': "\mRed2 s2 ?S2'" and tbisim'': "\t'. t' \ t \ tbisim (ws1 t' = None) t' (ts1 t') m1 (thr ?S2' t') m2" and cao2: "r2.cond_action_oks ?S2' t \ta2\\<^bsub>c\<^esub>" and tst'': "thr ?S2' t = \(x2, no_wait_locks)\" by(auto simp del: fun_upd_apply) note red21' moreover from tbisim'' tst'' tst have "\t'. ts1 t' = None \ thr ?S2' t' = None" by(force simp add: tbisim_def) from aoe thread_oks_bisim_inv[OF this nta] have "thread_oks (thr ?S2') \ta2\\<^bsub>t\<^esub>" by simp with cao2 aoe have aoe': "r2.actions_ok ?S2' t ta2" by auto with red2 tst'' s1' tasim have "?S2' -2-t\ta2\ ?s2'" by -(rule r2.redT_normal, auto simp add: ta_bisim_def) moreover from wakeup tasim have \2: "\ \move2 (x2, m2) ta2 (x2', m2')" by(auto dest: r2.silent_tl) hence "\ m\move2 ?S2' (t, ta2) ?s2'" proof(rule contrapos_nn) assume m\: "m\move2 ?S2' (t, ta2) ?s2'" thus "\move2 (x2, m2) ta2 (x2', m2')" using tst'' tst' by cases auto qed moreover { note s1' moreover have "redT_upd ?S2' t ta2 x2' m2' ?s2'" using s1' tasim by(auto simp add: ta_bisim_def) moreover have "wset s1 = wset ?S2'" "locks s1 = locks ?S2'" by simp_all moreover have "wset s1' = wset ?s2'" by simp moreover have "interrupts s1' = interrupts ?s2'" by simp moreover have "finite (dom (thr s1))" by simp moreover from mbisim have "wset_thread_ok (wset s1) (thr s1)" by(rule mbisim_wset_thread_ok1) moreover from tst have "thr s1 t = \(x1, no_wait_locks)\" by simp moreover from tst'' have "thr ?S2' t = \(x2, no_wait_locks)\" by simp moreover note aoe aoe' tasim bisim' moreover have "wset s1' t = None \ x1' \w x2'" proof(cases "wset s1' t") case None thus ?thesis .. next case (Some w') with redT_updWs_WokenUp_SuspendD[OF _ wakeup, of t "wset s1" "wset s1'" w'] s1' obtain w' where Suspend1: "Suspend w' \ set \ta1\\<^bsub>w\<^esub>" by(auto) with tasim have Suspend2: "Suspend w' \ set \ta2\\<^bsub>w\<^esub>" by(simp add: ta_bisim_def) with bisim rtranclp.rtrancl_refl red \ rtranclp.rtrancl_refl red2 \2 bisim' tasim Suspend1 have "x1' \w x2'" by(rule bisim_waitI) thus ?thesis .. qed moreover note rtranclp.rtrancl_refl moreover from red have "t \ (x1, shr s1) -1-ta1\ (x1', M1')" by simp moreover note rtranclp.rtrancl_refl moreover from red2 have "t \ (x2, shr ?S2') -2-ta2\ (x2', m2')" by simp moreover from bisim have "t \ (x1, shr s1) \ (x2, shr ?S2')" by simp moreover from \ have "\ \move1 (x1, shr s1) ta1 (x1', M1')" by simp moreover from \2 have "\ \move2 (x2, shr ?S2') ta2 (x2', m2')" by simp moreover from tbisim'' have "\t'. t \ t' \ tbisim (wset s1 t' = None) t' (thr s1 t') (shr s1) (thr ?S2' t') (shr ?S2')" by simp ultimately have "s1' \m ?s2'" by(rule mbisim_redT_upd) } moreover from tasim have "tl1 \T (t, ta2)" by simp ultimately show ?thesis unfolding s1' by blast qed next case (redT_acquire x1 n ln) hence [simp]: "ta1 = (K$ [], [], [], [], [], convert_RA ln)" and tst: "thr s1 t = \(x1, ln)\" and wst: "\ waiting (wset s1 t)" and maa: "may_acquire_all (locks s1) t ln" and ln: "0 < ln $ n" and s1': "s1' = (acquire_all ls1 t ln, (ts1(t \ (x1, no_wait_locks)), m1), ws1, is1)" by auto from tst mbisim obtain x2 where tst': "ts2 t = \(x2, ln)\" and bisim: "t \ (x1, m1) \ (x2, m2)" by(auto dest: mbisim_thrD1) let ?s2' = "(acquire_all ls1 t ln, (ts2(t \ (x2, no_wait_locks)), m2), ws1, is1)" from tst' wst maa ln have "s2 -2-t\(K$ [], [], [], [], [], convert_RA ln)\ ?s2'" by-(rule r2.redT.redT_acquire, auto) moreover from tst' ln have "\ m\move2 s2 (t, (K$ [], [], [], [], [], convert_RA ln)) ?s2'" by(auto simp add: acquire_all_def fun_eq_iff elim!: r2.m\move.cases) moreover have "mbisim s1' ?s2'" proof(rule mbisimI) from s1' show "locks s1' = locks ?s2'" by auto next from s1' show "wset s1' = wset ?s2'" by auto next from s1' show "interrupts s1' = interrupts ?s2'" by auto next fix t' assume "thr s1' t' = None" with s1' have "thr s1 t' = None" by(auto split: if_split_asm) with mbisim_thrNone_eq[OF mbisim] have "ts2 t' = None" by simp with tst' show "thr ?s2' t' = None" by auto next fix t' X1 LN assume ts't: "thr s1' t' = \(X1, LN)\" show "\x2. thr ?s2' t' = \(x2, LN)\ \ t' \ (X1, shr s1') \ (x2, shr ?s2') \ (wset ?s2' t' = None \ X1 \w x2)" proof(cases "t' = t") case True with s1' tst ts't have [simp]: "X1 = x1" "LN = no_wait_locks" by simp_all with mbisim_thrD1[OF mbisim tst] bisim tst tst' True s1' wst show ?thesis by(auto) next case False with ts't s1' have "ts1 t' = \(X1, LN)\" by auto with mbisim obtain X2 where "ts2 t' = \(X2, LN)\" "t' \ (X1, m1) \ (X2, m2)" "wset ?s2' t' = None \ X1 \w X2" by(auto dest: mbisim_thrD1) with False s1' show ?thesis by auto qed next from s1' show "finite (dom (thr s1'))" by auto next from mbisim_wset_thread_ok1[OF mbisim] show "wset_thread_ok (wset s1') (thr s1')" using s1' by(auto intro: wset_thread_ok_upd) qed moreover have "(t, K$ [], [], [], [], [], convert_RA ln) \T (t, K$ [], [], [], [], [], convert_RA ln)" by(simp add: ta_bisim_def) ultimately show ?thesis by fastforce qed qed theorem mbisim_simulation2: "\ mbisim s1 s2; r2.redT s2 tl2 s2'; \ m\move2 s2 tl2 s2' \ \ \s1' s1'' tl1. r1.mthr.silent_moves s1 s1' \ r1.redT s1' tl1 s1'' \ \ m\move1 s1' tl1 s1'' \ mbisim s1'' s2' \ mta_bisim tl1 tl2" using FWdelay_bisimulation_obs.mbisim_simulation1[OF FWdelay_bisimulation_obs_flip] unfolding flip_simps . end locale FWdelay_bisimulation_diverge = FWdelay_bisimulation_obs _ _ _ _ _ _ _ \move1 \move2 for \move1 :: "('l,'t,'x1,'m1,'w,'o) \moves" and \move2 :: "('l,'t,'x2,'m2,'w,'o) \moves" + assumes delay_bisimulation_diverge_locale: "delay_bisimulation_diverge (r1 t) (r2 t) (bisim t) (ta_bisim bisim) \move1 \move2" sublocale FWdelay_bisimulation_diverge < delay_bisimulation_diverge "r1 t" "r2 t" "bisim t" "ta_bisim bisim" \move1 \move2 for t by(rule delay_bisimulation_diverge_locale) context FWdelay_bisimulation_diverge begin lemma FWdelay_bisimulation_diverge_flip: "FWdelay_bisimulation_diverge final2 r2 final1 r1 (\t. flip (bisim t)) (flip bisim_wait) \move2 \move1" apply(rule FWdelay_bisimulation_diverge.intro) apply(rule FWdelay_bisimulation_obs_flip) apply(rule FWdelay_bisimulation_diverge_axioms.intro) apply(unfold flip_simps) apply(rule delay_bisimulation_diverge_axioms) done end lemma FWdelay_bisimulation_diverge_flip_simps [flip_simps]: "FWdelay_bisimulation_diverge final2 r2 final1 r1 (\t. flip (bisim t)) (flip bisim_wait) \move2 \move1 = FWdelay_bisimulation_diverge final1 r1 final2 r2 bisim bisim_wait \move1 \move2" by(auto dest: FWdelay_bisimulation_diverge.FWdelay_bisimulation_diverge_flip simp only: flip_flip) context FWdelay_bisimulation_diverge begin lemma bisim_inv1: assumes bisim: "t \ s1 \ s2" and red: "t \ s1 -1-ta1\ s1'" obtains s2' where "t \ s1' \ s2'" proof(atomize_elim) show "\s2'. t \ s1' \ s2'" proof(cases "\move1 s1 ta1 s1'") case True with red have "r1.silent_move t s1 s1'" by auto from simulation_silent1[OF bisim this] show ?thesis by auto next case False from simulation1[OF bisim red False] show ?thesis by auto qed qed lemma bisim_inv2: assumes "t \ s1 \ s2" "t \ s2 -2-ta2\ s2'" obtains s1' where "t \ s1' \ s2'" using assms FWdelay_bisimulation_diverge.bisim_inv1[OF FWdelay_bisimulation_diverge_flip] unfolding flip_simps by blast lemma bisim_inv: "bisim_inv" by(blast intro!: bisim_invI elim: bisim_inv1 bisim_inv2) lemma bisim_inv_\s1: assumes "t \ s1 \ s2" and "r1.silent_moves t s1 s1'" obtains s2' where "t \ s1' \ s2'" using assms by(rule bisim_inv_\s1_inv[OF bisim_inv]) lemma bisim_inv_\s2: assumes "t \ s1 \ s2" and "r2.silent_moves t s2 s2'" obtains s1' where "t \ s1' \ s2'" using assms by(rule bisim_inv_\s2_inv[OF bisim_inv]) lemma red1_rtrancl_\_into_RedT_\: assumes "r1.silent_moves t (x1, shr s1) (x1', m1')" "t \ (x1, shr s1) \ (x2, m2)" and "thr s1 t = \(x1, no_wait_locks)\" "wset s1 t = None" shows "\mRed1 s1 (redT_upd_\ s1 t x1' m1')" using assms by(blast intro: r1.silent_moves_into_RedT_\_inv) lemma red2_rtrancl_\_into_RedT_\: assumes "r2.silent_moves t (x2, shr s2) (x2', m2')" and "t \ (x1, m1) \ (x2, shr s2)" "thr s2 t = \(x2, no_wait_locks)\" "wset s2 t = None" shows "\mRed2 s2 (redT_upd_\ s2 t x2' m2')" using assms by(blast intro: r2.silent_moves_into_RedT_\_inv) lemma red1_rtrancl_\_heapD: "\ r1.silent_moves t s1 s1'; t \ s1 \ s2 \ \ snd s1' = snd s1" by(blast intro: r1.red_rtrancl_\_heapD_inv) lemma red2_rtrancl_\_heapD: "\ r2.silent_moves t s2 s2'; t \ s1 \ s2 \ \ snd s2' = snd s2" by(blast intro: r2.red_rtrancl_\_heapD_inv) lemma mbisim_simulation_silent1: assumes m\': "r1.mthr.silent_move s1 s1'" and mbisim: "s1 \m s2" shows "\s2'. r2.mthr.silent_moves s2 s2' \ s1' \m s2'" proof - from m\' obtain tl1 where m\: "m\move1 s1 tl1 s1'" "r1.redT s1 tl1 s1'" by auto obtain ls1 ts1 m1 ws1 is1 where [simp]: "s1 = (ls1, (ts1, m1), ws1, is1)" by(cases s1) fastforce obtain ls1' ts1' m1' ws1' is1' where [simp]: "s1' = (ls1', (ts1', m1'), ws1', is1')" by(cases s1') fastforce obtain ls2 ts2 m2 ws2 is2 where [simp]: "s2 = (ls2, (ts2, m2), ws2, is2)" by(cases s2) fastforce from m\ obtain t where "tl1 = (t, \)" by(auto elim!: r1.m\move.cases dest: r1.silent_tl) with m\ have m\: "m\move1 s1 (t, \) s1'" and redT1: "s1 -1-t\\\ s1'" by simp_all from m\ obtain x x' ln' where tst: "ts1 t = \(x, no_wait_locks)\" and ts't: "ts1' t = \(x', ln')\" and \: "\move1 (x, m1) \ (x', m1')" by(fastforce elim: r1.m\move.cases) from mbisim have [simp]: "ls2 = ls1" "ws2 = ws1" "is2 = is1" "finite (dom ts1)" by(auto simp add: mbisim_def) from redT1 show ?thesis proof cases case (redT_normal x1 x1' M') with tst ts't have [simp]: "x = x1" "x' = x1'" and red: "t \ (x1, m1) -1-\\ (x1', M')" and tst: "thr s1 t = \(x1, no_wait_locks)\" and wst: "wset s1 t = None" and s1': "redT_upd s1 t \ x1' M' s1'" by(auto) from s1' tst have [simp]: "ls1' = ls1" "ws1' = ws1" "is1' = is1" "M' = m1'" "ts1' = ts1(t \ (x1', no_wait_locks))" by(auto simp add: redT_updLs_def redT_updLns_def o_def redT_updWs_def elim!: rtrancl3p_cases) from mbisim tst obtain x2 where tst': "ts2 t = \(x2, no_wait_locks)\" and bisim: "t \ (x1, m1) \ (x2, m2)" by(auto dest: mbisim_thrD1) from r1.\move_heap[OF red] \ have [simp]: "m1 = M'" by simp from red \ have "r1.silent_move t (x1, m1) (x1', M')" by auto from simulation_silent1[OF bisim this] obtain x2' m2' where red: "r2.silent_moves t (x2, m2) (x2', m2')" and bisim': "t \ (x1', m1) \ (x2', m2')" by auto from red bisim have [simp]: "m2' = m2" by(auto dest: red2_rtrancl_\_heapD) let ?s2' = "redT_upd_\ s2 t x2' m2'" from red tst' wst bisim have "\mRed2 s2 ?s2'" by -(rule red2_rtrancl_\_into_RedT_\, auto) moreover have "mbisim s1' ?s2'" proof(rule mbisimI) show "locks s1' = locks ?s2'" "wset s1' = wset ?s2'" "interrupts s1' = interrupts ?s2'" by auto next fix t' assume "thr s1' t' = None" hence "ts1 t' = None" by(auto split: if_split_asm) with mbisim_thrNone_eq[OF mbisim] have "ts2 t' = None" by simp with tst' show "thr ?s2' t' = None" by auto next fix t' X1 LN assume ts't': "thr s1' t' = \(X1, LN)\" show "\x2. thr ?s2' t' = \(x2, LN)\ \ t' \ (X1, shr s1') \ (x2, shr ?s2') \ (wset ?s2' t' = None \ X1 \w x2)" proof(cases "t' = t") case True note this[simp] with s1' tst ts't' have [simp]: "X1 = x1'" "LN = no_wait_locks" by(simp_all)(auto simp add: redT_updLns_def o_def finfun_Diag_const2) with bisim' tst' wst show ?thesis by(auto simp add: redT_updLns_def o_def finfun_Diag_const2) next case False with ts't' have "ts1 t' = \(X1, LN)\" by auto with mbisim obtain X2 where "ts2 t' = \(X2, LN)\" "t' \ (X1, m1) \ (X2, m2)" "ws1 t' = None \ X1 \w X2" by(auto dest: mbisim_thrD1) with False show ?thesis by auto qed next show "finite (dom (thr s1'))" by simp next from mbisim_wset_thread_ok1[OF mbisim] show "wset_thread_ok (wset s1') (thr s1')" by(auto intro: wset_thread_ok_upd) qed ultimately show ?thesis by(auto) next case redT_acquire with tst have False by auto thus ?thesis .. qed qed lemma mbisim_simulation_silent2: "\ mbisim s1 s2; r2.mthr.silent_move s2 s2' \ \ \s1'. r1.mthr.silent_moves s1 s1' \ mbisim s1' s2'" using FWdelay_bisimulation_diverge.mbisim_simulation_silent1[OF FWdelay_bisimulation_diverge_flip] unfolding flip_simps . lemma mbisim_simulation1': assumes mbisim: "mbisim s1 s2" and "\ m\move1 s1 tl1 s1'" "r1.redT s1 tl1 s1'" shows "\s2' s2'' tl2. r2.mthr.silent_moves s2 s2' \ r2.redT s2' tl2 s2'' \ \ m\move2 s2' tl2 s2'' \ mbisim s1' s2'' \ mta_bisim tl1 tl2" using mbisim_simulation1 assms . lemma mbisim_simulation2': "\ mbisim s1 s2; r2.redT s2 tl2 s2'; \ m\move2 s2 tl2 s2' \ \ \s1' s1'' tl1. r1.mthr.silent_moves s1 s1' \ r1.redT s1' tl1 s1'' \ \ m\move1 s1' tl1 s1'' \ mbisim s1'' s2' \ mta_bisim tl1 tl2" using FWdelay_bisimulation_diverge.mbisim_simulation1'[OF FWdelay_bisimulation_diverge_flip] unfolding flip_simps . lemma m\diverge_simulation1: assumes "s1 \m s2" and "r1.mthr.\diverge s1" shows "r2.mthr.\diverge s2" proof - from \s1 \m s2\ have "finite (dom (thr s1))" by(rule mbisim_finite1)+ from r1.\diverge_\mredTD[OF \r1.mthr.\diverge s1\ this] obtain t x where "thr s1 t = \(x, no_wait_locks)\" "wset s1 t = None" "r1.\diverge t (x, shr s1)" by blast from \s1 \m s2\ \thr s1 t = \(x, no_wait_locks)\\ obtain x' where "thr s2 t = \(x', no_wait_locks)\" "t \ (x, shr s1) \ (x', shr s2)" by(auto dest: mbisim_thrD1) from \s1 \m s2\ \wset s1 t = None\ have "wset s2 t = None" by(simp add: mbisim_def) from \t \ (x, shr s1) \ (x', shr s2)\ \r1.\diverge t (x, shr s1)\ have "r2.\diverge t (x', shr s2)" by(simp add: \diverge_bisim_inv) thus ?thesis using \thr s2 t = \(x', no_wait_locks)\\ \wset s2 t = None\ by(rule r2.\diverge_into_\mredT) qed lemma \diverge_mbisim_inv: "s1 \m s2 \ r1.mthr.\diverge s1 \ r2.mthr.\diverge s2" apply(rule iffI) apply(erule (1) m\diverge_simulation1) by(rule FWdelay_bisimulation_diverge.m\diverge_simulation1[OF FWdelay_bisimulation_diverge_flip, unfolded flip_simps]) lemma mbisim_delay_bisimulation: "delay_bisimulation_diverge r1.redT r2.redT mbisim mta_bisim m\move1 m\move2" apply(unfold_locales) apply(rule mbisim_simulation1 mbisim_simulation2 mbisim_simulation_silent1 mbisim_simulation_silent2 \diverge_mbisim_inv|assumption)+ done theorem mdelay_bisimulation_final_base: "delay_bisimulation_final_base r1.redT r2.redT mbisim m\move1 m\move2 r1.mfinal r2.mfinal" apply(unfold_locales) apply(blast dest: mfinal1_simulation mfinal2_simulation)+ done end sublocale FWdelay_bisimulation_diverge < mthr: delay_bisimulation_diverge r1.redT r2.redT mbisim mta_bisim m\move1 m\move2 by(rule mbisim_delay_bisimulation) sublocale FWdelay_bisimulation_diverge < mthr: delay_bisimulation_final_base r1.redT r2.redT mbisim mta_bisim m\move1 m\move2 r1.mfinal r2.mfinal by(rule mdelay_bisimulation_final_base) context FWdelay_bisimulation_diverge begin lemma mthr_delay_bisimulation_diverge_final: "delay_bisimulation_diverge_final r1.redT r2.redT mbisim mta_bisim m\move1 m\move2 r1.mfinal r2.mfinal" by(unfold_locales) end sublocale FWdelay_bisimulation_diverge < mthr: delay_bisimulation_diverge_final r1.redT r2.redT mbisim mta_bisim m\move1 m\move2 r1.mfinal r2.mfinal by(rule mthr_delay_bisimulation_diverge_final) subsection \Strong bisimulation as corollary\ locale FWbisimulation = FWbisimulation_base _ _ _ r2 convert_RA bisim "\x1 x2. True" + r1: multithreaded final1 r1 convert_RA + r2: multithreaded final2 r2 convert_RA for r2 :: "('l,'t,'x2,'m2,'w,'o) semantics" ("_ \ _ -2-_\ _" [50,0,0,50] 80) and convert_RA :: "'l released_locks \ 'o list" and bisim :: "'t \ ('x1 \ 'm1, 'x2 \ 'm2) bisim" ("_ \ _/ \ _" [50, 50, 50] 60) + assumes bisimulation_locale: "bisimulation (r1 t) (r2 t) (bisim t) (ta_bisim bisim)" and bisim_final: "t \ (x1, m1) \ (x2, m2) \ final1 x1 \ final2 x2" and bisim_inv_red_other: "\ t' \ (x, m1) \ (xx, m2); t \ (x1, m1) \ (x2, m2); t \ (x1, m1) -1-ta1\ (x1', m1'); t \ (x2, m2) -2-ta2\ (x2', m2'); t \ (x1', m1') \ (x2', m2'); ta_bisim bisim ta1 ta2 \ \ t' \ (x, m1') \ (xx, m2')" and ex_final1_conv_ex_final2: "(\x1. final1 x1) \ (\x2. final2 x2)" sublocale FWbisimulation < bisim?: bisimulation "r1 t" "r2 t" "bisim t" "ta_bisim bisim" for t by(rule bisimulation_locale) sublocale FWbisimulation < bisim_diverge?: FWdelay_bisimulation_diverge final1 r1 final2 r2 convert_RA bisim "\x1 x2. True" "\s ta s'. False" "\s ta s'. False" proof - interpret biw: bisimulation_into_delay "r1 t" "r2 t" "bisim t" "ta_bisim bisim" "\s ta s'. False" "\s ta s'. False" for t by(unfold_locales) simp show "FWdelay_bisimulation_diverge final1 r1 final2 r2 bisim (\x1 x2. True) (\s ta s'. False) (\s ta s'. False)" proof(unfold_locales) fix t' x m1 xx m2 x1 x2 t x1' ta1 x1'' m1' x2' ta2 x2'' m2' assume bisim: "t' \ (x, m1) \ (xx, m2)" and bisim12: "t \ (x1, m1) \ (x2, m2)" and \1: "\trsys.silent_moves (r1 t) (\s ta s'. False) (x1, m1) (x1', m1)" and red1: "t \ (x1', m1) -1-ta1\ (x1'', m1')" and \2: "\trsys.silent_moves (r2 t) (\s ta s'. False) (x2, m2) (x2', m2)" and red2: "t \ (x2', m2) -2-ta2\ (x2'', m2')" and bisim12': "t \ (x1'', m1') \ (x2'', m2')" and tasim: "ta1 \m ta2" from \1 \2 have [simp]: "x1' = x1" "x2' = x2" by(simp_all add: rtranclp_False \moves_False) from bisim12 bisim_inv_red_other[OF bisim _ red1 red2 bisim12' tasim] show "t' \ (x, m1') \ (xx, m2')" by simp next fix t x1 m1 x2 m2 ta1 x1' m1' assume "t \ (x1, m1) \ (x2, m2)" "t \ (x1, m1) -1-ta1\ (x1', m1')" from simulation1[OF this] show "\ta2 x2' m2'. t \ (x2, m2) -2-ta2\ (x2', m2') \ t \ (x1', m1') \ (x2', m2') \ ta1 \m ta2" by auto next fix t x1 m1 x2 m2 ta2 x2' m2' assume "t \ (x1, m1) \ (x2, m2)" "t \ (x2, m2) -2-ta2\ (x2', m2')" from simulation2[OF this] show "\ta1 x1' m1'. t \ (x1, m1) -1-ta1\ (x1', m1') \ t \ (x1', m1') \ (x2', m2') \ ta1 \m ta2" by auto next show "(\x1. final1 x1) \ (\x2. final2 x2)" by(rule ex_final1_conv_ex_final2) qed(fastforce simp add: bisim_final)+ qed context FWbisimulation begin lemma FWbisimulation_flip: "FWbisimulation final2 r2 final1 r1 (\t. flip (bisim t))" apply(rule FWbisimulation.intro) apply(rule r2.multithreaded_axioms) apply(rule r1.multithreaded_axioms) apply(rule FWbisimulation_axioms.intro) apply(unfold flip_simps) apply(rule bisimulation_axioms) apply(erule bisim_final[symmetric]) apply(erule (5) bisim_inv_red_other) apply(rule ex_final1_conv_ex_final2[symmetric]) done end lemma FWbisimulation_flip_simps [flip_simps]: "FWbisimulation final2 r2 final1 r1 (\t. flip (bisim t)) = FWbisimulation final1 r1 final2 r2 bisim" by(auto dest: FWbisimulation.FWbisimulation_flip simp only: flip_flip) context FWbisimulation begin text \ The notation for mbisim is lost because @{term "bisim_wait"} is instantiated to @{term "\x1 x2. True"}. This reintroduces the syntax, but it does not work for output mode. This would require a new abbreviation. \ notation mbisim ("_ \m _" [50, 50] 60) theorem mbisim_bisimulation: "bisimulation r1.redT r2.redT mbisim mta_bisim" proof fix s1 s2 tta1 s1' assume mbisim: "s1 \m s2" and "r1.redT s1 tta1 s1'" from mthr.simulation1[OF this] show "\s2' tta2. r2.redT s2 tta2 s2' \ s1' \m s2' \ tta1 \T tta2" by(auto simp add: \moves_False m\move_False) next fix s2 s1 tta2 s2' assume "s1 \m s2" and "r2.redT s2 tta2 s2'" from mthr.simulation2[OF this] show "\s1' tta1. r1.redT s1 tta1 s1' \ s1' \m s2' \ tta1 \T tta2" by(auto simp add: \moves_False m\move_False) qed lemma mbisim_wset_eq: "s1 \m s2 \ wset s1 = wset s2" by(simp add: mbisim_def) lemma mbisim_mfinal: "s1 \m s2 \ r1.mfinal s1 \ r2.mfinal s2" apply(auto intro!: r2.mfinalI r1.mfinalI dest: mbisim_thrD2 mbisim_thrD1 bisim_final elim: r1.mfinalE r2.mfinalE) apply(frule (1) mbisim_thrD2, drule mbisim_wset_eq, auto elim: r1.mfinalE) apply(frule (1) mbisim_thrD1, drule mbisim_wset_eq, auto elim: r2.mfinalE) done end sublocale FWbisimulation < mthr: bisimulation r1.redT r2.redT mbisim mta_bisim by(rule mbisim_bisimulation) sublocale FWbisimulation < mthr: bisimulation_final r1.redT r2.redT mbisim mta_bisim r1.mfinal r2.mfinal by(unfold_locales)(rule mbisim_mfinal) end diff --git a/thys/JinjaThreads/Framework/FWCondAction.thy b/thys/JinjaThreads/Framework/FWCondAction.thy --- a/thys/JinjaThreads/Framework/FWCondAction.thy +++ b/thys/JinjaThreads/Framework/FWCondAction.thy @@ -1,107 +1,107 @@ (* Title: JinjaThreads/Framework/FWCondAction.thy Author: Andreas Lochbihler *) section \Semantics of the thread actions for purely conditional purpose such as Join\ theory FWCondAction imports FWState begin locale final_thread = fixes final :: "'x \ bool" begin primrec cond_action_ok :: "('l,'t,'x,'m,'w) state \ 't \ 't conditional_action \ bool" where "\ln. cond_action_ok s t (Join T) = (case thr s T of None \ True | \(x, ln)\ \ t \ T \ final x \ ln = no_wait_locks \ wset s T = None)" | "cond_action_ok s t Yield = True" primrec cond_action_oks :: "('l,'t,'x,'m,'w) state \ 't \ 't conditional_action list \ bool" where "cond_action_oks s t [] = True" | "cond_action_oks s t (ct#cts) = (cond_action_ok s t ct \ cond_action_oks s t cts)" lemma cond_action_oks_append [simp]: "cond_action_oks s t (cts @ cts') \ cond_action_oks s t cts \ cond_action_oks s t cts'" by(induct cts, auto) lemma cond_action_oks_conv_set: "cond_action_oks s t cts \ (\ct \ set cts. cond_action_ok s t ct)" by(induct cts) simp_all lemma cond_action_ok_Join: "\ln. \ cond_action_ok s t (Join T); thr s T = \(x, ln)\ \ \ final x \ ln = no_wait_locks \ wset s T = None" by(auto) lemma cond_action_oks_Join: "\ln. \ cond_action_oks s t cas; Join T \ set cas; thr s T = \(x, ln)\ \ \ final x \ ln = no_wait_locks \ wset s T = None \ t \ T" by(induct cas)(auto) lemma cond_action_oks_upd: assumes tst: "thr s t = \xln\" - shows "cond_action_oks (locks s, (thr s(t \ xln'), shr s), wset s, interrupts s) t cas = cond_action_oks s t cas" + shows "cond_action_oks (locks s, ((thr s)(t \ xln'), shr s), wset s, interrupts s) t cas = cond_action_oks s t cas" proof(induct cas) case Nil thus ?case by simp next case (Cons ca cas) - from tst have eq: "cond_action_ok (locks s, (thr s(t \ xln'), shr s), wset s, interrupts s) t ca = cond_action_ok s t ca" + from tst have eq: "cond_action_ok (locks s, ((thr s)(t \ xln'), shr s), wset s, interrupts s) t ca = cond_action_ok s t ca" by(cases ca) auto with Cons show ?case by(auto simp del: fun_upd_apply) qed lemma cond_action_ok_shr_change: "cond_action_ok (ls, (ts, m), ws, is) t ct \ cond_action_ok (ls, (ts, m'), ws, is) t ct" by(cases ct) auto lemma cond_action_oks_shr_change: "cond_action_oks (ls, (ts, m), ws, is) t cts \ cond_action_oks (ls, (ts, m'), ws, is) t cts" by(auto simp add: cond_action_oks_conv_set intro: cond_action_ok_shr_change) primrec cond_action_ok' :: "('l,'t,'x,'m,'w) state \ 't \ 't conditional_action \ bool" where "cond_action_ok' _ _ (Join t) = True" | "cond_action_ok' _ _ Yield = True" primrec cond_action_oks' :: "('l,'t,'x,'m,'w) state \ 't \ 't conditional_action list \ bool" where "cond_action_oks' s t [] = True" | "cond_action_oks' s t (ct#cts) = (cond_action_ok' s t ct \ cond_action_oks' s t cts)" lemma cond_action_oks'_append [simp]: "cond_action_oks' s t (cts @ cts') \ cond_action_oks' s t cts \ cond_action_oks' s t cts'" by(induct cts, auto) lemma cond_action_oks'_subset_Join: "set cts \ insert Yield (range Join) \ cond_action_oks' s t cts" apply(induct cts) apply(auto) done end definition collect_cond_actions :: "'t conditional_action list \ 't set" where "collect_cond_actions cts = {t. Join t \ set cts}" declare collect_cond_actions_def [simp] lemma cond_action_ok_final_change: "\ final_thread.cond_action_ok final1 s1 t ca; \t. thr s1 t = None \ thr s2 t = None; \t x1. \ thr s1 t = \(x1, no_wait_locks)\; final1 x1; wset s1 t = None \ \ \x2. thr s2 t = \(x2, no_wait_locks)\ \ final2 x2 \ ln2 = no_wait_locks \ wset s2 t = None \ \ final_thread.cond_action_ok final2 s2 t ca" apply(cases ca) apply(fastforce simp add: final_thread.cond_action_ok.simps)+ done lemma cond_action_oks_final_change: assumes major: "final_thread.cond_action_oks final1 s1 t cas" and minor: "\t. thr s1 t = None \ thr s2 t = None" "\t x1. \ thr s1 t = \(x1, no_wait_locks)\; final1 x1; wset s1 t = None \ \ \x2. thr s2 t = \(x2, no_wait_locks)\ \ final2 x2 \ ln2 = no_wait_locks \ wset s2 t = None" shows "final_thread.cond_action_oks final2 s2 t cas" using major by(induct cas)(auto simp add: final_thread.cond_action_oks.simps intro: cond_action_ok_final_change[OF _ minor]) end diff --git a/thys/JinjaThreads/Framework/FWDeadlock.thy b/thys/JinjaThreads/Framework/FWDeadlock.thy --- a/thys/JinjaThreads/Framework/FWDeadlock.thy +++ b/thys/JinjaThreads/Framework/FWDeadlock.thy @@ -1,1024 +1,1024 @@ (* Title: JinjaThreads/Framework/FWDeadlock.thy Author: Andreas Lochbihler *) section \Deadlock formalisation\ theory FWDeadlock imports FWProgressAux begin context final_thread begin definition all_final_except :: "('l,'t,'x,'m,'w) state \ 't set \ bool" where "all_final_except s Ts \ \t. not_final_thread s t \ t \ Ts" lemma all_final_except_mono [mono]: "(\x. x \ A \ x \ B) \ all_final_except ts A \ all_final_except ts B" by(auto simp add: all_final_except_def) lemma all_final_except_mono': "\ all_final_except ts A; \x. x \ A \ x \ B \ \ all_final_except ts B" by(blast intro: all_final_except_mono[rule_format]) lemma all_final_exceptI: "(\t. not_final_thread s t \ t \ Ts) \ all_final_except s Ts" by(auto simp add: all_final_except_def) lemma all_final_exceptD: "\ all_final_except s Ts; not_final_thread s t \ \ t \ Ts" by(auto simp add: all_final_except_def) inductive must_wait :: "('l,'t,'x,'m,'w) state \ 't \ ('l + 't + 't) \ 't set \ bool" for s :: "('l,'t,'x,'m,'w) state" and t :: "'t" where \ \Lock l\ "\ has_lock (locks s $ l) t'; t' \ t; t' \ Ts \ \ must_wait s t (Inl l) Ts" | \ \Join t'\ "\ not_final_thread s t'; t' \ Ts \ \ must_wait s t (Inr (Inl t')) Ts" | \ \IsInterrupted t' True\ "\ all_final_except s Ts; t' \ interrupts s \ \ must_wait s t (Inr (Inr t')) Ts" declare must_wait.cases [elim] declare must_wait.intros [intro] lemma must_wait_elims [consumes 1, case_names lock join interrupt, cases pred]: assumes "must_wait s t lt Ts" obtains l t' where "lt = Inl l" "has_lock (locks s $ l) t'" "t' \ t" "t' \ Ts" | t' where "lt = Inr (Inl t')" "not_final_thread s t'" "t' \ Ts" | t' where "lt = Inr (Inr t')" "all_final_except s Ts" "t' \ interrupts s" using assms by(auto) inductive_cases must_wait_elims2 [elim!]: "must_wait s t (Inl l) Ts" "must_wait s t (Inr (Inl t'')) Ts" "must_wait s t (Inr (Inr t'')) Ts" lemma must_wait_iff: "must_wait s t lt Ts \ (case lt of Inl l \ \t'\Ts. t \ t' \ has_lock (locks s $ l) t' | Inr (Inl t') \ not_final_thread s t' \ t' \ Ts | Inr (Inr t') \ all_final_except s Ts \ t' \ interrupts s)" by(auto simp add: must_wait.simps split: sum.splits) end text\Deadlock as a system-wide property\ context multithreaded_base begin definition deadlock :: "('l,'t,'x,'m,'w) state \ bool" where "deadlock s \ (\t x. thr s t = \(x, no_wait_locks)\ \ \ final x \ wset s t = None \ t \ \x, shr s\ \ \ (\LT. t \ \x, shr s\ LT \ \ (\lt \ LT. must_wait s t lt (dom (thr s))))) \ (\t x ln. thr s t = \(x, ln)\ \ (\l. ln $ l > 0) \ \ waiting (wset s t) \ (\l t'. ln $ l > 0 \ t \ t' \ thr s t' \ None \ has_lock (locks s $ l) t')) \ (\t x w. thr s t = \(x, no_wait_locks)\ \ wset s t \ \PostWS w\)" lemma deadlockI: "\ \t x. \ thr s t = \(x, no_wait_locks)\; \ final x; wset s t = None \ \ t \ \x, shr s\ \ \ (\LT. t \ \x, shr s\ LT \ \ (\lt \ LT. must_wait s t lt (dom (thr s)))); \t x ln l. \ thr s t = \(x, ln)\; ln $ l > 0; \ waiting (wset s t) \ \ \l t'. ln $ l > 0 \ t \ t' \ thr s t' \ None \ has_lock (locks s $ l) t'; \t x w. thr s t = \(x, no_wait_locks)\ \ wset s t \ \PostWS w\ \ \ deadlock s" by(auto simp add: deadlock_def) lemma deadlockE: assumes "deadlock s" obtains "\t x. thr s t = \(x, no_wait_locks)\ \ \ final x \ wset s t = None \ t \ \x, shr s\ \ \ (\LT. t \ \x, shr s\ LT \ \ (\lt \ LT. must_wait s t lt (dom (thr s))))" and "\t x ln. thr s t = \(x, ln)\ \ (\l. ln $ l > 0) \ \ waiting (wset s t) \ (\l t'. ln $ l > 0 \ t \ t' \ thr s t' \ None \ has_lock (locks s $ l) t')" and "\t x w. thr s t = \(x, no_wait_locks)\ \ wset s t \ \PostWS w\" using assms unfolding deadlock_def by(blast) lemma deadlockD1: assumes "deadlock s" and "thr s t = \(x, no_wait_locks)\" and "\ final x" and "wset s t = None" obtains "t \ \x, shr s\ \" and "\LT. t \ \x, shr s\ LT \ \ (\lt \ LT. must_wait s t lt (dom (thr s)))" using assms unfolding deadlock_def by(blast) lemma deadlockD2: fixes ln assumes "deadlock s" and "thr s t = \(x, ln)\" and "ln $ l > 0" and "\ waiting (wset s t)" obtains l' t' where "ln $ l' > 0" "t \ t'" "thr s t' \ None" "has_lock (locks s $ l') t'" using assms unfolding deadlock_def by blast lemma deadlockD3: assumes "deadlock s" and "thr s t = \(x, no_wait_locks)\" shows "\w. wset s t \ \PostWS w\" using assms unfolding deadlock_def by blast lemma deadlock_def2: "deadlock s \ (\t x. thr s t = \(x, no_wait_locks)\ \ \ final x \ wset s t = None \ t \ \x, shr s\ \ \ (\LT. t \ \x, shr s\ LT \ \ (\lt \ LT. must_wait s t lt (dom (thr s))))) \ (\t x ln. thr s t = \(x, ln)\ \ ln \ no_wait_locks \ \ waiting (wset s t) \ (\l. ln $ l > 0 \ must_wait s t (Inl l) (dom (thr s)))) \ (\t x w. thr s t = \(x, no_wait_locks)\ \ wset s t \ \PostWS WSNotified\ \ wset s t \ \PostWS WSWokenUp\)" unfolding neq_no_wait_locks_conv apply(rule iffI) apply(intro strip conjI) apply(blast dest: deadlockD1) apply(blast dest: deadlockD1) apply(blast elim: deadlockD2) apply(blast dest: deadlockD3) apply(blast dest: deadlockD3) apply(elim conjE exE) apply(rule deadlockI) apply blast apply(rotate_tac 1) apply(erule allE, rotate_tac -1) apply(erule allE, rotate_tac -1) apply(erule allE, rotate_tac -1) apply(erule impE, blast) apply(elim exE conjE) apply(erule must_wait.cases) apply(clarify) apply(rotate_tac 3) apply(rule exI conjI|erule not_sym|assumption)+ apply blast apply blast apply blast apply blast apply(case_tac w) apply blast apply blast done lemma all_waiting_implies_deadlock: assumes "lock_thread_ok (locks s) (thr s)" and normal: "\t x. \ thr s t = \(x, no_wait_locks)\; \ final x; wset s t = None \ \ t \ \x, shr s\ \ \ (\LT. t \ \x, shr s\ LT \ \ (\lt \ LT. must_wait s t lt (dom (thr s))))" and acquire: "\t x ln l. \ thr s t = \(x, ln)\; \ waiting (wset s t); ln $ l > 0 \ \ \l'. ln $ l' > 0 \ \ may_lock (locks s $ l') t" and wakeup: "\t x w. thr s t = \(x, no_wait_locks)\ \ wset s t \ \PostWS w\" shows "deadlock s" proof(rule deadlockI) fix T X assume "thr s T = \(X, no_wait_locks)\" "\ final X" "wset s T = None" thus "T \ \X, shr s\ \ \ (\LT. T \ \X, shr s\ LT \ \ (\lt\LT. must_wait s T lt (dom (thr s))))" by(rule normal) next fix T X LN l' assume "thr s T = \(X, LN)\" and "0 < LN $ l'" and wset: "\ waiting (wset s T)" from acquire[OF \thr s T = \(X, LN)\\ wset, OF \0 < LN $ l'\] obtain l' where "0 < LN $ l'" "\ may_lock (locks s $ l') T" by blast then obtain t' where "T \ t'" "has_lock (locks s $ l') t'" unfolding not_may_lock_conv by fastforce moreover with \lock_thread_ok (locks s) (thr s)\ have "thr s t' \ None" by(auto dest: lock_thread_okD) ultimately show "\l t'. 0 < LN $ l \ T \ t' \ thr s t' \ None \ has_lock (locks s $ l) t'" using \0 < LN $ l'\ by(auto) qed(rule wakeup) lemma mfinal_deadlock: "mfinal s \ deadlock s" unfolding mfinal_def2 by(rule deadlockI)(auto simp add: final_thread_def) text \Now deadlock for single threads\ lemma must_wait_mono: "(\x. x \ A \ x \ B) \ must_wait s t lt A \ must_wait s t lt B" by(auto simp add: must_wait_iff split: sum.split elim: all_final_except_mono') lemma must_wait_mono': "\ must_wait s t lt A; A \ B \ \ must_wait s t lt B" using must_wait_mono[of A B s t lt] by blast end lemma UN_mono: "\ x \ A \ x \ A'; x \ B \ x \ B' \ \ x \ A \ B \ x \ A' \ B'" by blast lemma Collect_mono_conv [mono]: "x \ {x. P x} \ P x" by blast context multithreaded_base begin coinductive_set deadlocked :: "('l,'t,'x,'m,'w) state \ 't set" for s :: "('l,'t,'x,'m,'w) state" where deadlockedLock: "\ thr s t = \(x, no_wait_locks)\; t \ \x, shr s\ \; wset s t = None; \LT. t \ \x, shr s\ LT \ \ \lt \ LT. must_wait s t lt (deadlocked s \ final_threads s) \ \ t \ deadlocked s" | deadlockedWait: "\ln. \ thr s t = \(x, ln)\; all_final_except s (deadlocked s); waiting (wset s t) \ \ t \ deadlocked s" | deadlockedAcquire: "\ln. \ thr s t = \(x, ln)\; \ waiting (wset s t); ln $ l > 0; has_lock (locks s $ l) t'; t' \ t; t' \ deadlocked s \ final_thread s t' \ \ t \ deadlocked s" monos must_wait_mono UN_mono lemma deadlockedAcquire_must_wait: "\ln. \ thr s t = \(x, ln)\; \ waiting (wset s t); ln $ l > 0; must_wait s t (Inl l) (deadlocked s \ final_threads s) \ \ t \ deadlocked s" apply(erule must_wait_elims) apply(erule (2) deadlockedAcquire) apply auto done lemma deadlocked_elims [consumes 1, case_names lock wait acquire]: assumes "t \ deadlocked s" and lock: "\x. \ thr s t = \(x, no_wait_locks)\; t \ \x, shr s\ \; wset s t = None; \LT. t \ \x, shr s\ LT \ \ \lt \ LT. must_wait s t lt (deadlocked s \ final_threads s) \ \ thesis" and wait: "\x ln. \ thr s t = \(x, ln)\; all_final_except s (deadlocked s); waiting (wset s t) \ \ thesis" and acquire: "\x ln l t'. \ thr s t = \(x, ln)\; \ waiting (wset s t); 0 < ln $ l; has_lock (locks s $ l) t'; t \ t'; t' \ deadlocked s \ final_thread s t' \ \ thesis" shows thesis using assms by cases blast+ lemma deadlocked_coinduct [consumes 1, case_names deadlocked, case_conclusion deadlocked Lock Wait Acquire, coinduct set: deadlocked]: assumes major: "t \ X" and step: "\t. t \ X \ (\x. thr s t = \(x, no_wait_locks)\ \ t \ \x, shr s\ \ \ wset s t = None \ (\LT. t \ \x, shr s\ LT \ \ (\lt\LT. must_wait s t lt (X \ deadlocked s \ final_threads s)))) \ (\x ln. thr s t = \(x, ln)\ \ all_final_except s (X \ deadlocked s) \ waiting (wset s t)) \ (\x l t' ln. thr s t = \(x, ln)\ \ \ waiting (wset s t) \ 0 < ln $ l \ has_lock (locks s $ l) t' \ t' \ t \ ((t' \ X \ t' \ deadlocked s) \ final_thread s t'))" shows "t \ deadlocked s" using major proof(coinduct) case (deadlocked t) have "X \ deadlocked s \ final_threads s = {x. x \ X \ x \ deadlocked s \ x \ final_threads s}" by auto moreover have "X \ deadlocked s = {x. x \ X \ x \ deadlocked s}" by blast ultimately show ?case using step[OF deadlocked] by(elim disjE) simp_all qed definition deadlocked' :: "('l,'t,'x,'m,'w) state \ bool" where "deadlocked' s \ (\t. not_final_thread s t \ t \ deadlocked s)" lemma deadlocked'I: "(\t. not_final_thread s t \ t \ deadlocked s) \ deadlocked' s" by(auto simp add: deadlocked'_def) lemma deadlocked'D2: "\ deadlocked' s; not_final_thread s t; t \ deadlocked s \ thesis \ \ thesis" by(auto simp add: deadlocked'_def) lemma not_deadlocked'I: "\ not_final_thread s t; t \ deadlocked s \ \ \ deadlocked' s" by(auto dest: deadlocked'D2) lemma deadlocked'_intro: "\ \t. not_final_thread s t \ t \ deadlocked s \ \ deadlocked' s" by(rule deadlocked'I)(blast)+ lemma deadlocked_thread_exists: assumes "t \ deadlocked s" and "\x ln. thr s t = \(x, ln)\ \ thesis" shows thesis using assms by cases blast+ end context multithreaded begin lemma red_no_deadlock: assumes P: "s -t\ta\ s'" and dead: "t \ deadlocked s" shows False proof - from P show False proof(cases) case (redT_normal x x' m') note red = \t \ \x, shr s\ -ta\ \x', m'\\ note tst = \thr s t = \(x, no_wait_locks)\\ note aok = \actions_ok s t ta\ show False proof(cases "\w. wset s t = \InWS w\") case True with aok show ?thesis by(auto simp add: wset_actions_ok_def split: if_split_asm) next case False with dead tst have mle: "t \ \x, shr s\ \" and cledead: "\LT. t \ \x, shr s\ LT \ \ (\lt \ LT. must_wait s t lt (deadlocked s \ final_threads s))" by(cases, auto simp add: waiting_def)+ let ?LT = "collect_waits ta" from red have "t \ \x, shr s\ ?LT \" by(auto intro: can_syncI) then obtain lt where lt: "lt \ ?LT" and mw: "must_wait s t lt (deadlocked s \ final_threads s)" by(blast dest: cledead[rule_format]) from mw show False proof(cases rule: must_wait_elims) case (lock l t') from \lt = Inl l\ lt have "l \ collect_locks \ta\\<^bsub>l\<^esub>" by(auto) with aok have "may_lock (locks s $ l) t" by(auto elim!: collect_locksE lock_ok_las_may_lock) with \has_lock (locks s $ l) t'\ have "t' = t" by(auto dest: has_lock_may_lock_t_eq) with \t' \ t\ show False by contradiction next case (join t') from \lt = Inr (Inl t')\ lt have "Join t' \ set \ta\\<^bsub>c\<^esub>" by auto from \not_final_thread s t'\ obtain x'' ln'' where "thr s t' = \(x'', ln'')\" by(rule not_final_thread_existsE) moreover with \Join t' \ set \ta\\<^bsub>c\<^esub>\ aok have "final x''" "ln'' = no_wait_locks" "wset s t' = None" by(auto dest: cond_action_oks_Join) ultimately show False using \not_final_thread s t'\ by(auto) next case (interrupt t') from aok lt \lt = Inr (Inr t')\ have "t' \ interrupts s" by(auto intro: collect_interrupts_interrupted) with \t' \ interrupts s\ show False by contradiction qed qed next case (redT_acquire x n ln) show False proof(cases "\w. wset s t = \InWS w\") case True with \\ waiting (wset s t)\ show ?thesis by(auto simp add: not_waiting_iff) next case False with dead \thr s t = \(x, ln)\\ \0 < ln $ n\ obtain l t' where "0 < ln $ l" "t \ t'" and "has_lock (locks s $ l) t'" by(cases)(fastforce simp add: waiting_def)+ hence "\ may_acquire_all (locks s) t ln" by(auto elim: may_acquire_allE dest: has_lock_may_lock_t_eq) with \may_acquire_all (locks s) t ln\ show ?thesis by contradiction qed qed qed lemma deadlocked'_no_red: "\ s -t\ta\ s'; deadlocked' s \ \ False" apply(rule red_no_deadlock) apply(assumption) apply(erule deadlocked'D2) by(rule red_not_final_thread) lemma not_final_thread_deadlocked_final_thread [iff]: "thr s t = \xln\ \ not_final_thread s t \ t \ deadlocked s \ final_thread s t" by(auto simp add: not_final_thread_final_thread_conv[symmetric]) lemma all_waiting_deadlocked: assumes "not_final_thread s t" and "lock_thread_ok (locks s) (thr s)" and normal: "\t x. \ thr s t = \(x, no_wait_locks)\; \ final x; wset s t = None \ \ t \ \x, shr s\ \ \ (\LT. t \ \x, shr s\ LT \ \ (\lt\LT. must_wait s t lt (final_threads s)))" and acquire: "\t x ln l. \ thr s t = \(x, ln)\; \ waiting (wset s t); ln $ l > 0 \ \ \l'. ln $ l' > 0 \ \ may_lock (locks s $ l') t" and wakeup: "\t x w. thr s t = \(x, no_wait_locks)\ \ wset s t \ \PostWS w\" shows "t \ deadlocked s" proof - from \not_final_thread s t\ have "t \ {t. not_final_thread s t}" by simp thus ?thesis proof(coinduct) case (deadlocked z) hence "not_final_thread s z" by simp then obtain x' ln' where "thr s z = \(x', ln')\" by(fastforce elim!: not_final_thread_existsE) { assume "wset s z = None" "\ final x'" and [simp]: "ln' = no_wait_locks" with \thr s z = \(x', ln')\\ have "z \ \x', shr s\ \ \ (\LT. z \ \x', shr s\ LT \ \ (\lt \ LT. must_wait s z lt (final_threads s)))" by(auto dest: normal) then obtain "z \ \x', shr s\ \" and clnml: "\LT. z \ \x', shr s\ LT \ \ \lt \ LT. must_wait s z lt (final_threads s)" by(blast) { fix LT assume "z \ \x', shr s\ LT \" then obtain lt where mw: "must_wait s z lt (final_threads s)" and lt: "lt \ LT" by(blast dest: clnml) from mw have "must_wait s z lt ({t. not_final_thread s t} \ deadlocked s \ final_threads s)" by(blast intro: must_wait_mono') with lt have "\lt \ LT. must_wait s z lt ({t. not_final_thread s t} \ deadlocked s \ final_threads s)" by blast } with \z \ \x', shr s\ \\ \thr s z = \(x', ln')\\ \wset s z = None\ have ?case by(simp) } note c1 = this { assume wsz: "\ waiting (wset s z)" and "ln' \ no_wait_locks" from \ln' \ no_wait_locks\ obtain l where "0 < ln' $ l" by(auto simp add: neq_no_wait_locks_conv) with wsz \thr s z = \(x', ln')\\ obtain l' where "0 < ln' $ l'" "\ may_lock (locks s $ l') z" by(blast dest: acquire) then obtain t'' where "t'' \ z" "has_lock (locks s $ l') t''" unfolding not_may_lock_conv by blast with \lock_thread_ok (locks s) (thr s)\ obtain x'' ln'' where "thr s t'' = \(x'', ln'')\" by(auto elim!: lock_thread_ok_has_lockE) hence "(not_final_thread s t'' \ t'' \ deadlocked s) \ final_thread s t''" by(clarsimp simp add: not_final_thread_iff final_thread_def) with wsz \0 < ln' $ l'\ \thr s z = \(x', ln')\\ \t'' \ z\ \has_lock (locks s $ l') t''\ have ?Acquire by simp blast hence ?case by simp } note c2 = this { fix w assume "waiting (wset s z)" with \thr s z = \(x', ln')\\ have "?Wait" by(clarsimp simp add: all_final_except_def) hence ?case by simp } note c3 = this from \not_final_thread s z\ \thr s z = \(x', ln')\\ show ?case proof(cases rule: not_final_thread_cases2) case final show ?thesis proof(cases "wset s z") case None show ?thesis proof(cases "ln' = no_wait_locks") case True with None final show ?thesis by(rule c1) next case False from None have "\ waiting (wset s z)" by(simp add: not_waiting_iff) thus ?thesis using False by(rule c2) qed next case (Some w) show ?thesis proof(cases w) case (InWS w') with Some have "waiting (wset s z)" by(simp add: waiting_def) thus ?thesis by(rule c3) next case (PostWS w') with Some have "\ waiting (wset s z)" by(simp add: not_waiting_iff) moreover from PostWS \thr s z = \(x', ln')\\ Some have "ln' \ no_wait_locks" by(auto dest: wakeup) ultimately show ?thesis by(rule c2) qed qed next case wait_locks show ?thesis proof(cases "wset s z") case None hence "\ waiting (wset s z)" by(simp add: not_waiting_iff) thus ?thesis using wait_locks by(rule c2) next case (Some w) show ?thesis proof(cases w) case (InWS w') with Some have "waiting (wset s z)" by(simp add: waiting_def) thus ?thesis by(rule c3) next case (PostWS w') with Some have "\ waiting (wset s z)" by(simp add: not_waiting_iff) moreover from PostWS \thr s z = \(x', ln')\\ Some have "ln' \ no_wait_locks" by(auto dest: wakeup) ultimately show ?thesis by(rule c2) qed qed next case (wait_set w) show ?thesis proof(cases w) case (InWS w') with wait_set have "waiting (wset s z)" by(simp add: waiting_def) thus ?thesis by(rule c3) next case (PostWS w') with wait_set have "\ waiting (wset s z)" by(simp add: not_waiting_iff) moreover from PostWS \thr s z = \(x', ln')\\ wait_set have "ln' \ no_wait_locks" by(auto dest: wakeup[simplified]) ultimately show ?thesis by(rule c2) qed qed qed qed text \Equivalence proof for both notions of deadlock\ lemma deadlock_implies_deadlocked': assumes dead: "deadlock s" shows "deadlocked' s" proof - show ?thesis proof(rule deadlocked'I) fix t assume "not_final_thread s t" hence "t \ {t. not_final_thread s t}" .. thus "t \ deadlocked s" proof(coinduct) case (deadlocked t'') hence "not_final_thread s t''" .. then obtain x'' ln'' where tst'': "thr s t'' = \(x'', ln'')\" by(rule not_final_thread_existsE) { assume "waiting (wset s t'')" moreover with tst'' have nfine: "not_final_thread s t''" unfolding waiting_def by(blast intro: not_final_thread.intros) ultimately have ?case using tst'' by(blast intro: all_final_exceptI not_final_thread_final) } note c1 = this { assume wst'': "\ waiting (wset s t'')" and "ln'' \ no_wait_locks" then obtain l where l: "ln'' $ l > 0" by(auto simp add: neq_no_wait_locks_conv) with dead wst'' tst'' obtain l' T where "ln'' $ l' > 0" "t'' \ T" and hl: "has_lock (locks s $ l') T" and tsT: "thr s T \ None" by - (erule deadlockD2) moreover from \thr s T \ None\ obtain xln where tsT: "thr s T = \xln\" by auto then obtain X LN where "thr s T = \(X, LN)\" by(cases xln, auto) moreover hence "not_final_thread s T \ final_thread s T" by(auto simp add: final_thread_def not_final_thread_iff) ultimately have ?case using wst'' tst'' by blast } note c2 = this { assume "wset s t'' = None" and [simp]: "ln'' = no_wait_locks" moreover with \not_final_thread s t''\ tst'' have "\ final x''" by(auto) ultimately obtain "t'' \ \x'', shr s\ \" and clnml: "\LT. t'' \ \x'', shr s\ LT \ \ \t'. thr s t' \ None \ (\lt\LT. must_wait s t'' lt (dom (thr s)))" using \thr s t'' = \(x'', ln'')\\ \deadlock s\ by(blast elim: deadlockD1) { fix LT assume "t'' \ \x'', shr s\ LT \" then obtain lt where lt: "lt \ LT" and mw: "must_wait s t'' lt (dom (thr s))" by(blast dest: clnml) note mw also have "dom (thr s) = {t. not_final_thread s t} \ deadlocked s \ final_threads s" by(auto simp add: not_final_thread_conv dest: deadlocked_thread_exists elim: final_threadE) finally have "\lt\LT. must_wait s t'' lt ({t. not_final_thread s t} \ deadlocked s \ final_threads s)" using lt by blast } with \t'' \ \x'', shr s\ \\ tst'' \wset s t'' = None\ have ?case by(simp) } note c3 = this from \not_final_thread s t''\ tst'' show ?case proof(cases rule: not_final_thread_cases2) case final show ?thesis proof(cases "wset s t''") case None show ?thesis proof(cases "ln'' = no_wait_locks") case True with None show ?thesis by(rule c3) next case False from None have "\ waiting (wset s t'')" by(simp add: not_waiting_iff) thus ?thesis using False by(rule c2) qed next case (Some w) show ?thesis proof(cases w) case (InWS w') with Some have "waiting (wset s t'')" by(simp add: waiting_def) thus ?thesis by(rule c1) next case (PostWS w') hence "\ waiting (wset s t'')" using Some by(simp add: not_waiting_iff) moreover from PostWS Some tst'' have "ln'' \ no_wait_locks" by(auto dest: deadlockD3[OF dead]) ultimately show ?thesis by(rule c2) qed qed next case wait_locks show ?thesis proof(cases "waiting (wset s t'')") case False thus ?thesis using wait_locks by(rule c2) next case True thus ?thesis by(rule c1) qed next case (wait_set w) show ?thesis proof(cases w) case InWS with wait_set have "waiting (wset s t'')" by(simp add: waiting_def) thus ?thesis by(rule c1) next case (PostWS w') hence "\ waiting (wset s t'')" using wait_set by(simp add: not_waiting_iff) moreover from PostWS wait_set tst'' have "ln'' \ no_wait_locks" by(auto dest: deadlockD3[OF dead]) ultimately show ?thesis by(rule c2) qed qed qed qed qed lemma deadlocked'_implies_deadlock: assumes dead: "deadlocked' s" shows "deadlock s" proof - have deadlocked: "\t. not_final_thread s t \ t \ deadlocked s" using dead by(rule deadlocked'D2) show ?thesis proof(rule deadlockI) fix t' x' assume "thr s t' = \(x', no_wait_locks)\" and "\ final x'" and "wset s t' = None" hence "not_final_thread s t'" by(auto intro: not_final_thread_final) hence "t' \ deadlocked s" by(rule deadlocked) thus "t' \ \x', shr s\ \ \ (\LT. t' \ \x', shr s\ LT \ \ (\lt \ LT. must_wait s t' lt (dom (thr s))))" proof(cases rule: deadlocked_elims) case (lock x'') note lock = \\LT. t' \ \x'', shr s\ LT \ \ \lt \ LT. must_wait s t' lt (deadlocked s \ final_threads s)\ from \thr s t' = \(x'', no_wait_locks)\\ \thr s t' = \(x', no_wait_locks)\\ have [simp]: "x' = x''" by auto { fix LT assume "t' \ \x'', shr s\ LT \" from lock[OF this] obtain lt where lt: "lt \ LT" and mw: "must_wait s t' lt (deadlocked s \ final_threads s)" by blast have "deadlocked s \ final_threads s \ dom (thr s)" by(auto elim: final_threadE dest: deadlocked_thread_exists) with mw have "must_wait s t' lt (dom (thr s))" by(rule must_wait_mono') with lt have "\lt\LT. must_wait s t' lt (dom (thr s))" by blast } with \t' \ \x'', shr s\ \\ show ?thesis by(auto) next case (wait x'' ln'') from \wset s t' = None\ \waiting (wset s t')\ have False by(simp add: waiting_def) thus ?thesis .. next case (acquire x'' ln'' l'' T) from \thr s t' = \(x'', ln'')\\ \thr s t' = \(x', no_wait_locks)\\ \0 < ln'' $ l''\ have False by(auto) thus ?thesis .. qed next fix t' x' ln' l assume "thr s t' = \(x', ln')\" and "0 < ln' $ l" and wst': "\ waiting (wset s t')" hence "not_final_thread s t'" by(auto intro: not_final_thread_wait_locks) hence "t' \ deadlocked s" by(rule deadlocked) thus "\l T. 0 < ln' $ l \ t' \ T \ thr s T \ None \ has_lock (locks s $ l) T" proof(cases rule: deadlocked_elims) case (lock x'') from \thr s t' = \(x', ln')\\ \thr s t' = \(x'', no_wait_locks)\\ \0 < ln' $ l\ have False by auto thus ?thesis .. next case (wait x' ln') from wst' \waiting (wset s t')\ have False by contradiction thus ?thesis .. next case (acquire x'' ln'' l'' t'') from \thr s t' = \(x'', ln'')\\ \thr s t' = \(x', ln')\\ have [simp]: "x' = x''" "ln' = ln''" by auto moreover from \t'' \ deadlocked s \ final_thread s t''\ have "thr s t'' \ None" by(auto elim: deadlocked_thread_exists simp add: final_thread_def) with \0 < ln'' $ l''\ \has_lock (locks s $ l'') t''\ \t' \ t''\ \thr s t' = \(x'', ln'')\\ show ?thesis by auto qed next fix t x w assume tst: "thr s t = \(x, no_wait_locks)\" show "wset s t \ \PostWS w\" proof assume "wset s t = \PostWS w\" moreover with tst have "not_final_thread s t" by(auto simp add: not_final_thread_iff) hence "t \ deadlocked s" by(rule deadlocked) ultimately show False using tst by(auto elim: deadlocked.cases simp add: waiting_def) qed qed qed lemma deadlock_eq_deadlocked': "deadlock = deadlocked'" by(rule ext)(auto intro: deadlock_implies_deadlocked' deadlocked'_implies_deadlock) lemma deadlock_no_red: "\ s -t\ta\ s'; deadlock s \ \ False" unfolding deadlock_eq_deadlocked' by(rule deadlocked'_no_red) lemma deadlock_no_active_threads: assumes dead: "deadlock s" shows "active_threads s = {}" proof(rule equals0I) fix t assume active: "t \ active_threads s" then obtain ta s' where "s -t\ta\ s'" by(auto dest: active_thread_ex_red) thus False using dead by(rule deadlock_no_red) qed end locale preserve_deadlocked = multithreaded final r convert_RA for final :: "'x \ bool" and r :: "('l,'t,'x,'m,'w,'o) semantics" ("_ \ _ -_\ _" [50,0,0,50] 80) and convert_RA :: "'l released_locks \ 'o list" + fixes wf_state :: "('l,'t,'x,'m,'w) state set" assumes invariant3p_wf_state: "invariant3p redT wf_state" assumes can_lock_preserved: "\ s \ wf_state; s -t'\ta'\ s'; thr s t = \(x, no_wait_locks)\; t \ \x, shr s\ \ \ \ t \ \x, shr s'\ \" and can_lock_devreserp: "\ s \ wf_state; s -t'\ta'\ s'; thr s t = \(x, no_wait_locks)\; t \ \x, shr s'\ L \ \ \ \L'\L. t \ \x, shr s\ L' \" begin lemma redT_deadlocked_subset: assumes wfs: "s \ wf_state" and Red: "s -t\ta\ s'" shows "deadlocked s \ deadlocked s'" proof fix t' assume t'dead: "t' \ deadlocked s" from Red have tndead: "t \ deadlocked s" by(auto dest: red_no_deadlock) with t'dead have t't: "t' \ t" by auto { fix t' assume "final_thread s t'" then obtain x' ln' where tst': "thr s t' = \(x', ln')\" by(auto elim!: final_threadE) with \final_thread s t'\ have "final x'" and "wset s t' = None" and [simp]: "ln' = no_wait_locks" by(auto elim: final_threadE) with Red tst' have "t \ t'" by cases(auto dest: final_no_red) with Red tst' have "thr s' t' = \(x', ln')\" by cases(auto intro: redT_updTs_Some) moreover from Red \t \ t'\ \wset s t' = None\ have "wset s' t' = None" by cases(auto simp: redT_updWs_None_implies_None) ultimately have "final_thread s' t'" using tst' \final x'\ by(auto simp add: final_thread_def) } hence subset: "deadlocked s \ final_threads s \ deadlocked s \ deadlocked s' \ final_threads s'" by(auto) from Red show "t' \ deadlocked s'" proof(cases) case (redT_normal x x' m') note red = \t \ \x, shr s\ -ta\ \x', m'\\ and tst = \thr s t = \(x, no_wait_locks)\\ and aok = \actions_ok s t ta\ and s' = \redT_upd s t ta x' m' s'\ from red have "\ final x" by(auto dest: final_no_red) with tndead tst have nafe: "\ all_final_except s (deadlocked s)" by(fastforce simp add: all_final_except_def not_final_thread_iff) from t'dead show ?thesis proof(coinduct) case (deadlocked t'') note t''dead = this with Red have t''t: "t'' \ t" by(auto dest: red_no_deadlock) from t''dead show ?case proof(cases rule: deadlocked_elims) case (lock X) hence est'': "thr s t'' = \(X, no_wait_locks)\" and msE: "t'' \ \X, shr s\ \" and csexdead: "\LT. t'' \ \X, shr s\ LT \ \ \lt \ LT. must_wait s t'' lt (deadlocked s \ final_threads s)" by auto from t''t Red est'' have es't'': "thr s' t'' = \(X, no_wait_locks)\" by(cases s)(cases s', auto elim!: redT_ts_Some_inv) note es't'' moreover from wfs Red est'' msE have msE': "t'' \ \X, shr s'\ \" by(rule can_lock_preserved) moreover { fix LT assume clL'': "t'' \ \X, shr s'\ LT \" with est'' have "\LT'\LT. t'' \ \X, shr s\ LT' \" by(rule can_lock_devreserp[OF wfs Red]) then obtain LT' where clL': "t'' \ \X, shr s\ LT' \" and LL': "LT' \ LT" by blast with csexdead obtain lt where lt: "lt \ LT" and mw: "must_wait s t'' lt (deadlocked s \ final_threads s)" by blast from mw have "must_wait s' t'' lt (deadlocked s \ deadlocked s' \ final_threads s')" proof(cases rule: must_wait_elims) case (lock l t') from \t' \ deadlocked s \ final_threads s\ Red have tt': "t \ t'" by(auto dest: red_no_deadlock final_no_redT elim: final_threadE) from aok have "lock_actions_ok (locks s $ l) t (\ta\\<^bsub>l\<^esub> $ l)" by(auto simp add: lock_ok_las_def) with tt' \has_lock (locks s $ l) t'\ s' have hl't': "has_lock (locks s' $ l) t'" by(auto) moreover note \t' \ t''\ moreover from \t' \ deadlocked s \ final_threads s\ have "t' \ (deadlocked s \ deadlocked s' \ final_threads s')" using subset by blast ultimately show ?thesis unfolding \lt = Inl l\ .. next case (join t') note t'dead = \t' \ deadlocked s \ final_threads s\ with Red have tt': "t \ t'" by(auto dest: red_no_deadlock final_no_redT elim: final_threadE) note nftt' = \not_final_thread s t'\ from t'dead Red aok s' tt' have ts't': "thr s' t' = thr s t'" by(auto elim!: deadlocked_thread_exists final_threadE intro: redT_updTs_Some) from nftt' have "thr s t' \ None" by auto with nftt' t'dead have "t' \ deadlocked s" by(simp add: not_final_thread_final_thread_conv[symmetric]) hence "not_final_thread s' t'" proof(cases rule: deadlocked_elims) case (lock x'') from \t' \ \x'', shr s\ \\ have "\ final x''" by(auto elim: must_syncE dest: final_no_red) with \thr s t' = \(x'', no_wait_locks)\\ ts't' show ?thesis by(auto intro: not_final_thread.intros) next case (wait x'' ln'') from \\ final x\ tst \all_final_except s (deadlocked s)\ have "t \ deadlocked s" by(fastforce dest: all_final_exceptD simp add: not_final_thread_iff) with Red have False by(auto dest: red_no_deadlock) thus ?thesis .. next case (acquire x'' ln'' l'' T'') from \thr s t' = \(x'', ln'')\\ \0 < ln'' $ l''\ ts't' show ?thesis by(auto intro: not_final_thread.intros(2)) qed moreover from t'dead subset have "t' \ deadlocked s \ deadlocked s' \ final_threads s'" .. ultimately show ?thesis unfolding \lt = Inr (Inl t')\ .. next case (interrupt t') from tst red aok have "not_final_thread s t" by(auto simp add: wset_actions_ok_def not_final_thread_iff split: if_split_asm dest: final_no_red) with \all_final_except s (deadlocked s \ final_threads s)\ have "t \ deadlocked s \ final_threads s" by(rule all_final_exceptD) moreover have "t \ deadlocked s" using Red by(blast dest: red_no_deadlock) moreover have "\ final_thread s t" using red tst by(auto simp add: final_thread_def dest: final_no_red) ultimately have False by blast thus ?thesis .. qed with lt have "\lt\LT. must_wait s' t'' lt (deadlocked s \ deadlocked s' \ final_threads s')" by blast } moreover have "wset s' t'' = None" using s' t''t \wset s t'' = None\ by(auto intro: redT_updWs_None_implies_None) ultimately show ?thesis by(auto) next case (wait x ln) from \all_final_except s (deadlocked s)\ nafe have False by simp thus ?thesis by simp next case (acquire X ln l T) from t''t Red \thr s t'' = \(X, ln)\\ s' have es't'': "thr s' t'' = \(X, ln)\" by(cases s)(auto dest: redT_ts_Some_inv) moreover from \T \ deadlocked s \ final_thread s T\ have "T \ t" proof(rule disjE) assume "T \ deadlocked s" with Red show ?thesis by(auto dest: red_no_deadlock) next assume "final_thread s T" with Red show ?thesis by(auto dest!: final_no_redT simp add: final_thread_def) qed with s' tst Red \has_lock (locks s $ l) T\ have "has_lock (locks s' $ l) T" by -(cases s, auto dest: redT_has_lock_inv[THEN iffD2]) moreover from s' \T \ t\ have wset: "wset s T = None \ wset s' T = None" by(auto intro: redT_updWs_None_implies_None) { fix x assume "thr s T = \(x, no_wait_locks)\" with \T \ t\ Red s' aok tst have "thr s' T = \(x, no_wait_locks)\" by(auto intro: redT_updTs_Some) } moreover hence "final_thread s T \ final_thread s' T" by(auto simp add: final_thread_def intro: wset) moreover from \\ waiting (wset s t'')\ s' t''t have "\ waiting (wset s' t'')" by(auto simp add: redT_updWs_None_implies_None redT_updWs_PostWS_imp_PostWS not_waiting_iff) ultimately have ?Acquire using \0 < ln $ l\ \t'' \ T\ \T \ deadlocked s \ final_thread s T\ by(auto) thus ?thesis by simp qed qed next case (redT_acquire x n ln) hence [simp]: "ta = (K$ [], [], [], [], [], convert_RA ln)" - and s': "s' = (acquire_all (locks s) t ln, (thr s(t \ (x, no_wait_locks)), shr s), wset s, interrupts s)" + and s': "s' = (acquire_all (locks s) t ln, ((thr s)(t \ (x, no_wait_locks)), shr s), wset s, interrupts s)" and tst: "thr s t = \(x, ln)\" and wst: "\ waiting (wset s t)" by auto from t'dead show ?thesis proof(coinduct) case (deadlocked t'') note t''dead = this with Red have t''t: "t'' \ t" by(auto dest: red_no_deadlock) from t''dead show ?case proof(cases rule: deadlocked_elims) case (lock X) note clnml = \\LT. t'' \ \X, shr s\ LT \ \ \lt \ LT. must_wait s t'' lt (deadlocked s \ final_threads s)\ note tst'' = \thr s t'' = \(X, no_wait_locks)\\ with s' t''t have ts't'': "thr s' t'' = \(X, no_wait_locks)\" by simp moreover { fix LT assume "t'' \ \X, shr s'\ LT \" hence "t'' \ \X, shr s\ LT \" using s' by simp then obtain lt where lt: "lt \ LT" and hlnft: "must_wait s t'' lt (deadlocked s \ final_threads s)" by(blast dest: clnml) from hlnft have "must_wait s' t'' lt (deadlocked s \ deadlocked s' \ final_threads s')" proof(cases rule: must_wait_elims) case (lock l' T) from \has_lock (locks s $ l') T\ s' have "has_lock (locks s' $ l') T" by(auto intro: has_lock_has_lock_acquire_locks) moreover note \T \ t''\ moreover from \T \ deadlocked s \ final_threads s\ have "T \ deadlocked s \ deadlocked s' \ final_threads s'" using subset by blast ultimately show ?thesis unfolding \lt = Inl l'\ .. next case (join T) from \not_final_thread s T\ have "thr s T \ None" by(auto simp add: not_final_thread_iff) moreover from \T \ deadlocked s \ final_threads s\ have "T \ t" proof assume "T \ deadlocked s" with Red show ?thesis by(auto dest: red_no_deadlock) next assume "T \ final_threads s" with \0 < ln $ n\ tst show ?thesis by(auto simp add: final_thread_def) qed ultimately have "not_final_thread s' T" using \not_final_thread s T\ s' by(auto simp add: not_final_thread_iff) moreover from \T \ deadlocked s \ final_threads s\ have "T \ deadlocked s \ deadlocked s' \ final_threads s'" using subset by blast ultimately show ?thesis unfolding \lt = Inr (Inl T)\ .. next case (interrupt T) from tst wst \0 < ln $ n\ have "not_final_thread s t" by(auto simp add: waiting_def not_final_thread_iff) with \all_final_except s (deadlocked s \ final_threads s)\ have "t \ deadlocked s \ final_threads s" by(rule all_final_exceptD) moreover have "t \ deadlocked s" using Red by(blast dest: red_no_deadlock) moreover have "\ final_thread s t" using tst \0 < ln $ n\ by(auto simp add: final_thread_def) ultimately have False by blast thus ?thesis .. qed with lt have "\lt\LT. must_wait s' t'' lt (deadlocked s \ deadlocked s' \ final_threads s')" by blast } moreover from \wset s t'' = None\ s' have "wset s' t'' = None" by simp ultimately show ?thesis using \thr s t'' = \(X, no_wait_locks)\\ \t'' \ \X, shr s\ \\ s' by fastforce next case (wait X LN) have "all_final_except s' (deadlocked s)" proof(rule all_final_exceptI) fix T assume "not_final_thread s' T" hence "not_final_thread s T" using wst tst s' by(auto simp add: not_final_thread_iff split: if_split_asm) with \all_final_except s (deadlocked s)\ \thr s t = \(x, ln)\\ show "T \ deadlocked s" by-(erule all_final_exceptD) qed hence "all_final_except s' (deadlocked s \ deadlocked s')" by(rule all_final_except_mono') blast with t''t \thr s t'' = \(X, LN)\\ \waiting (wset s t'')\ s' have ?Wait by simp thus ?thesis by simp next case (acquire X LN l T) from \thr s t'' = \(X, LN)\\ t''t s' have "thr s' t'' = \(X, LN)\" by(simp) moreover from \T \ deadlocked s \ final_thread s T\ s' tst have "T \ deadlocked s \ final_thread s' T" by(clarsimp simp add: final_thread_def) moreover from \has_lock (locks s $ l) T\ s' have "has_lock (locks s' $ l) T" by(auto intro: has_lock_has_lock_acquire_locks) moreover have "\ waiting (wset s' t'')" using \\ waiting (wset s t'')\ s' by simp ultimately show ?thesis using \0 < LN $ l\ \t'' \ T\ by blast qed qed qed qed corollary RedT_deadlocked_subset: assumes wfs: "s \ wf_state" and Red: "s -\ttas\* s'" shows "deadlocked s \ deadlocked s'" using Red apply(induct rule: RedT_induct') apply(unfold RedT_def) apply(blast dest: invariant3p_rtrancl3p[OF invariant3p_wf_state _ wfs] redT_deadlocked_subset)+ done end end diff --git a/thys/JinjaThreads/Framework/FWLiftingSem.thy b/thys/JinjaThreads/Framework/FWLiftingSem.thy --- a/thys/JinjaThreads/Framework/FWLiftingSem.thy +++ b/thys/JinjaThreads/Framework/FWLiftingSem.thy @@ -1,204 +1,204 @@ (* Title: JinjaThreads/Framework/FWLiftingSem.thy Author: Andreas Lochbihler *) section \Semantic properties of lifted predicates\ theory FWLiftingSem imports FWSemantics FWLifting begin context multithreaded_base begin lemma redT_preserves_ts_inv_ok: "\ s -t\ta\ s'; ts_inv_ok (thr s) I \ \ ts_inv_ok (thr s') (upd_invs I P \ta\\<^bsub>t\<^esub>)" by(erule redT.cases)(fastforce intro: ts_inv_ok_upd_invs ts_inv_ok_upd_ts redT_updTs_Some)+ lemma RedT_preserves_ts_inv_ok: "\ s -\ttas\* s'; ts_inv_ok (thr s) I \ \ ts_inv_ok (thr s') (upd_invs I Q (concat (map (thr_a \ snd) ttas)))" by(induct rule: RedT_induct)(auto intro: redT_preserves_ts_inv_ok) lemma redT_upd_inv_ext: fixes I :: "'t \ 'i" shows "\ s -t\ta\ s'; ts_inv_ok (thr s) I \ \ I \\<^sub>m upd_invs I P \ta\\<^bsub>t\<^esub>" by(erule redT.cases, auto intro: ts_inv_ok_inv_ext_upd_invs) lemma RedT_upd_inv_ext: fixes I :: "'t \ 'i" shows "\ s -\ttas\* s'; ts_inv_ok (thr s) I \ \ I \\<^sub>m upd_invs I P (concat (map (thr_a \ snd) ttas))" proof(induct rule: RedT_induct) case refl thus ?case by simp next case (step S TTAS S' T TA S'') hence "ts_inv_ok (thr S') (upd_invs I P (concat (map (thr_a \ snd) TTAS)))" by -(rule RedT_preserves_ts_inv_ok) hence "upd_invs I P (concat (map (thr_a \ snd) TTAS)) \\<^sub>m upd_invs (upd_invs I P (concat (map (thr_a \ snd) TTAS))) P \TA\\<^bsub>t\<^esub>" using step by -(rule redT_upd_inv_ext) with step show ?case by(auto elim!: map_le_trans simp add: comp_def) qed end locale lifting_inv = multithreaded final r convert_RA for final :: "'x \ bool" and r :: "('l,'t,'x,'m,'w,'o) semantics" ("_ \ _ -_\ _" [50,0,0,50] 80) and convert_RA :: "'l released_locks \ 'o list" + fixes P :: "'i \ 't \ 'x \ 'm \ bool" assumes invariant_red: "\ t \ \x, m\ -ta\ \x', m'\; P i t x m \ \ P i t x' m'" and invariant_NewThread: "\ t \ \x, m\ -ta\ \x', m'\; P i t x m; NewThread t'' x'' m' \ set \ta\\<^bsub>t\<^esub> \ \ \i''. P i'' t'' x'' m'" and invariant_other: "\ t \ \x, m\ -ta\ \x', m'\; P i t x m; P i'' t'' x'' m \ \ P i'' t'' x'' m'" begin lemma redT_updTs_invariant: fixes ln assumes tsiP: "ts_inv P I ts m" and red: "t \ \x, m\ -ta\ \x', m'\" and tao: "thread_oks ts \ta\\<^bsub>t\<^esub>" and tst: "ts t = \(x, ln)\" - shows "ts_inv P (upd_invs I P \ta\\<^bsub>t\<^esub>) (redT_updTs ts \ta\\<^bsub>t\<^esub>(t \ (x', ln'))) m'" + shows "ts_inv P (upd_invs I P \ta\\<^bsub>t\<^esub>) ((redT_updTs ts \ta\\<^bsub>t\<^esub>)(t \ (x', ln'))) m'" proof(rule ts_invI) fix T X LN - assume XLN: "(redT_updTs ts \ta\\<^bsub>t\<^esub>(t \ (x', ln'))) T = \(X, LN)\" + assume XLN: "((redT_updTs ts \ta\\<^bsub>t\<^esub>)(t \ (x', ln'))) T = \(X, LN)\" from tsiP \ts t = \(x, ln)\\ obtain i where "I t = \i\" "P i t x m" by(auto dest: ts_invD) show "\i. upd_invs I P \ta\\<^bsub>t\<^esub> T = \i\ \ P i T X m'" proof(cases "T = t") case True from red \P i t x m\ have "P i t x' m'" by(rule invariant_red) moreover from \I t = \i\\ \ts t = \(x, ln)\\ tao have "upd_invs I P \ta\\<^bsub>t\<^esub> t = \i\" by(simp add: upd_invs_Some) ultimately show ?thesis using True XLN by simp next case False show ?thesis proof(cases "ts T") case None with XLN tao False have "\m'. NewThread T X m' \ set \ta\\<^bsub>t\<^esub>" by(auto dest: redT_updTs_new_thread) with red have nt: "NewThread T X m' \ set \ta\\<^bsub>t\<^esub>" by(auto dest: new_thread_memory) with red \P i t x m\ have "\i''. P i'' T X m'" by(rule invariant_NewThread) hence "P (SOME i. P i T X m') T X m'" by(rule someI_ex) with nt tao show ?thesis by(auto intro: SOME_new_thread_upd_invs) next case (Some a) obtain X' LN' where [simp]: "a = (X', LN')" by(cases a) with \ts T = \a\\ have esT: "ts T = \(X', LN')\" by simp hence "redT_updTs ts \ta\\<^bsub>t\<^esub> T = \(X', LN')\" using \thread_oks ts \ta\\<^bsub>t\<^esub>\ by(auto intro: redT_updTs_Some) moreover from esT tsiP obtain i' where "I T = \i'\" "P i' T X' m" by(auto dest: ts_invD) from red \P i t x m\ \P i' T X' m\ have "P i' T X' m'" by(rule invariant_other) moreover from \I T = \i'\\ esT tao have "upd_invs I P \ta\\<^bsub>t\<^esub> T = \i'\" by(simp add: upd_invs_Some) ultimately show ?thesis using XLN False by simp qed qed qed theorem redT_invariant: assumes redT: "s -t\ta\ s'" and esinvP: "ts_inv P I (thr s) (shr s)" shows "ts_inv P (upd_invs I P \ta\\<^bsub>t\<^esub>) (thr s') (shr s')" using redT proof(cases rule: redT_elims) case acquire thus ?thesis using esinvP by(auto intro!: ts_invI split: if_split_asm dest: ts_invD) next case (normal x x' m') with esinvP - have "ts_inv P (upd_invs I P \ta\\<^bsub>t\<^esub>) (redT_updTs (thr s) \ta\\<^bsub>t\<^esub>(t \ (x', redT_updLns (locks s) t no_wait_locks \ta\\<^bsub>l\<^esub>))) m'" + have "ts_inv P (upd_invs I P \ta\\<^bsub>t\<^esub>) ((redT_updTs (thr s) \ta\\<^bsub>t\<^esub>)(t \ (x', redT_updLns (locks s) t no_wait_locks \ta\\<^bsub>l\<^esub>))) m'" by(auto intro: redT_updTs_invariant) thus ?thesis using normal by simp qed theorem RedT_invariant: assumes RedT: "s -\ttas\* s'" and esinvQ: "ts_inv P I (thr s) (shr s)" shows "ts_inv P (upd_invs I P (concat (map (thr_a \ snd) ttas))) (thr s') (shr s')" using RedT esinvQ proof(induct rule: RedT_induct) case refl thus ?case by(simp (no_asm)) next case (step S TTAS S' T TA S'') note IH = \ts_inv P I (thr S) (shr S) \ ts_inv P (upd_invs I P (concat (map (thr_a \ snd) TTAS))) (thr S') (shr S')\ with \ts_inv P I (thr S) (shr S)\ have "ts_inv P (upd_invs I P (concat (map (thr_a \ snd) TTAS))) (thr S') (shr S')" by blast with \S' -T\TA\ S''\ have "ts_inv P (upd_invs (upd_invs I P (concat (map (thr_a \ snd) TTAS))) P \TA\\<^bsub>t\<^esub>) (thr S'') (shr S'')" by(rule redT_invariant) thus ?case by(simp add: comp_def) qed lemma invariant3p_ts_inv: "invariant3p redT {s. \I. ts_inv P I (thr s) (shr s)}" by(auto intro!: invariant3pI dest: redT_invariant) end locale lifting_wf = multithreaded final r convert_RA for final :: "'x \ bool" and r :: "('l,'t,'x,'m,'w,'o) semantics" ("_ \ _ -_\ _" [50,0,0,50] 80) and convert_RA :: "'l released_locks \ 'o list" + fixes P :: "'t \ 'x \ 'm \ bool" assumes preserves_red: "\ t \ \x, m\ -ta\ \x', m'\; P t x m \ \ P t x' m'" and preserves_NewThread: "\ t \ \x, m\ -ta\ \x', m'\; P t x m; NewThread t'' x'' m' \ set \ta\\<^bsub>t\<^esub> \ \ P t'' x'' m'" and preserves_other: "\ t \ \x, m\ -ta\ \x', m'\; P t x m; P t'' x'' m \ \ P t'' x'' m'" begin lemma lifting_inv: "lifting_inv final r (\_ :: unit. P)" by(unfold_locales)(blast intro: preserves_red preserves_NewThread preserves_other)+ lemma redT_updTs_preserves: fixes ln assumes esokQ: "ts_ok P ts m" and red: "t \ \x, m\ -ta\ \x', m'\" and "ts t = \(x, ln)\" and "thread_oks ts \ta\\<^bsub>t\<^esub>" - shows "ts_ok P (redT_updTs ts \ta\\<^bsub>t\<^esub>(t \ (x', ln'))) m'" + shows "ts_ok P ((redT_updTs ts \ta\\<^bsub>t\<^esub>)(t \ (x', ln'))) m'" proof - interpret lifting_inv final r convert_RA "\_ :: unit. P" by(rule lifting_inv) from esokQ obtain I :: "'t \ unit" where "ts_inv (\_. P) I ts m" by(rule ts_ok_into_ts_inv_const) - hence "ts_inv (\_. P) (upd_invs I (\_. P) \ta\\<^bsub>t\<^esub>) (redT_updTs ts \ta\\<^bsub>t\<^esub>(t \ (x', ln'))) m'" + hence "ts_inv (\_. P) (upd_invs I (\_. P) \ta\\<^bsub>t\<^esub>) ((redT_updTs ts \ta\\<^bsub>t\<^esub>)(t \ (x', ln'))) m'" using red \thread_oks ts \ta\\<^bsub>t\<^esub>\ \ts t = \(x, ln)\\ by(rule redT_updTs_invariant) thus ?thesis by(rule ts_inv_const_into_ts_ok) qed theorem redT_preserves: assumes redT: "s -t\ta\ s'" and esokQ: "ts_ok P (thr s) (shr s)" shows "ts_ok P (thr s') (shr s')" proof - interpret lifting_inv final r convert_RA "\_ :: unit. P" by(rule lifting_inv) from esokQ obtain I :: "'t \ unit" where "ts_inv (\_. P) I (thr s) (shr s)" by(rule ts_ok_into_ts_inv_const) with redT have "ts_inv (\_. P) (upd_invs I (\_. P) \ta\\<^bsub>t\<^esub>) (thr s') (shr s')" by(rule redT_invariant) thus ?thesis by(rule ts_inv_const_into_ts_ok) qed theorem RedT_preserves: "\ s -\ttas\* s'; ts_ok P (thr s) (shr s) \ \ ts_ok P (thr s') (shr s')" by(erule (1) RedT_lift_preserveD)(fastforce elim: redT_preserves) lemma invariant3p_ts_ok: "invariant3p redT {s. ts_ok P (thr s) (shr s)}" by(auto intro!: invariant3pI intro: redT_preserves) end lemma lifting_wf_Const [intro!]: assumes "multithreaded final r" shows "lifting_wf final r (\t x m. k)" proof - interpret multithreaded final r using assms . show ?thesis by unfold_locales blast+ qed end diff --git a/thys/JinjaThreads/Framework/FWSemantics.thy b/thys/JinjaThreads/Framework/FWSemantics.thy --- a/thys/JinjaThreads/Framework/FWSemantics.thy +++ b/thys/JinjaThreads/Framework/FWSemantics.thy @@ -1,420 +1,420 @@ (* Title: JinjaThreads/Framework/FWSemantics.thy Author: Andreas Lochbihler *) section \The multithreaded semantics\ theory FWSemantics imports FWWellform FWLockingThread FWCondAction FWInterrupt begin inductive redT_upd :: "('l,'t,'x,'m,'w) state \ 't \ ('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm \ ('l,'t,'x,'m,'w) state \ bool" for s t ta x' m' where "redT_updWs t (wset s) \ta\\<^bsub>w\<^esub> ws' \ redT_upd s t ta x' m' (redT_updLs (locks s) t \ta\\<^bsub>l\<^esub>, ((redT_updTs (thr s) \ta\\<^bsub>t\<^esub>)(t \ (x', redT_updLns (locks s) t (snd (the (thr s t))) \ta\\<^bsub>l\<^esub>)), m'), ws', redT_updIs (interrupts s) \ta\\<^bsub>i\<^esub>)" inductive_simps redT_upd_simps [simp]: "redT_upd s t ta x' m' s'" definition redT_acq :: "('l,'t,'x,'m,'w) state \ 't \ ('l \f nat) \ ('l,'t,'x,'m,'w) state" where "\ln. redT_acq s t ln = (acquire_all (locks s) t ln, ((thr s)(t \ (fst (the (thr s t)), no_wait_locks)), shr s), wset s, interrupts s)" context final_thread begin inductive actions_ok :: "('l,'t,'x,'m,'w) state \ 't \ ('l,'t,'x','m,'w,'o) thread_action \ bool" for s :: "('l,'t,'x,'m,'w) state" and t :: 't and ta :: "('l,'t,'x','m,'w,'o) thread_action" where "\ 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> \ \ actions_ok s t ta" declare actions_ok.intros [intro!] declare actions_ok.cases [elim!] lemma actions_ok_iff [simp]: "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>" by(auto) lemma actions_ok_thread_oksD: "actions_ok s t ta \ thread_oks (thr s) \ta\\<^bsub>t\<^esub>" by(erule actions_ok.cases) inductive actions_ok' :: "('l,'t,'x,'m,'w) state \ 't \ ('l,'t,'x','m,'w,'o) thread_action \ bool" where "\ 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> \ \ actions_ok' s t ta" declare actions_ok'.intros [intro!] declare actions_ok'.cases [elim!] lemma actions_ok'_iff: "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>" by auto lemma actions_ok'_ta_upd_obs: "actions_ok' s t (ta_update_obs ta obs) \ actions_ok' s t ta" by(auto simp add: actions_ok'_iff lock_ok_las'_def ta_upd_simps wset_actions_ok_def) lemma actions_ok'_empty: "actions_ok' s t \ \ wset s t = None" by(simp add: actions_ok'_iff lock_ok_las'_def) lemma actions_ok'_convert_extTA: "actions_ok' s t (convert_extTA f ta) = actions_ok' s t ta" by(simp add: actions_ok'_iff) inductive actions_subset :: "('l,'t,'x,'m,'w,'o) thread_action \ ('l,'t,'x','m,'w,'o) thread_action \ bool" where "\ 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> \ \ actions_subset ta' ta" declare actions_subset.intros [intro!] declare actions_subset.cases [elim!] lemma actions_subset_iff: "actions_subset ta' ta \ 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>" by auto lemma actions_subset_refl [intro]: "actions_subset ta ta" by(auto intro: actions_subset.intros collect_locks'_subset_collect_locks del: subsetI) definition final_thread :: "('l,'t,'x,'m,'w) state \ 't \ bool" where "\ln. final_thread s t \ (case thr s t of None \ False | \(x, ln)\ \ final x \ ln = no_wait_locks \ wset s t = None)" definition final_threads :: "('l,'t,'x,'m,'w) state \ 't set" where "final_threads s \ {t. final_thread s t}" lemma [iff]: "t \ final_threads s = final_thread s t" by (simp add: final_threads_def) lemma [pred_set_conv]: "final_thread s = (\t. t \ final_threads s)" by simp definition mfinal :: "('l,'t,'x,'m,'w) state \ bool" where "mfinal s \ (\t x ln. thr s t = \(x, ln)\ \ final x \ ln = no_wait_locks \ wset s t = None)" lemma final_threadI: "\ thr s t = \(x, no_wait_locks)\; final x; wset s t = None \ \ final_thread s t" by(simp add: final_thread_def) lemma final_threadE: assumes "final_thread s t" obtains x where "thr s t = \(x, no_wait_locks)\" "final x" "wset s t = None" using assms by(auto simp add: final_thread_def) lemma mfinalI: "(\t x ln. thr s t = \(x, ln)\ \ final x \ ln = no_wait_locks \ wset s t = None) \ mfinal s" unfolding mfinal_def by blast lemma mfinalD: fixes ln assumes "mfinal s" "thr s t = \(x, ln)\" shows "final x" "ln = no_wait_locks" "wset s t = None" using assms unfolding mfinal_def by blast+ lemma mfinalE: fixes ln assumes "mfinal s" "thr s t = \(x, ln)\" obtains "final x" "ln = no_wait_locks" "wset s t = None" using mfinalD[OF assms] by(rule that) lemma mfinal_def2: "mfinal s \ dom (thr s) \ final_threads s" by(fastforce elim: mfinalE final_threadE intro: mfinalI final_threadI) end locale multithreaded_base = final_thread + constrains final :: "'x \ bool" fixes r :: "('l,'t,'x,'m,'w,'o) semantics" ("_ \ _ -_\ _" [50,0,0,50] 80) and convert_RA :: "'l released_locks \ 'o list" begin abbreviation r_syntax :: "'t \ 'x \ 'm \ ('l,'t,'x,'m,'w,'o) thread_action \ 'x \ 'm \ bool" ("_ \ \_, _\ -_\ \_, _\" [50,0,0,0,0,0] 80) where "t \ \x, m\ -ta\ \x', m'\ \ t \ (x, m) -ta\ (x', m')" inductive redT :: "('l,'t,'x,'m,'w) state \ 't \ ('l,'t,'x,'m,'w,'o) thread_action \ ('l,'t,'x,'m,'w) state \ bool" and redT_syntax1 :: "('l,'t,'x,'m,'w) state \ 't \ ('l,'t,'x,'m,'w,'o) thread_action \ ('l,'t,'x,'m,'w) state \ bool" ("_ -_\_\ _" [50,0,0,50] 80) where "s -t\ta\ s' \ redT s (t, ta) s'" | redT_normal: "\ t \ \x, shr s\ -ta\ \x', m'\; thr s t = \(x, no_wait_locks)\; actions_ok s t ta; redT_upd s t ta x' m' s' \ \ s -t\ta\ s'" | redT_acquire: "\ln. \ thr s t = \(x, ln)\; \ waiting (wset s t); may_acquire_all (locks s) t ln; ln $ n > 0; - s' = (acquire_all (locks s) t ln, (thr s(t \ (x, no_wait_locks)), shr s), wset s, interrupts s) \ + s' = (acquire_all (locks s) t ln, ((thr s)(t \ (x, no_wait_locks)), shr s), wset s, interrupts s) \ \ s -t\((K$ []), [], [], [], [], convert_RA ln)\ s'" abbreviation redT_syntax2 :: "('l,'t) locks \ ('l,'t,'x) thread_info \ 'm \ ('w,'t) wait_sets \ 't interrupts \ 't \ ('l,'t,'x,'m,'w,'o) thread_action \ ('l,'t) locks \ ('l,'t,'x) thread_info \ 'm \ ('w,'t) wait_sets \ 't interrupts \ bool" ("\_, _, _, _\ -_\_\ \_, _, _, _\" [0,0,0,0,0,0,0,0,0] 80) where "\ls, tsm, ws, is\ -t\ta\ \ls', tsm', ws', is'\ \ (ls, tsm, ws, is) -t\ta\ (ls', tsm', ws', is')" lemma redT_elims [consumes 1, case_names normal acquire]: assumes red: "s -t\ta\ s'" and normal: "\x x' m' ws'. \ t \ \x, shr s\ -ta\ \x', m'\; thr s t = \(x, no_wait_locks)\; 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>; redT_updWs t (wset s) \ta\\<^bsub>w\<^esub> ws'; - s' = (redT_updLs (locks s) t \ta\\<^bsub>l\<^esub>, (redT_updTs (thr s) \ta\\<^bsub>t\<^esub>(t \ (x', redT_updLns (locks s) t no_wait_locks \ta\\<^bsub>l\<^esub>)), m'), ws', redT_updIs (interrupts s) \ta\\<^bsub>i\<^esub>) \ + s' = (redT_updLs (locks s) t \ta\\<^bsub>l\<^esub>, ((redT_updTs (thr s) \ta\\<^bsub>t\<^esub>)(t \ (x', redT_updLns (locks s) t no_wait_locks \ta\\<^bsub>l\<^esub>)), m'), ws', redT_updIs (interrupts s) \ta\\<^bsub>i\<^esub>) \ \ thesis" and acquire: "\x ln n. \ thr s t = \(x, ln)\; ta = (K$ [], [], [], [], [], convert_RA ln); \ waiting (wset s t); may_acquire_all (locks s) t ln; 0 < ln $ n; - s' = (acquire_all (locks s) t ln, (thr s(t \ (x, no_wait_locks)), shr s), wset s, interrupts s) \ + s' = (acquire_all (locks s) t ln, ((thr s)(t \ (x, no_wait_locks)), shr s), wset s, interrupts s) \ \ thesis" shows thesis using red proof cases case redT_normal thus ?thesis using normal by(cases s')(auto) next case redT_acquire thus ?thesis by-(rule acquire, fastforce+) qed definition RedT :: "('l,'t,'x,'m,'w) state \ ('t \ ('l,'t,'x,'m,'w,'o) thread_action) list \ ('l,'t,'x,'m,'w) state \ bool" ("_ -\_\* _" [50,0,50] 80) where "RedT \ rtrancl3p redT" lemma RedTI: "rtrancl3p redT s ttas s' \ RedT s ttas s'" by(simp add: RedT_def) lemma RedTE: "\ RedT s ttas s'; rtrancl3p redT s ttas s' \ P \ \ P" by(auto simp add: RedT_def) lemma RedTD: "RedT s ttas s' \ rtrancl3p redT s ttas s'" by(simp add: RedT_def) lemma RedT_induct [consumes 1, case_names refl step]: "\ s -\ttas\* s'; \s. P s [] s; \s ttas s' t ta s''. \ s -\ttas\* s'; P s ttas s'; s' -t\ta\ s'' \ \ P s (ttas @ [(t, ta)]) s''\ \ P s ttas s'" unfolding RedT_def by(erule rtrancl3p.induct) auto lemma RedT_induct' [consumes 1, case_names refl step]: "\ s -\ttas\* s'; P s [] s; \ttas s' t ta s''. \ s -\ttas\* s'; P s ttas s'; s' -t\ta\ s'' \ \ P s (ttas @ [(t, ta)]) s''\ \ P s ttas s'" unfolding RedT_def apply(erule rtrancl3p_induct', blast) apply(case_tac b, blast) done lemma RedT_lift_preserveD: assumes Red: "s -\ttas\* s'" and P: "P s" and preserve: "\s t tas s'. \ s -t\tas\ s'; P s \ \ P s'" shows "P s'" using Red P by(induct rule: RedT_induct)(auto intro: preserve) lemma RedT_refl [intro, simp]: "s -\[]\* s" by(rule RedTI)(rule rtrancl3p_refl) lemma redT_has_locks_inv: "\ \ls, (ts, m), ws, is\ -t\ta\ \ls', (ts', m'), ws', is'\; t \ t' \ \ has_locks (ls $ l) t' = has_locks (ls' $ l) t'" by(auto elim!: redT.cases intro: redT_updLs_has_locks[THEN sym, simplified] may_acquire_all_has_locks_acquire_locks[symmetric]) lemma redT_has_lock_inv: "\ \ls, (ts, m), ws, is\ -t\ta\ \ls', (ts', m'), ws', is'\; t \ t' \ \ has_lock (ls' $ l) t' = has_lock (ls $ l) t'" by(auto simp add: redT_has_locks_inv) lemma redT_ts_Some_inv: "\ \ls, (ts, m), ws, is\ -t\ta\ \ls', (ts', m'), ws', is'\; t \ t'; ts t' = \x\ \ \ ts' t' = \x\" by(fastforce elim!: redT.cases simp: redT_updTs_upd[THEN sym] intro: redT_updTs_Some) lemma redT_thread_not_disappear: "\ s -t\ta\ s'; thr s' t' = None\ \ thr s t' = None" apply(cases "t \ t'") apply(auto elim!: redT_elims simp add: redT_updTs_upd[THEN sym] intro: redT_updTs_None) done lemma RedT_thread_not_disappear: "\ s -\ttas\* s'; thr s' t' = None\ \ thr s t' = None" apply(erule contrapos_pp[where Q="thr s' t' = None"]) apply(drule (1) RedT_lift_preserveD) apply(erule_tac Q="thr sa t' = None" in contrapos_nn) apply(erule redT_thread_not_disappear) apply(auto) done lemma redT_preserves_wset_thread_ok: "\ s -t\ta\ s'; wset_thread_ok (wset s) (thr s) \ \ wset_thread_ok (wset s') (thr s')" by(fastforce elim!: redT.cases intro: wset_thread_ok_upd redT_updTs_preserves_wset_thread_ok redT_updWs_preserve_wset_thread_ok) lemma RedT_preserves_wset_thread_ok: "\ s -\ttas\* s'; wset_thread_ok (wset s) (thr s) \ \ wset_thread_ok (wset s') (thr s')" by(erule (1) RedT_lift_preserveD)(erule redT_preserves_wset_thread_ok) lemma redT_new_thread_ts_Some: "\ s -t\ta\ s'; NewThread t' x m'' \ set \ta\\<^bsub>t\<^esub>; wset_thread_ok (wset s) (thr s) \ \ thr s' t' = \(x, no_wait_locks)\" by(erule redT_elims)(auto dest: thread_oks_new_thread elim: redT_updTs_new_thread_ts) lemma RedT_new_thread_ts_not_None: "\ s -\ttas\* s'; NewThread t x m'' \ set (concat (map (thr_a \ snd) ttas)); wset_thread_ok (wset s) (thr s) \ \ thr s' t \ None" proof(induct rule: RedT_induct) case refl thus ?case by simp next case (step S TTAS S' T TA S'') note Red = \S -\TTAS\* S'\ note IH = \\ NewThread t x m'' \ set (concat (map (thr_a \ snd) TTAS)); wset_thread_ok (wset S) (thr S) \ \ thr S' t \ None\ note red = \S' -T\TA\ S''\ note ins = \NewThread t x m'' \ set (concat (map (thr_a \ snd) (TTAS @ [(T, TA)])))\ note wto = \wset_thread_ok (wset S) (thr S)\ from Red wto have wto': "wset_thread_ok (wset S') (thr S')" by(auto dest: RedT_preserves_wset_thread_ok) show ?case proof(cases "NewThread t x m'' \ set \TA\\<^bsub>t\<^esub>") case True thus ?thesis using red wto' by(auto dest!: redT_new_thread_ts_Some) next case False hence "NewThread t x m'' \ set (concat (map (thr_a \ snd) TTAS))" using ins by(auto) hence "thr S' t \ None" using wto by(rule IH) with red show ?thesis by -(erule contrapos_nn, auto dest: redT_thread_not_disappear) qed qed lemma redT_preserves_lock_thread_ok: "\ s -t\ta\ s'; lock_thread_ok (locks s) (thr s) \ \ lock_thread_ok (locks s') (thr s')" by(auto elim!: redT_elims intro: redT_upds_preserves_lock_thread_ok acquire_all_preserves_lock_thread_ok) lemma RedT_preserves_lock_thread_ok: "\ s -\ttas\* s'; lock_thread_ok (locks s) (thr s) \ \ lock_thread_ok (locks s') (thr s')" by(erule (1) RedT_lift_preserveD)(erule redT_preserves_lock_thread_ok) lemma redT_ex_new_thread: assumes "s -t'\ta\ s'" "wset_thread_ok (wset s) (thr s)" "thr s' t = \(x, w)\" "thr s t = None" shows "\m. NewThread t x m \ set \ta\\<^bsub>t\<^esub> \ w = no_wait_locks" using assms by cases (fastforce split: if_split_asm dest: wset_thread_okD redT_updTs_new_thread)+ lemma redT_ex_new_thread': assumes "s -t'\ta\ s'" "thr s' t = \(x, w)\" "thr s t = None" shows "\m x. NewThread t x m \ set \ta\\<^bsub>t\<^esub>" using assms by(cases)(fastforce split: if_split_asm dest!: redT_updTs_new_thread)+ definition deterministic :: "('l,'t,'x,'m,'w) state set \ bool" where "deterministic I \ (\s t x ta' x' m' ta'' x'' m''. s \ I \ thr s t = \(x, no_wait_locks)\ \ t \ \x, shr s\ -ta'\ \x', m'\ \ t \ \x, shr s\ -ta''\ \x'', m''\ \ actions_ok s t ta' \ actions_ok s t ta'' \ ta' = ta'' \ x' = x'' \ m' = m'') \ invariant3p redT I" lemma determisticI: "\\s t x ta' x' m' ta'' x'' m''. \ s \ I; thr s t = \(x, no_wait_locks)\; t \ \x, shr s\ -ta'\ \x', m'\; t \ \x, shr s\ -ta''\ \x'', m''\; actions_ok s t ta'; actions_ok s t ta'' \ \ ta' = ta'' \ x' = x'' \ m' = m''; invariant3p redT I \ \ deterministic I" unfolding deterministic_def by blast lemma deterministicD: "\ deterministic I; t \ \x, shr s\ -ta'\ \x', m'\; t \ \x, shr s\ -ta''\ \x'', m''\; thr s t = \(x, no_wait_locks)\; actions_ok s t ta'; actions_ok s t ta''; s \ I \ \ ta' = ta'' \ x' = x'' \ m' = m''" unfolding deterministic_def by blast lemma deterministic_invariant3p: "deterministic I \ invariant3p redT I" unfolding deterministic_def by blast lemma deterministic_THE: "\ deterministic I; thr s t = \(x, no_wait_locks)\; t \ \x, shr s\ -ta\ \x', m'\; actions_ok s t ta; s \ I \ \ (THE (ta, x', m'). t \ \x, shr s\ -ta\ \x', m'\ \ actions_ok s t ta) = (ta, x', m')" by(rule the_equality)(blast dest: deterministicD)+ end locale multithreaded = multithreaded_base + constrains final :: "'x \ bool" and r :: "('l,'t,'x,'m,'w,'o) semantics" and convert_RA :: "'l released_locks \ 'o list" assumes new_thread_memory: "\ t \ s -ta\ s'; NewThread t' x m \ set \ta\\<^bsub>t\<^esub> \ \ m = snd s'" and final_no_red: "\ t \ (x, m) -ta\ (x', m'); final x \ \ False" begin lemma redT_new_thread_common: "\ s -t\ta\ s'; NewThread t' x m'' \ set \ta\\<^bsub>t\<^esub>; \ta\\<^bsub>w\<^esub> = [] \ \ m'' = shr s'" by(auto elim!: redT_elims rtrancl3p_cases dest: new_thread_memory) lemma redT_new_thread: assumes "s -t'\ta\ s'" "thr s' t = \(x, w)\" "thr s t = None" "\ta\\<^bsub>w\<^esub> = []" shows "NewThread t x (shr s') \ set \ta\\<^bsub>t\<^esub> \ w = no_wait_locks" using assms apply(cases rule: redT_elims) apply(auto split: if_split_asm del: conjI elim!: rtrancl3p_cases) apply(drule (2) redT_updTs_new_thread) apply(auto dest: new_thread_memory) done lemma final_no_redT: "\ s -t\ta\ s'; thr s t = \(x, no_wait_locks)\ \ \ \ final x" by(auto elim!: redT_elims dest: final_no_red) lemma mfinal_no_redT: assumes redT: "s -t\ta\ s'" and mfinal: "mfinal s" shows False using redT mfinalD[OF mfinal, of t] by cases (metis final_no_red, metis neq_no_wait_locks_conv) end end diff --git a/thys/JinjaThreads/Framework/FWThread.thy b/thys/JinjaThreads/Framework/FWThread.thy --- a/thys/JinjaThreads/Framework/FWThread.thy +++ b/thys/JinjaThreads/Framework/FWThread.thy @@ -1,247 +1,247 @@ (* Title: JinjaThreads/Framework/FWThread.thy Author: Andreas Lochbihler *) section \Semantics of the thread actions for thread creation\ theory FWThread imports FWState begin text\Abstractions for thread ids\ context notes [[inductive_internals]] begin inductive free_thread_id :: "('l,'t,'x) thread_info \ 't \ bool" for ts :: "('l,'t,'x) thread_info" and t :: 't where "ts t = None \ free_thread_id ts t" declare free_thread_id.cases [elim] end lemma free_thread_id_iff: "free_thread_id ts t = (ts t = None)" by(auto elim: free_thread_id.cases intro: free_thread_id.intros) text\Update functions for the multithreaded state\ fun redT_updT :: "('l,'t,'x) thread_info \ ('t,'x,'m) new_thread_action \ ('l,'t,'x) thread_info" where "redT_updT ts (NewThread t' x m) = ts(t' \ (x, no_wait_locks))" | "redT_updT ts _ = ts" fun redT_updTs :: "('l,'t,'x) thread_info \ ('t,'x,'m) new_thread_action list \ ('l,'t,'x) thread_info" where "redT_updTs ts [] = ts" | "redT_updTs ts (ta#tas) = redT_updTs (redT_updT ts ta) tas" lemma redT_updTs_append [simp]: "redT_updTs ts (tas @ tas') = redT_updTs (redT_updTs ts tas) tas'" by(induct ts tas rule: redT_updTs.induct) auto lemma redT_updT_None: "redT_updT ts ta t = None \ ts t = None" by(cases ta)(auto split: if_splits) lemma redT_updTs_None: "redT_updTs ts tas t = None \ ts t = None" by(induct ts tas rule: redT_updTs.induct)(auto intro: redT_updT_None) lemma redT_updT_Some1: "ts t = \xw\ \ \xw. redT_updT ts ta t = \xw\" by(cases ta) auto lemma redT_updTs_Some1: "ts t = \xw\ \ \xw. redT_updTs ts tas t = \xw\" unfolding not_None_eq[symmetric] by(induct ts tas arbitrary: xw rule: redT_updTs.induct)(simp_all del: split_paired_Ex, blast dest: redT_updT_Some1) lemma redT_updT_finite_dom_inv: "finite (dom (redT_updT ts ta)) = finite (dom ts)" by(cases ta) auto lemma redT_updTs_finite_dom_inv: "finite (dom (redT_updTs ts tas)) = finite (dom ts)" by(induct ts tas rule: redT_updTs.induct)(simp_all add: redT_updT_finite_dom_inv) text\Preconditions for thread creation actions\ text\These primed versions are for checking preconditions only. They allow the thread actions to have a type for thread-local information that is different than the thread info state itself.\ fun redT_updT' :: "('l,'t,'x) thread_info \ ('t,'x','m) new_thread_action \ ('l,'t,'x) thread_info" where "redT_updT' ts (NewThread t' x m) = ts(t' \ (undefined, no_wait_locks))" | "redT_updT' ts _ = ts" fun redT_updTs' :: "('l,'t,'x) thread_info \ ('t,'x','m) new_thread_action list \ ('l,'t,'x) thread_info" where "redT_updTs' ts [] = ts" | "redT_updTs' ts (ta#tas) = redT_updTs' (redT_updT' ts ta) tas" lemma redT_updT'_None: "redT_updT' ts ta t = None \ ts t = None" by(cases ta)(auto split: if_splits) primrec thread_ok :: "('l,'t,'x) thread_info \ ('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)" fun thread_oks :: "('l,'t,'x) thread_info \ ('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)" lemma thread_ok_ts_change: "(\t. ts t = None \ ts' t = None) \ thread_ok ts ta \ thread_ok ts' ta" by(cases ta)(auto simp add: free_thread_id_iff) lemma thread_oks_ts_change: "(\t. ts t = None \ ts' t = None) \ thread_oks ts tas \ thread_oks ts' tas" proof(induct tas arbitrary: ts ts') case Nil thus ?case by simp next case (Cons ta tas ts ts') note IH = \\ts ts'. (\t. (ts t = None) = (ts' t = None)) \ thread_oks ts tas = thread_oks ts' tas\ note eq = \\t. (ts t = None) = (ts' t = None)\ from eq have "thread_ok ts ta \ thread_ok ts' ta" by(rule thread_ok_ts_change) moreover from eq have "\t. (redT_updT' ts ta t = None) = (redT_updT' ts' ta t = None)" by(cases ta)(auto) hence "thread_oks (redT_updT' ts ta) tas = thread_oks (redT_updT' ts' ta) tas" by(rule IH) ultimately show ?case by simp qed lemma redT_updT'_eq_None_conv: "(\t. ts t = None \ ts' t = None) \ redT_updT' ts ta t = None \ redT_updT ts' ta t = None" by(cases ta) simp_all lemma redT_updTs'_eq_None_conv: "(\t. ts t = None \ ts' t = None) \ redT_updTs' ts tas t = None \ redT_updTs ts' tas t = None" apply(induct tas arbitrary: ts ts') apply simp_all apply(blast intro: redT_updT'_eq_None_conv del: iffI) done lemma thread_oks_redT_updT_conv [simp]: "thread_oks (redT_updT' ts ta) tas = thread_oks (redT_updT ts ta) tas" by(rule thread_oks_ts_change)(rule redT_updT'_eq_None_conv refl)+ lemma thread_oks_append [simp]: "thread_oks ts (tas @ tas') = (thread_oks ts tas \ thread_oks (redT_updTs' ts tas) tas')" by(induct tas arbitrary: ts, auto) lemma thread_oks_redT_updTs_conv [simp]: "thread_oks (redT_updTs' ts ta) tas = thread_oks (redT_updTs ts ta) tas" by(rule thread_oks_ts_change)(rule redT_updTs'_eq_None_conv refl)+ lemma redT_updT_Some: "\ ts t = \xw\; thread_ok ts ta \ \ redT_updT ts ta t = \xw\" by(cases ta) auto lemma redT_updTs_Some: "\ ts t = \xw\; thread_oks ts tas \ \ redT_updTs ts tas t = \xw\" by(induct ts tas rule: redT_updTs.induct)(auto intro: redT_updT_Some) lemma redT_updT'_Some: "\ ts t = \xw\; thread_ok ts ta \ \ redT_updT' ts ta t = \xw\" by(cases ta) auto lemma redT_updTs'_Some: "\ ts t = \xw\; thread_oks ts tas \ \ redT_updTs' ts tas t = \xw\" by(induct ts tas rule: redT_updTs'.induct)(auto intro: redT_updT'_Some) lemma thread_ok_new_thread: "thread_ok ts (NewThread t m' x) \ ts t = None" by(auto) lemma thread_oks_new_thread: "\ thread_oks ts tas; NewThread t x m \ set tas \ \ ts t = None" by(induct ts tas rule: thread_oks.induct)(auto intro: redT_updT'_None) lemma redT_updT_new_thread_ts: "thread_ok ts (NewThread t x m) \ redT_updT ts (NewThread t x m) t = \(x, no_wait_locks)\" by(simp) lemma redT_updTs_new_thread_ts: "\ thread_oks ts tas; NewThread t x m \ set tas \ \ redT_updTs ts tas t = \(x, no_wait_locks)\" by(induct ts tas rule: redT_updTs.induct)(auto intro: redT_updTs_Some) lemma redT_updT_new_thread: "\ redT_updT ts ta t = \(x, w)\; thread_ok ts ta; ts t = None \ \ \m. ta = NewThread t x m \ w = no_wait_locks" by(cases ta)(auto split: if_split_asm) lemma redT_updTs_new_thread: "\ redT_updTs ts tas t = \(x, w)\; thread_oks ts tas; ts t = None \ \ \m .NewThread t x m \ set tas \ w = no_wait_locks" proof(induct tas arbitrary: ts) case Nil thus ?case by simp next case (Cons TA TAS TS) note IH = \\ts. \redT_updTs ts TAS t = \(x, w)\; thread_oks ts TAS; ts t = None\ \ \m. NewThread t x m \ set TAS \ w = no_wait_locks\ note es't = \redT_updTs TS (TA # TAS) t = \(x, w)\\ note cct = \thread_oks TS (TA # TAS)\ hence cctta: "thread_ok TS TA" and ccts: "thread_oks (redT_updT TS TA) TAS" by auto note est = \TS t = None\ { fix X W assume rest: "redT_updT TS TA t = \(X, W)\" then obtain m where "TA = NewThread t X m \ W = no_wait_locks" using cctta est by (auto dest!: redT_updT_new_thread) then obtain "TA = NewThread t X m" "W = no_wait_locks" .. moreover from rest ccts have "redT_updTs TS (TA # TAS) t = \(X, W)\" by(auto intro:redT_updTs_Some) with es't have "X = x" "W = w" by auto ultimately have ?case by auto } moreover { assume rest: "redT_updT TS TA t = None" hence "\m. TA \ NewThread t x m" using est cct by(clarsimp) with rest ccts es't have ?case by(auto dest: IH) } ultimately show ?case by(cases "redT_updT TS TA t", auto) qed lemma redT_updT_upd: - "\ ts t = \xw\; thread_ok ts ta \ \ redT_updT ts ta(t \ xw') = redT_updT (ts(t \ xw')) ta" + "\ ts t = \xw\; thread_ok ts ta \ \ (redT_updT ts ta)(t \ xw') = redT_updT (ts(t \ xw')) ta" by(cases ta)(fastforce intro: fun_upd_twist)+ lemma redT_updTs_upd: - "\ ts t = \xw\; thread_oks ts tas \ \ redT_updTs ts tas(t \ xw') = redT_updTs (ts(t \ xw')) tas" + "\ ts t = \xw\; thread_oks ts tas \ \ (redT_updTs ts tas)(t \ xw') = redT_updTs (ts(t \ xw')) tas" by(induct ts tas rule: redT_updTs.induct)(auto simp del: fun_upd_apply simp add: redT_updT_upd dest: redT_updT_Some) lemma thread_ok_upd: "ts t = \xln\ \ thread_ok (ts(t \ xln')) ta = thread_ok ts ta" by(rule thread_ok_ts_change) simp lemma thread_oks_upd: "ts t = \xln\ \ thread_oks (ts(t \ xln')) tas = thread_oks ts tas" by(rule thread_oks_ts_change) simp lemma thread_ok_convert_new_thread_action [simp]: "thread_ok ts (convert_new_thread_action f ta) = thread_ok ts ta" by(cases ta) auto lemma redT_updT'_convert_new_thread_action_eq_None: "redT_updT' ts (convert_new_thread_action f ta) t = None \ redT_updT' ts ta t = None" by(cases ta) auto lemma thread_oks_convert_new_thread_action [simp]: "thread_oks ts (map (convert_new_thread_action f) tas) = thread_oks ts tas" by(induct ts tas rule: thread_oks.induct)(simp_all add: thread_oks_ts_change[OF redT_updT'_convert_new_thread_action_eq_None]) lemma map_redT_updT: "map_option (map_prod f id) (redT_updT ts ta t) = redT_updT (\t. map_option (map_prod f id) (ts t)) (convert_new_thread_action f ta) t" by(cases ta) auto lemma map_redT_updTs: "map_option (map_prod f id) (redT_updTs ts tas t) = redT_updTs (\t. map_option (map_prod f id) (ts t)) (map (convert_new_thread_action f) tas) t" by(induct tas arbitrary: ts)(auto simp add: map_redT_updT) end diff --git a/thys/JinjaThreads/Framework/FWWellform.thy b/thys/JinjaThreads/Framework/FWWellform.thy --- a/thys/JinjaThreads/Framework/FWWellform.thy +++ b/thys/JinjaThreads/Framework/FWWellform.thy @@ -1,198 +1,198 @@ (* Title: JinjaThreads/Framework/FWWellform.thy Author: Andreas Lochbihler *) section \Wellformedness conditions for the multithreaded state\ theory FWWellform imports FWLocking FWThread FWWait FWCondAction begin text\Well-formedness property: Locks are held by real threads\ definition lock_thread_ok :: "('l, 't) locks \ ('l, 't,'x) thread_info \ bool" where [code del]: "lock_thread_ok ls ts \ \l t. has_lock (ls $ l) t \ (\xw. ts t = \xw\)" lemma lock_thread_ok_code [code]: "lock_thread_ok ls ts = finfun_All ((\l. case l of None \ True | \(t, n)\ \ (ts t \ None)) \$ ls)" by(simp add: lock_thread_ok_def finfun_All_All has_lock_has_locks_conv has_locks_iff o_def) lemma lock_thread_okI: "(\l t. has_lock (ls $ l) t \ \xw. ts t = \xw\) \ lock_thread_ok ls ts" by(auto simp add: lock_thread_ok_def) lemma lock_thread_okD: "\ lock_thread_ok ls ts; has_lock (ls $ l) t \ \ \xw. ts t = \xw\" by(fastforce simp add: lock_thread_ok_def) lemma lock_thread_okD': "\ lock_thread_ok ls ts; has_locks (ls $ l) t = Suc n \ \ \xw. ts t = \xw\" by(auto elim: lock_thread_okD[where l=l] simp del: split_paired_Ex) lemma lock_thread_okE: "\ lock_thread_ok ls ts; \l t. has_lock (ls $ l) t \ (\xw. ts t = \xw\) \ P \ \ P" by(auto simp add: lock_thread_ok_def simp del: split_paired_Ex) lemma lock_thread_ok_upd: "lock_thread_ok ls ts \ lock_thread_ok ls (ts(t \ xw))" by(auto intro!: lock_thread_okI dest: lock_thread_okD) lemma lock_thread_ok_has_lockE: assumes "lock_thread_ok ls ts" and "has_lock (ls $ l) t" obtains x ln' where "ts t = \(x, ln')\" using assms by(auto dest!: lock_thread_okD) lemma redT_updLs_preserves_lock_thread_ok: assumes lto: "lock_thread_ok ls ts" and tst: "ts t = \xw\" shows "lock_thread_ok (redT_updLs ls t las) ts" proof(rule lock_thread_okI) fix L T assume ru: "has_lock (redT_updLs ls t las $ L) T" show "\xw. ts T = \xw\" proof(cases "t = T") case True thus ?thesis using tst lto by(auto elim: lock_thread_okE) next case False with ru have "has_lock (ls $ L) T" by(rule redT_updLs_Some_thread_idD) thus ?thesis using lto by(auto elim!: lock_thread_okE simp del: split_paired_Ex) qed qed lemma redT_updTs_preserves_lock_thread_ok: assumes lto: "lock_thread_ok ls ts" shows "lock_thread_ok ls (redT_updTs ts nts)" proof(rule lock_thread_okI) fix l t assume "has_lock (ls $ l) t" with lto have "\xw. ts t = \xw\" by(auto elim!: lock_thread_okE simp del: split_paired_Ex) thus "\xw. redT_updTs ts nts t = \xw\" by(auto intro: redT_updTs_Some1 simp del: split_paired_Ex) qed lemma lock_thread_ok_has_lock: assumes "lock_thread_ok ls ts" and "has_lock (ls $ l) t" obtains xw where "ts t = \xw\" using assms by(auto dest!: lock_thread_okD) lemma lock_thread_ok_None_has_locks_0: "\ lock_thread_ok ls ts; ts t = None \ \ has_locks (ls $ l) t = 0" by(rule ccontr)(auto dest: lock_thread_okD) lemma redT_upds_preserves_lock_thread_ok: "\lock_thread_ok ls ts; ts t = \xw\; thread_oks ts tas\ - \ lock_thread_ok (redT_updLs ls t las) (redT_updTs ts tas(t \ xw'))" + \ lock_thread_ok (redT_updLs ls t las) ((redT_updTs ts tas)(t \ xw'))" apply(rule lock_thread_okI) apply(clarsimp simp del: split_paired_Ex) apply(drule has_lock_upd_locks_implies_has_lock, simp) apply(drule lock_thread_okD, assumption) apply(erule exE) by(rule redT_updTs_Some1) lemma acquire_all_preserves_lock_thread_ok: fixes ln shows "\ lock_thread_ok ls ts; ts t = \(x, ln)\ \ \ lock_thread_ok (acquire_all ls t ln) (ts(t \ xw))" by(rule lock_thread_okI)(auto dest!: has_lock_acquire_locks_implies_has_lock dest: lock_thread_okD) text \Well-formedness condition: Wait sets contain only real threads\ definition wset_thread_ok :: "('w, 't) wait_sets \ ('l, 't, 'x) thread_info \ bool" where "wset_thread_ok ws ts \ \t. ts t = None \ ws t = None" lemma wset_thread_okI: "(\t. ts t = None \ ws t = None) \ wset_thread_ok ws ts" by(simp add: wset_thread_ok_def) lemma wset_thread_okD: "\ wset_thread_ok ws ts; ts t = None \ \ ws t = None" by(simp add: wset_thread_ok_def) lemma wset_thread_ok_conv_dom: "wset_thread_ok ws ts \ dom ws \ dom ts" by(auto simp add: wset_thread_ok_def) lemma wset_thread_ok_upd: "wset_thread_ok ls ts \ wset_thread_ok ls (ts(t \ xw))" by(auto intro!: wset_thread_okI dest: wset_thread_okD split: if_split_asm) lemma wset_thread_ok_upd_None: "wset_thread_ok ws ts \ wset_thread_ok (ws(t := None)) (ts(t := None))" by(auto intro!: wset_thread_okI dest: wset_thread_okD split: if_split_asm) lemma wset_thread_ok_upd_Some: "wset_thread_ok ws ts \ wset_thread_ok (ws(t := wo)) (ts(t \ xln))" by(auto intro!: wset_thread_okI dest: wset_thread_okD split: if_split_asm) lemma wset_thread_ok_upd_ws: "\ wset_thread_ok ws ts; ts t = \xln\ \ \ wset_thread_ok (ws(t := w)) ts" by(auto intro!: wset_thread_okI dest: wset_thread_okD) lemma wset_thread_ok_NotifyAllI: "wset_thread_ok ws ts \ wset_thread_ok (\t. if ws t = \w t\ then \w' t\ else ws t) ts" by(simp add: wset_thread_ok_def) lemma redT_updTs_preserves_wset_thread_ok: assumes wto: "wset_thread_ok ws ts" shows "wset_thread_ok ws (redT_updTs ts nts)" proof(rule wset_thread_okI) fix t assume "redT_updTs ts nts t = None" hence "ts t = None" by(rule redT_updTs_None) with wto show "ws t = None" by(rule wset_thread_okD) qed lemma redT_updW_preserve_wset_thread_ok: "\ wset_thread_ok ws ts; redT_updW t ws wa ws'; ts t = \xln\ \ \ wset_thread_ok ws' ts" by(fastforce simp add: redT_updW.simps intro: wset_thread_okI wset_thread_ok_NotifyAllI wset_thread_ok_upd_ws dest: wset_thread_okD) lemma redT_updWs_preserve_wset_thread_ok: "\ wset_thread_ok ws ts; redT_updWs t ws was ws'; ts t = \xln\ \ \ wset_thread_ok ws' ts" unfolding redT_updWs_def apply(rotate_tac 1) by(induct rule: rtrancl3p_converse_induct)(auto intro: redT_updW_preserve_wset_thread_ok) text \Well-formedness condition: Wait sets contain only non-final threads\ context final_thread begin definition wset_final_ok :: "('w, 't) wait_sets \ ('l, 't, 'x) thread_info \ bool" where "wset_final_ok ws ts \ (\t \ dom ws. \x ln. ts t = \(x, ln)\ \ \ final x)" lemma wset_final_okI: "(\t w. ws t = \w\ \ \x ln. ts t = \(x, ln)\ \ \ final x) \ wset_final_ok ws ts" unfolding wset_final_ok_def by(blast) lemma wset_final_okD: "\ wset_final_ok ws ts; ws t = \w\ \ \ \x ln. ts t = \(x, ln)\ \ \ final x" unfolding wset_final_ok_def by(blast) lemma wset_final_okE: assumes "wset_final_ok ws ts" "ws t = \w\" and "\x ln. ts t = \(x, ln)\ \ \ final x \ thesis" shows thesis using assms by(blast dest: wset_final_okD) lemma wset_final_ok_imp_wset_thread_ok: "wset_final_ok ws ts \ wset_thread_ok ws ts" apply(rule wset_thread_okI) apply(rule ccontr) apply(auto elim: wset_final_okE) done end end diff --git a/thys/JinjaThreads/J/Threaded.thy b/thys/JinjaThreads/J/Threaded.thy --- a/thys/JinjaThreads/J/Threaded.thy +++ b/thys/JinjaThreads/J/Threaded.thy @@ -1,765 +1,765 @@ (* Title: JinjaThreads/J/Threaded.thy Author: Andreas Lochbihler *) section \The source language as an instance of the framework\ theory Threaded imports SmallStep JWellForm "../Common/ConformThreaded" "../Common/ExternalCallWF" "../Framework/FWLiftingSem" "../Framework/FWProgressAux" begin context heap_base begin \ \Move to ?? - also used in BV\ lemma wset_Suspend_ok_start_state: fixes final r convert_RA assumes "start_state f P C M vs \ I" shows "start_state f P C M vs \ multithreaded_base.wset_Suspend_ok final r convert_RA I" using assms by(rule multithreaded_base.wset_Suspend_okI)(simp add: start_state_def split_beta) end abbreviation final_expr :: "'addr expr \ 'addr locals \ bool"where "final_expr \ \(e, x). final e" lemma final_locks: "final e \ expr_locks e l = 0" by(auto elim: finalE) context J_heap_base begin abbreviation mred :: "'addr J_prog \ ('addr, 'thread_id, 'addr expr \ 'addr locals, 'heap, 'addr, ('addr, 'thread_id) obs_event) semantics" where "mred P t \ (\((e, l), h) ta ((e', l'), h'). P,t \ \e, (h, l)\ -ta\ \e', (h', l')\)" lemma red_new_thread_heap: "\ convert_extTA extNTA,P,t \ \e, s\ -ta\ \e', s'\; NewThread t'' ex'' h'' \ set \ta\\<^bsub>t\<^esub> \ \ h'' = hp s'" and reds_new_thread_heap: "\ convert_extTA extNTA,P,t \ \es, s\ [-ta\] \es', s'\; NewThread t'' ex'' h'' \ set \ta\\<^bsub>t\<^esub> \ \ h'' = hp s'" apply(induct rule: red_reds.inducts) apply(fastforce dest: red_ext_new_thread_heap simp add: ta_upd_simps)+ done lemma red_ta_Wakeup_no_Join_no_Lock_no_Interrupt: "\ convert_extTA extNTA,P,t \ \e, s\ -ta\ \e', s'\; Notified \ set \ta\\<^bsub>w\<^esub> \ WokenUp \ set \ta\\<^bsub>w\<^esub> \ \ collect_locks \ta\\<^bsub>l\<^esub> = {} \ collect_cond_actions \ta\\<^bsub>c\<^esub> = {} \ collect_interrupts \ta\\<^bsub>i\<^esub> = {}" and reds_ta_Wakeup_no_Join_no_Lock_no_Interrupt: "\ convert_extTA extNTA,P,t \ \es, s\ [-ta\] \es', s'\; Notified \ set \ta\\<^bsub>w\<^esub> \ WokenUp \ set \ta\\<^bsub>w\<^esub> \ \ collect_locks \ta\\<^bsub>l\<^esub> = {} \ collect_cond_actions \ta\\<^bsub>c\<^esub> = {} \ collect_interrupts \ta\\<^bsub>i\<^esub> = {}" apply(induct rule: red_reds.inducts) apply(auto simp add: ta_upd_simps dest: red_external_Wakeup_no_Join_no_Lock_no_Interrupt del: conjI) done lemma final_no_red: "final e \ \ P,t \ \e, (h, l)\ -ta\ \e', (h', l')\" by(auto elim: red.cases finalE) lemma red_mthr: "multithreaded final_expr (mred P)" by(unfold_locales)(auto dest: red_new_thread_heap) end sublocale J_heap_base < red_mthr: multithreaded "final_expr" "mred P" convert_RA for P by(rule red_mthr) context J_heap_base begin abbreviation mredT :: "'addr J_prog \ ('addr,'thread_id,'addr expr \ 'addr locals,'heap,'addr) state \ ('thread_id \ ('addr, 'thread_id, 'addr expr \ 'addr locals,'heap) Jinja_thread_action) \ ('addr,'thread_id,'addr expr \ 'addr locals,'heap,'addr) state \ bool" where "mredT P \ red_mthr.redT P" abbreviation mredT_syntax1 :: "'addr J_prog \ ('addr,'thread_id,'addr expr \ 'addr locals,'heap,'addr) state \ 'thread_id \ ('addr, 'thread_id, 'addr expr \ 'addr locals,'heap) Jinja_thread_action \ ('addr,'thread_id,'addr expr \ 'addr locals,'heap,'addr) state \ bool" ("_ \ _ -_\_\ _" [50,0,0,0,50] 80) where "mredT_syntax1 P s t ta s' \ mredT P s (t, ta) s'" abbreviation mRedT_syntax1 :: "'addr J_prog \ ('addr,'thread_id,'addr expr \ 'addr locals,'heap,'addr) state \ ('thread_id \ ('addr, 'thread_id, 'addr expr \ 'addr locals,'heap) Jinja_thread_action) list \ ('addr,'thread_id,'addr expr \ 'addr locals,'heap,'addr) state \ bool" ("_ \ _ -\_\* _" [50,0,0,50] 80) where "P \ s -\ttas\* s' \ red_mthr.RedT P s ttas s'" end context J_heap begin lemma redT_hext_incr: "P \ s -t\ta\ s' \ shr s \ shr s'" by(erule red_mthr.redT.cases)(auto dest!: red_hext_incr intro: hext_trans) lemma RedT_hext_incr: assumes "P \ s -\tta\* s'" shows "shr s \ shr s'" using assms unfolding red_mthr.RedT_def by(induct)(auto dest: redT_hext_incr intro: hext_trans) end subsection \Lifting @{term "tconf"} to multithreaded states\ context J_heap begin lemma red_NewThread_Thread_Object: "\ convert_extTA extNTA,P,t \ \e, s\ -ta\ \e', s'\; NewThread t' x m \ set \ta\\<^bsub>t\<^esub> \ \ \C. typeof_addr (hp s') (thread_id2addr t') = \Class_type C\ \ P \ C \\<^sup>* Thread" and reds_NewThread_Thread_Object: "\ convert_extTA extNTA,P,t \ \es, s\ [-ta\] \es', s'\; NewThread t' x m \ set \ta\\<^bsub>t\<^esub> \ \ \C. typeof_addr (hp s') (thread_id2addr t') = \Class_type C\ \ P \ C \\<^sup>* Thread" apply(induct rule: red_reds.inducts) apply(fastforce dest: red_external_new_thread_exists_thread_object simp add: ta_upd_simps)+ done lemma lifting_wf_tconf: "lifting_wf final_expr (mred P) (\t ex h. P,h \ t \t)" by(unfold_locales)(fastforce dest: red_hext_incr red_NewThread_Thread_Object elim!: tconf_hext_mono intro: tconfI)+ end sublocale J_heap < red_tconf: lifting_wf final_expr "mred P" convert_RA "\t ex h. P,h \ t \t" by(rule lifting_wf_tconf) subsection \Towards agreement between the framework semantics' lock state and the locks stored in the expressions\ primrec sync_ok :: "('a,'b,'addr) exp \ bool" and sync_oks :: "('a,'b,'addr) exp list \ bool" where "sync_ok (new C) = True" | "sync_ok (newA T\i\) = sync_ok i" | "sync_ok (Cast T e) = sync_ok e" | "sync_ok (e instanceof T) = sync_ok e" | "sync_ok (Val v) = True" | "sync_ok (Var v) = True" | "sync_ok (e \bop\ e') = (sync_ok e \ sync_ok e' \ (contains_insync e' \ is_val e))" | "sync_ok (V := e) = sync_ok e" | "sync_ok (a\i\) = (sync_ok a \ sync_ok i \ (contains_insync i \ is_val a))" | "sync_ok (AAss a i e) = (sync_ok a \ sync_ok i \ sync_ok e \ (contains_insync i \ is_val a) \ (contains_insync e \ is_val a \ is_val i))" | "sync_ok (a\length) = sync_ok a" | "sync_ok (e\F{D}) = sync_ok e" | "sync_ok (FAss e F D e') = (sync_ok e \ sync_ok e' \ (contains_insync e' \ is_val e))" | "sync_ok (e\compareAndSwap(D\F, e', e'')) = (sync_ok e \ sync_ok e' \ sync_ok e'' \ (contains_insync e' \ is_val e) \ (contains_insync e'' \ is_val e \ is_val e'))" | "sync_ok (e\m(pns)) = (sync_ok e \ sync_oks pns \ (contains_insyncs pns \ is_val e))" | "sync_ok ({V : T=vo; e}) = sync_ok e" | "sync_ok (sync\<^bsub>V\<^esub> (o') e) = (sync_ok o' \ \ contains_insync e)" | "sync_ok (insync\<^bsub>V\<^esub> (a) e) = sync_ok e" | "sync_ok (e;;e') = (sync_ok e \ \ contains_insync e')" | "sync_ok (if (b) e else e') = (sync_ok b \ \ contains_insync e \ \ contains_insync e')" | "sync_ok (while (b) e) = (\ contains_insync b \ \ contains_insync e)" | "sync_ok (throw e) = sync_ok e" | "sync_ok (try e catch(C v) e') = (sync_ok e \ \ contains_insync e')" | "sync_oks [] = True" | "sync_oks (x # xs) = (sync_ok x \ sync_oks xs \ (contains_insyncs xs \ is_val x))" lemma sync_oks_append [simp]: "sync_oks (xs @ ys) \ sync_oks xs \ sync_oks ys \ (contains_insyncs ys \ (\vs. xs = map Val vs))" by(induct xs)(auto simp add: Cons_eq_map_conv) lemma fixes e :: "('a,'b,'addr) exp" and es :: "('a,'b,'addr) exp list" shows not_contains_insync_sync_ok: "\ contains_insync e \ sync_ok e" and not_contains_insyncs_sync_oks: "\ contains_insyncs es \ sync_oks es" by(induct e and es rule: sync_ok.induct sync_oks.induct)(auto) lemma expr_locks_sync_ok: "(\ad. expr_locks e ad = 0) \ sync_ok e" and expr_lockss_sync_oks: "(\ad. expr_lockss es ad = 0) \ sync_oks es" by(auto intro!: not_contains_insync_sync_ok not_contains_insyncs_sync_oks simp add: contains_insync_conv contains_insyncs_conv) lemma sync_ok_extRet2J [simp, intro!]: "sync_ok e \ sync_ok (extRet2J e va)" by(cases va) auto abbreviation sync_es_ok :: "('addr,'thread_id,('a,'b,'addr) exp\'c) thread_info \ 'heap \ bool" where "sync_es_ok \ ts_ok (\t (e, x) m. sync_ok e)" lemma sync_es_ok_blocks [simp]: "\ length pns = length Ts; length Ts = length vs \ \ sync_ok (blocks pns Ts vs e) = sync_ok e" by(induct pns Ts vs e rule: blocks.induct) auto context J_heap_base begin lemma assumes wf: "wf_J_prog P" shows red_preserve_sync_ok: "\ extTA,P,t \ \e, s\ -ta\ \e', s'\; sync_ok e \ \ sync_ok e'" and reds_preserve_sync_oks: "\ extTA,P,t \ \es, s\ [-ta\] \es', s'\; sync_oks es \ \ sync_oks es'" proof(induct rule: red_reds.inducts) case (RedCall s a U M Ts T pns body D vs) from wf \P \ class_type_of U sees M: Ts\T = \(pns, body)\ in D\ have "wf_mdecl wf_J_mdecl P D (M,Ts,T,\(pns,body)\)" by(rule sees_wf_mdecl) then obtain T where "P,[this\Class D,pns[\]Ts] \ body :: T" by(auto simp add: wf_mdecl_def) hence "expr_locks body = (\ad. 0)" by(rule WT_expr_locks) with \length vs = length pns\ \length Ts = length pns\ have "expr_locks (blocks pns Ts vs body) = (\ad. 0)" by(simp add: expr_locks_blocks) thus ?case by(auto intro: expr_locks_sync_ok) qed(fastforce intro: not_contains_insync_sync_ok)+ lemma assumes wf: "wf_J_prog P" shows expr_locks_new_thread: "\ P,t \ \e, s\ -ta\ \e', s'\; NewThread t'' (e'', x'') h \ set \ta\\<^bsub>t\<^esub> \ \ expr_locks e'' = (\ad. 0)" and expr_locks_new_thread': "\ P,t \ \es, s\ [-ta\] \es', s'\; NewThread t'' (e'', x'') h \ set \ta\\<^bsub>t\<^esub> \ \ expr_locks e'' = (\ad. 0)" proof(induct rule: red_reds.inducts) case (RedCallExternal s a U M Ts T D vs ta va h' ta' e' s') then obtain C fs a where subThread: "P \ C \\<^sup>* Thread" and ext: "extNTA2J P (C, run, a) = (e'', x'')" by(fastforce dest: red_external_new_thread_sub_thread) from sub_Thread_sees_run[OF wf subThread] obtain D pns body where sees: "P \ C sees run: []\Void = \(pns, body)\ in D" by auto from sees_wf_mdecl[OF wf this] obtain T where "P,[this \ Class D] \ body :: T" by(auto simp add: wf_mdecl_def) hence "expr_locks body = (\ad. 0)" by(rule WT_expr_locks) with sees ext show ?case by(auto simp add: extNTA2J_def) qed(auto simp add: ta_upd_simps) lemma assumes wf: "wf_J_prog P" shows red_new_thread_sync_ok: "\ P,t \ \e, s\ -ta\ \e', s'\; NewThread t'' (e'', x'') h'' \ set \ta\\<^bsub>t\<^esub> \ \ sync_ok e''" and reds_new_thread_sync_ok: "\ P,t \ \es, s\ [-ta\] \es', s'\; NewThread t'' (e'', x'') h'' \ set \ta\\<^bsub>t\<^esub> \ \ sync_ok e''" by(auto dest!: expr_locks_new_thread[OF wf] expr_locks_new_thread'[OF wf] intro: expr_locks_sync_ok expr_lockss_sync_oks) lemma lifting_wf_sync_ok: "wf_J_prog P \ lifting_wf final_expr (mred P) (\t (e, x) m. sync_ok e)" by(unfold_locales)(auto intro: red_preserve_sync_ok red_new_thread_sync_ok) lemma redT_preserve_sync_ok: assumes red: "P \ s -t\ta\ s'" shows "\ wf_J_prog P; sync_es_ok (thr s) (shr s) \ \ sync_es_ok (thr s') (shr s')" by(rule lifting_wf.redT_preserves[OF lifting_wf_sync_ok red]) lemma RedT_preserves_sync_ok: "\wf_J_prog P; P \ s -\ttas\* s'; sync_es_ok (thr s) (shr s)\ \ sync_es_ok (thr s') (shr s')" by(rule lifting_wf.RedT_preserves[OF lifting_wf_sync_ok]) lemma sync_es_ok_J_start_state: "\ wf_J_prog P; P \ C sees M:Ts\T=\(pns, body)\ in D; length Ts = length vs \ \ sync_es_ok (thr (J_start_state P C M vs)) m" apply(rule ts_okI) apply(clarsimp simp add: start_state_def split_beta split: if_split_asm) apply(drule (1) sees_wf_mdecl) apply(clarsimp simp add: wf_mdecl_def) apply(drule WT_expr_locks) apply(rule expr_locks_sync_ok) apply simp done end text \Framework lock state agrees with locks stored in the expression\ definition lock_ok :: "('addr,'thread_id) locks \ ('addr,'thread_id,('a, 'b,'addr) exp \ 'x) thread_info \ bool" where "\ln. lock_ok ls ts \ \t. (case (ts t) of None \ (\l. has_locks (ls $ l) t = 0) | \((e, x), ln)\ \ (\l. has_locks (ls $ l) t + ln $ l = expr_locks e l))" lemma lock_okI: "\ \t l. ts t = None \ has_locks (ls $ l) t = 0; \t e x ln l. ts t = \((e, x), ln)\ \ has_locks (ls $ l) t + ln $ l= expr_locks e l \ \ lock_ok ls ts" apply(fastforce simp add: lock_ok_def) done lemma lock_okE: "\ lock_ok ls ts; \t. ts t = None \ (\l. has_locks (ls $ l) t = 0) \ Q; \t e x ln. ts t = \((e, x), ln)\ \ (\l. has_locks (ls $ l) t + ln $ l = expr_locks e l) \ Q \ \ Q" by(fastforce simp add: lock_ok_def) lemma lock_okD1: "\ lock_ok ls ts; ts t = None \ \ \l. has_locks (ls $ l) t = 0" apply(simp add: lock_ok_def) apply(erule_tac x="t" in allE) apply(auto) done lemma lock_okD2: "\ln. \ lock_ok ls ts; ts t = \((e, x), ln)\ \ \ \l. has_locks (ls $ l) t + ln $ l = expr_locks e l" apply(fastforce simp add: lock_ok_def) done lemma lock_ok_lock_thread_ok: assumes lock: "lock_ok ls ts" shows "lock_thread_ok ls ts" proof(rule lock_thread_okI) fix l t assume lsl: "has_lock (ls $ l) t" show "\xw. ts t = \xw\" proof(cases "ts t") case None with lock have "has_locks (ls $ l) t = 0" by(auto dest: lock_okD1) with lsl show ?thesis by simp next case (Some a) thus ?thesis by blast qed qed lemma (in J_heap_base) lock_ok_J_start_state: "\ wf_J_prog P; P \ C sees M:Ts\T=\(pns, body)\ in D; length Ts = length vs \ \ lock_ok (locks (J_start_state P C M vs)) (thr (J_start_state P C M vs))" apply(rule lock_okI) apply(auto simp add: start_state_def split: if_split_asm) apply(drule (1) sees_wf_mdecl) apply(clarsimp simp add: wf_mdecl_def) apply(drule WT_expr_locks) apply(simp add: expr_locks_blocks) done subsection \Preservation of lock state agreement\ fun upd_expr_lock_action :: "int \ lock_action \ int" where "upd_expr_lock_action i Lock = i + 1" | "upd_expr_lock_action i Unlock = i - 1" | "upd_expr_lock_action i UnlockFail = i" | "upd_expr_lock_action i ReleaseAcquire = i" fun upd_expr_lock_actions :: "int \ lock_action list \ int" where "upd_expr_lock_actions n [] = n" | "upd_expr_lock_actions n (L # Ls) = upd_expr_lock_actions (upd_expr_lock_action n L) Ls" lemma upd_expr_lock_actions_append [simp]: "upd_expr_lock_actions n (Ls @ Ls') = upd_expr_lock_actions (upd_expr_lock_actions n Ls) Ls'" by(induct Ls arbitrary: n, auto) definition upd_expr_locks :: "('l \ int) \ 'l lock_actions \ 'l \ int" where "upd_expr_locks els las \ \l. upd_expr_lock_actions (els l) (las $ l)" lemma upd_expr_locks_iff [simp]: "upd_expr_locks els las l = upd_expr_lock_actions (els l) (las $ l)" by(simp add: upd_expr_locks_def) lemma upd_expr_lock_action_add [simp]: "upd_expr_lock_action (l + l') L = upd_expr_lock_action l L + l'" by(cases L, auto) lemma upd_expr_lock_actions_add [simp]: "upd_expr_lock_actions (l + l') Ls = upd_expr_lock_actions l Ls + l'" by(induct Ls arbitrary: l, auto) lemma upd_expr_locks_add [simp]: "upd_expr_locks (\a. x a + y a) las = (\a. upd_expr_locks x las a + y a)" by(auto intro: ext) lemma expr_locks_extRet2J [simp, intro!]: "expr_locks e = (\ad. 0) \ expr_locks (extRet2J e va) = (\ad. 0)" by(cases va) auto lemma (in J_heap_base) assumes wf: "wf_J_prog P" shows red_update_expr_locks: "\ convert_extTA extNTA,P,t \ \e, s\ -ta\ \e', s'\; sync_ok e \ \ upd_expr_locks (int o expr_locks e) \ta\\<^bsub>l\<^esub> = int o expr_locks e'" and reds_update_expr_lockss: "\ convert_extTA extNTA,P,t \ \es, s\ [-ta\] \es', s'\; sync_oks es \ \ upd_expr_locks (int o expr_lockss es) \ta\\<^bsub>l\<^esub> = int o expr_lockss es'" proof - have "\ convert_extTA extNTA,P,t \ \e, s\ -ta\ \e', s'\; sync_ok e \ \ upd_expr_locks (\ad. 0) \ta\\<^bsub>l\<^esub> = (\ad. (int o expr_locks e') ad - (int o expr_locks e) ad)" and "\ convert_extTA extNTA,P,t \ \es, s\ [-ta\] \es', s'\; sync_oks es \ \ upd_expr_locks (\ad. 0) \ta\\<^bsub>l\<^esub> = (\ad. (int o expr_lockss es') ad - (int o expr_lockss es) ad)" proof(induct rule: red_reds.inducts) case (RedCall s a U M Ts T pns body D vs) from wf \P \ class_type_of U sees M: Ts\T = \(pns, body)\ in D\ have "wf_mdecl wf_J_mdecl P D (M,Ts,T,\(pns,body)\)" by(rule sees_wf_mdecl) then obtain T where "P,[this\Class D,pns[\]Ts] \ body :: T" by(auto simp add: wf_mdecl_def) hence "expr_locks body = (\ad. 0)" by(rule WT_expr_locks) with \length vs = length pns\ \length Ts = length pns\ have "expr_locks (blocks pns Ts vs body) = (\ad. 0)" by(simp add: expr_locks_blocks) thus ?case by(auto intro: expr_locks_sync_ok) next case RedCallExternal thus ?case by(auto simp add: fun_eq_iff contains_insync_conv contains_insyncs_conv finfun_upd_apply ta_upd_simps elim!: red_external.cases) qed(fastforce simp add: fun_eq_iff contains_insync_conv contains_insyncs_conv finfun_upd_apply ta_upd_simps)+ hence "\ convert_extTA extNTA,P,t \ \e, s\ -ta\ \e', s'\; sync_ok e \ \ upd_expr_locks (\ad. 0 + (int \ expr_locks e) ad) \ta\\<^bsub>l\<^esub> = int \ expr_locks e'" and "\ convert_extTA extNTA,P,t \ \es, s\ [-ta\] \es', s'\; sync_oks es \ \ upd_expr_locks (\ad. 0 + (int \ expr_lockss es) ad) \ta\\<^bsub>l\<^esub> = int \ expr_lockss es'" by(auto intro: ext simp only: upd_expr_locks_add) thus "\ convert_extTA extNTA,P,t \ \e, s\ -ta\ \e', s'\; sync_ok e \ \ upd_expr_locks (int o expr_locks e) \ta\\<^bsub>l\<^esub> = int o expr_locks e'" and "\ convert_extTA extNTA,P,t \ \es, s\ [-ta\] \es', s'\; sync_oks es \ \ upd_expr_locks (int o expr_lockss es) \ta\\<^bsub>l\<^esub> = int o expr_lockss es'" by(auto simp add: o_def) qed definition lock_expr_locks_ok :: "'t FWState.lock \ 't \ nat \ int \ bool" where "lock_expr_locks_ok l t n i \ (i = int (has_locks l t) + int n) \ i \ 0" lemma upd_lock_upd_expr_lock_action_preserve_lock_expr_locks_ok: assumes lao: "lock_action_ok l t L" and lelo: "lock_expr_locks_ok l t n i" shows "lock_expr_locks_ok (upd_lock l t L) t (upd_threadR n l t L) (upd_expr_lock_action i L)" proof - from lelo have i: "i \ 0" and hl: "i = int (has_locks l t) + int n" by(auto simp add: lock_expr_locks_ok_def) from lelo show ?thesis proof(cases L) case Lock with lao have "may_lock l t" by(simp) with hl have "has_locks (lock_lock l t) t = (Suc (has_locks l t))" by(auto) with Lock i hl show ?thesis by(simp add: lock_expr_locks_ok_def) next case Unlock with lao have "has_lock l t" by simp then obtain n' where hl': "has_locks l t = Suc n'" by(auto dest: has_lock_has_locks_Suc) hence "has_locks (unlock_lock l) t = n'" by simp with Unlock i hl hl' show ?thesis by(simp add: lock_expr_locks_ok_def) qed(auto simp add: lock_expr_locks_ok_def) qed lemma upd_locks_upd_expr_lock_preserve_lock_expr_locks_ok: "\ lock_actions_ok l t Ls; lock_expr_locks_ok l t n i \ \ lock_expr_locks_ok (upd_locks l t Ls) t (upd_threadRs n l t Ls) (upd_expr_lock_actions i Ls)" by(induct Ls arbitrary: l i n)(auto intro: upd_lock_upd_expr_lock_action_preserve_lock_expr_locks_ok) definition ls_els_ok :: "('addr,'thread_id) locks \ 'thread_id \ ('addr \f nat) \ ('addr \ int) \ bool" where "\ln. ls_els_ok ls t ln els \ \l. lock_expr_locks_ok (ls $ l) t (ln $ l) (els l)" lemma ls_els_okI: "\ln. (\l. lock_expr_locks_ok (ls $ l) t (ln $ l) (els l)) \ ls_els_ok ls t ln els" by(auto simp add: ls_els_ok_def) lemma ls_els_okE: "\ln. \ ls_els_ok ls t ln els; \l. lock_expr_locks_ok (ls $ l) t (ln $ l) (els l) \ P \ \ P" by(auto simp add: ls_els_ok_def) lemma ls_els_okD: "\ln. ls_els_ok ls t ln els \ lock_expr_locks_ok (ls $ l) t (ln $ l) (els l)" by(auto simp add: ls_els_ok_def) lemma redT_updLs_upd_expr_locks_preserves_ls_els_ok: "\ln. \ ls_els_ok ls t ln els; lock_ok_las ls t las \ \ ls_els_ok (redT_updLs ls t las) t (redT_updLns ls t ln las) (upd_expr_locks els las)" by(auto intro!: ls_els_okI upd_locks_upd_expr_lock_preserve_lock_expr_locks_ok elim!: ls_els_okE simp add: redT_updLs_def lock_ok_las_def) lemma sync_ok_redT_updT: assumes "sync_es_ok ts h" and nt: "\t e x h''. ta = NewThread t (e, x) h'' \ sync_ok e" shows "sync_es_ok (redT_updT ts ta) h'" using assms proof(cases ta) case (NewThread T x m) obtain E X where [simp]: "x = (E, X)" by (cases x, auto) with NewThread have "sync_ok E" by(simp)(rule nt) with NewThread \sync_es_ok ts h\ show ?thesis apply - apply(rule ts_okI) apply(case_tac "t=T") by(auto dest: ts_okD) qed(auto intro: ts_okI dest: ts_okD) lemma sync_ok_redT_updTs: "\ sync_es_ok ts h; \t e x h. NewThread t (e, x) h \ set tas \ sync_ok e \ \ sync_es_ok (redT_updTs ts tas) h'" proof(induct tas arbitrary: ts) case Nil thus ?case by(auto intro: ts_okI dest: ts_okD) next case (Cons TA TAS TS) note IH = \\ts. \sync_es_ok ts h; \t e x h''. NewThread t (e, x) h'' \ set TAS \ sync_ok e\ \ sync_es_ok (redT_updTs ts TAS) h'\ note nt = \\t e x h. NewThread t (e, x) h \ set (TA # TAS) \ sync_ok e\ from \sync_es_ok TS h\ nt have "sync_es_ok (redT_updT TS TA) h" by(auto elim!: sync_ok_redT_updT) hence "sync_es_ok (redT_updTs (redT_updT TS TA) TAS) h'" by(rule IH)(auto intro: nt) thus ?case by simp qed lemma lock_ok_thr_updI: "\ln. \ lock_ok ls ts; ts t = \((e, xs), ln)\; expr_locks e = expr_locks e' \ \ lock_ok ls (ts(t \ ((e', xs'), ln)))" by(rule lock_okI)(auto split: if_split_asm dest: lock_okD2 lock_okD1) context J_heap_base begin lemma redT_preserves_lock_ok: assumes wf: "wf_J_prog P" and "P \ s -t\ta\ s'" and "lock_ok (locks s) (thr s)" and "sync_es_ok (thr s) (shr s)" shows "lock_ok (locks s') (thr s')" proof - obtain ls ts h ws "is" where s [simp]: "s = (ls, (ts, h), ws, is)" by(cases s) fastforce obtain ls' ts' h' ws' is' where s' [simp]: "s' = (ls', (ts', h'), ws', is')" by(cases s') fastforce from assms have redT: "P \ (ls, (ts, h), ws, is) -t\ta\ (ls', (ts', h'), ws', is')" and loes: "lock_ok ls ts" and aoes: "sync_es_ok ts h" by auto from redT have "lock_ok ls' ts'" proof(cases rule: red_mthr.redT_elims) case (normal a a' m') moreover obtain e x where "a = (e, x)" by (cases a, auto) moreover obtain e' x' where "a' = (e', x')" by (cases a', auto) ultimately have P: "P,t \ \e,(h, x)\ -ta\ \e',(m', x')\" and est: "ts t = \((e, x), no_wait_locks)\" and lota: "lock_ok_las ls t \ta\\<^bsub>l\<^esub>" and cctta: "thread_oks ts \ta\\<^bsub>t\<^esub>" and ls': "ls' = redT_updLs ls t \ta\\<^bsub>l\<^esub>" - and s': "ts' = redT_updTs ts \ta\\<^bsub>t\<^esub>(t \ ((e', x'), redT_updLns ls t no_wait_locks \ta\\<^bsub>l\<^esub>))" + and s': "ts' = (redT_updTs ts \ta\\<^bsub>t\<^esub>)(t \ ((e', x'), redT_updLns ls t no_wait_locks \ta\\<^bsub>l\<^esub>))" by auto - let ?ts' = "redT_updTs ts \ta\\<^bsub>t\<^esub>(t \ ((e', x'), redT_updLns ls t no_wait_locks \ta\\<^bsub>l\<^esub>))" + let ?ts' = "(redT_updTs ts \ta\\<^bsub>t\<^esub>)(t \ ((e', x'), redT_updLns ls t no_wait_locks \ta\\<^bsub>l\<^esub>))" from est aoes have aoe: "sync_ok e" by(auto dest: ts_okD) from aoe P have aoe': "sync_ok e'" by(auto dest: red_preserve_sync_ok[OF wf]) from aoes red_new_thread_sync_ok[OF wf P] have "sync_es_ok (redT_updTs ts \ta\\<^bsub>t\<^esub>) h'" by(rule sync_ok_redT_updTs) with aoe' have aoes': "sync_es_ok ?ts' m'" by(auto intro!: ts_okI dest: ts_okD split: if_split_asm) have "lock_ok ls' ?ts'" proof(rule lock_okI) fix t'' l assume "?ts' t'' = None" hence "ts t'' = None" by(auto split: if_split_asm intro: redT_updTs_None) with loes have "has_locks (ls $ l) t'' = 0" by(auto dest: lock_okD1) moreover from \?ts' t'' = None\ have "t \ t''" by(simp split: if_split_asm) ultimately show "has_locks (ls' $ l) t'' = 0" by(simp add: red_mthr.redT_has_locks_inv[OF redT]) next fix t'' e'' x'' l ln'' assume ts't'': "?ts' t'' = \((e'', x''), ln'')\" with aoes' have aoe'': "sync_ok e''" by(auto dest: ts_okD) show "has_locks (ls' $ l) t'' + ln'' $ l = expr_locks e'' l" proof(cases "t = t''") case True note tt'' = \t = t''\ with ts't'' have e'': "e'' = e'" and x'': "x'' = x'" and ln'': "ln'' = redT_updLns ls t no_wait_locks \ta\\<^bsub>l\<^esub>" by auto have "ls_els_ok ls t no_wait_locks (int o expr_locks e)" proof(rule ls_els_okI) fix l note lock_okD2[OF loes, OF est] thus "lock_expr_locks_ok (ls $ l) t (no_wait_locks $ l) ((int \ expr_locks e) l)" by(simp add: lock_expr_locks_ok_def) qed hence "ls_els_ok (redT_updLs ls t \ta\\<^bsub>l\<^esub>) t (redT_updLns ls t no_wait_locks \ta\\<^bsub>l\<^esub>) (upd_expr_locks (int o expr_locks e) \ta\\<^bsub>l\<^esub>)" by(rule redT_updLs_upd_expr_locks_preserves_ls_els_ok[OF _ lota]) hence "ls_els_ok (redT_updLs ls t \ta\\<^bsub>l\<^esub>) t (redT_updLns ls t no_wait_locks \ta\\<^bsub>l\<^esub>) (int o expr_locks e')" by(simp only: red_update_expr_locks[OF wf P aoe]) thus ?thesis using ls' e'' tt'' ln'' by(auto dest: ls_els_okD[where l = l] simp: lock_expr_locks_ok_def) next case False note tt'' = \t \ t''\ from lota have lao: "lock_actions_ok (ls $ l) t (\ta\\<^bsub>l\<^esub> $ l)" by(simp add: lock_ok_las_def) show ?thesis proof(cases "ts t''") case None with est ts't'' tt'' cctta obtain m where "NewThread t'' (e'', x'') m \ set \ta\\<^bsub>t\<^esub>" and ln'': "ln'' = no_wait_locks" by(auto dest: redT_updTs_new_thread) moreover with P have "m' = m" by(auto dest: red_new_thread_heap) ultimately have "NewThread t'' (e'', x'') m' \ set \ta\\<^bsub>t\<^esub>" by simp with wf P ln'' have "expr_locks e'' = (\ad. 0)" by -(rule expr_locks_new_thread) hence elel: "expr_locks e'' l = 0" by simp from loes None have "has_locks (ls $ l) t'' = 0" by(auto dest: lock_okD1) moreover note lock_actions_ok_has_locks_upd_locks_eq_has_locks[OF lao tt''[symmetric]] ultimately have "has_locks (redT_updLs ls t \ta\\<^bsub>l\<^esub> $ l) t'' = 0" by(auto simp add: fun_eq_iff) with elel ls' ln'' show ?thesis by(auto) next case (Some a) then obtain E X LN where est'': "ts t'' = \((E, X), LN)\" by(cases a, auto) with loes have IH: "has_locks (ls $ l) t'' + LN $ l = expr_locks E l" by(auto dest: lock_okD2) from est est'' tt'' cctta have "?ts' t'' = \((E, X), LN)\" by(simp)(rule redT_updTs_Some, simp_all) with ts't'' have e'': "E = e''" and x'': "X = x''" and ln'': "ln'' = LN" by(simp_all) with lock_actions_ok_has_locks_upd_locks_eq_has_locks[OF lao tt''[symmetric]] IH ls' show ?thesis by(clarsimp simp add: redT_updLs_def fun_eq_iff) qed qed qed with s' show ?thesis by simp next case (acquire a ln n) hence [simp]: "ta = (K$ [], [], [], [], [], convert_RA ln)" "ws' = ws" "h' = h" and ls': "ls' = acquire_all ls t ln" and ts': "ts' = ts(t \ (a, no_wait_locks))" and "ts t = \(a, ln)\" and "may_acquire_all ls t ln" by auto obtain e x where [simp]: "a = (e, x)" by (cases a, auto) from ts' have ts': "ts' = ts(t \ ((e, x), no_wait_locks))" by simp from \ts t = \(a, ln)\\ have tst: "ts t = \((e, x), ln)\" by simp show ?thesis proof(rule lock_okI) fix t'' l assume rtutes: "ts' t'' = None" with ts' have tst'': "ts t'' = None" by(simp split: if_split_asm) with tst have tt'': "t \ t''" by auto from tst'' loes have "has_locks (ls $ l) t'' = 0" by(auto dest: lock_okD1) thus "has_locks (ls' $ l) t'' = 0" by(simp add: red_mthr.redT_has_locks_inv[OF redT tt'']) next fix t'' e'' x'' ln'' l assume ts't'': "ts' t'' = \((e'', x''), ln'')\" show "has_locks (ls' $ l) t'' + ln'' $ l = expr_locks e'' l" proof(cases "t = t''") case True note [simp] = this with ts't'' ts' tst have [simp]: "ln'' = no_wait_locks" "e = e''" by auto from tst loes have "has_locks (ls $ l) t + ln $ l = expr_locks e l" by(auto dest: lock_okD2) show ?thesis proof(cases "ln $ l > 0") case True with \may_acquire_all ls t ln\ ls' have "may_lock (ls $ l) t" by(auto elim: may_acquire_allE) with ls' have "has_locks (ls' $ l) t = has_locks (ls $ l) t + ln $ l" by(simp add: has_locks_acquire_locks_conv) with \has_locks (ls $ l) t + ln $ l = expr_locks e l\ show ?thesis by(simp) next case False hence "ln $ l = 0" by simp with ls' have "has_locks (ls' $ l) t = has_locks (ls $ l) t" by(simp) with \has_locks (ls $ l) t + ln $ l = expr_locks e l\ \ln $ l = 0\ show ?thesis by(simp) qed next case False with ts' ts't'' have tst'': "ts t'' = \((e'', x''), ln'')\" by(simp) with loes have "has_locks (ls $ l) t'' + ln'' $ l = expr_locks e'' l" by(auto dest: lock_okD2) show ?thesis proof(cases "ln $ l > 0") case False with \t \ t''\ ls' have "has_locks (ls' $ l) t'' = has_locks (ls $ l) t''" by(simp) with \has_locks (ls $ l) t'' + ln'' $ l = expr_locks e'' l\ show ?thesis by(simp) next case True with \may_acquire_all ls t ln\ have "may_lock (ls $ l) t" by(auto elim: may_acquire_allE) with ls' \t \ t''\ have "has_locks (ls' $ l) t'' = has_locks (ls $ l) t''" by(simp add: has_locks_acquire_locks_conv') with ls' \has_locks (ls $ l) t'' + ln'' $ l = expr_locks e'' l\ show ?thesis by(simp) qed qed qed qed thus ?thesis by simp qed lemma invariant3p_sync_es_ok_lock_ok: assumes wf: "wf_J_prog P" shows "invariant3p (mredT P) {s. sync_es_ok (thr s) (shr s) \ lock_ok (locks s) (thr s)}" apply(rule invariant3pI) apply clarify apply(rule conjI) apply(rule lifting_wf.redT_preserves[OF lifting_wf_sync_ok[OF wf]], blast) apply(assumption) apply(erule (2) redT_preserves_lock_ok[OF wf]) done lemma RedT_preserves_lock_ok: assumes wf: "wf_J_prog P" and Red: "P \ s -\ttas\* s'" and ae: "sync_es_ok (thr s) (shr s)" and loes: "lock_ok (locks s) (thr s)" shows "lock_ok (locks s') (thr s')" using invariant3p_rtrancl3p[OF invariant3p_sync_es_ok_lock_ok[OF wf] Red[unfolded red_mthr.RedT_def]] ae loes by simp end subsection \Determinism\ context J_heap_base begin lemma fixes final assumes det: "deterministic_heap_ops" shows red_deterministic: "\ convert_extTA extTA,P,t \ \e, (shr s, xs)\ -ta\ \e', s'\; convert_extTA extTA,P,t \ \e, (shr s, xs)\ -ta'\ \e'', s''\; final_thread.actions_ok final s t ta; final_thread.actions_ok final s t ta' \ \ ta = ta' \ e' = e'' \ s' = s''" and reds_deterministic: "\ convert_extTA extTA,P,t \ \es, (shr s, xs)\ [-ta\] \es', s'\; convert_extTA extTA,P,t \ \es, (shr s, xs)\ [-ta'\] \es'', s''\; final_thread.actions_ok final s t ta; final_thread.actions_ok final s t ta' \ \ ta = ta' \ es' = es'' \ s' = s''" proof(induct e "(shr s, xs)" ta e' s' and es "(shr s, xs)" ta es' s' arbitrary: e'' s'' xs and es'' s'' xs rule: red_reds.inducts) case RedNew thus ?case by(auto elim!: red_cases dest: deterministic_heap_ops_allocateD[OF det]) next case RedNewArray thus ?case by(auto elim!: red_cases dest: deterministic_heap_ops_allocateD[OF det]) next case RedCall thus ?case by(auto elim!: red_cases dest: sees_method_fun simp add: map_eq_append_conv) next case RedCallExternal thus ?case by(auto elim!: red_cases dest: red_external_deterministic[OF det] simp add: final_thread.actions_ok_iff map_eq_append_conv dest: sees_method_fun) next case RedCallNull thus ?case by(auto elim!: red_cases dest: sees_method_fun simp add: map_eq_append_conv) next case CallThrowParams thus ?case by(auto elim!: red_cases dest: sees_method_fun simp add: map_eq_append_conv append_eq_map_conv append_eq_append_conv2 reds_map_Val_Throw Cons_eq_append_conv append_eq_Cons_conv) qed(fastforce elim!: red_cases reds_cases dest: deterministic_heap_ops_readD[OF det] deterministic_heap_ops_writeD[OF det] iff: reds_map_Val_Throw)+ lemma red_mthr_deterministic: assumes det: "deterministic_heap_ops" shows "red_mthr.deterministic P UNIV" proof(rule red_mthr.determisticI) fix s t x ta' x' m' ta'' x'' m'' assume "thr s t = \(x, no_wait_locks)\" and red: "mred P t (x, shr s) ta' (x', m')" "mred P t (x, shr s) ta'' (x'', m'')" and aok: "red_mthr.actions_ok s t ta'" "red_mthr.actions_ok s t ta''" moreover obtain e xs where [simp]: "x = (e, xs)" by(cases x) moreover obtain e' xs' where [simp]: "x' = (e', xs')" by(cases x') moreover obtain e'' xs'' where [simp]: "x'' = (e'', xs'')" by(cases x'') ultimately have "extTA2J P,P,t \ \e,(shr s, xs)\ -ta'\ \e',(m', xs')\" and "extTA2J P,P,t \ \e,(shr s, xs)\ -ta''\ \e'',(m'', xs'')\" by simp_all from red_deterministic[OF det this aok] show "ta' = ta'' \ x' = x'' \ m' = m''" by simp qed simp end end