(* * Copyright 2014, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory NonDetMonadLemmaBucket imports "Monad_WP/NonDetMonadVCG" "MonadEq" "Monad_WP/WhileLoopRulesCompleteness" Distinct_Prop "~~/src/HOL/Word/Word_Miscellaneous" begin setup \AutoLevity_Base.add_attribute_test "wp" WeakestPre.is_wp_rule\ lemma no_fail_assume_pre: "(\s. P s \ no_fail P f) \ no_fail P f" by (simp add: no_fail_def) lemma no_fail_liftM_eq [simp]: "no_fail P (liftM f m) = no_fail P m" by (auto simp: liftM_def no_fail_def bind_def return_def) lemma mapME_Cons: "mapME m (x # xs) = (doE y \ m x; ys \ (mapME m xs); returnOk (y # ys) odE)" by (simp add: mapME_def sequenceE_def Let_def) lemma mapME_Nil : "mapME f [] = returnOk []" unfolding mapME_def by (simp add: sequenceE_def) lemma hoare_take_disjunct: "\P\ f \\rv s. P' rv s \ (False \ P'' rv s)\ \ \P\ f \P''\" by (erule hoare_strengthen_post, simp) lemma hoare_post_add: "\P\ S \\r s. R r s \ Q r s\ \ \P\ S \Q\" by (erule hoare_strengthen_post, simp) lemma hoare_disjI1: "\R\ f \P\ \ \R\ f \\r s. P r s \ Q r s\" apply (erule hoare_post_imp [rotated]) apply simp done lemma hoare_disjI2: "\R\ f \Q\ \ \R\ f \\r s. P r s \ Q r s \" by (rule hoare_post_imp [OF _ hoare_disjI1, where P1=Q], auto) lemma hoare_name_pre_state: "\ \s. P s \ \op = s\ f \Q\ \ \ \P\ f \Q\" by (clarsimp simp: valid_def) lemma hoare_name_pre_stateE: "\\s. P s \ \op = s\ f \Q\, \E\\ \ \P\ f \Q\, \E\" by (clarsimp simp: validE_def2) lemma valid_prove_more: "\P\ f \\rv s. Q rv s \ Q' rv s\ \ \P\ f \Q'\" by (erule hoare_strengthen_post, simp) lemma hoare_vcg_if_lift: "\R\ f \\rv s. (P \ X rv s) \ (\P \ Y rv s)\ \ \R\ f \\rv s. if P then X rv s else Y rv s\" "\R\ f \\rv s. (P \ X rv s) \ (\P \ Y rv s)\ \ \R\ f \\rv. if P then X rv else Y rv\" by (auto simp: valid_def split_def) lemma no_fail_bind [wp]: assumes f: "no_fail P f" assumes g: "\rv. no_fail (R rv) (g rv)" assumes v: "\Q\ f \R\" shows "no_fail (P and Q) (f >>= (\rv. g rv))" apply (clarsimp simp: no_fail_def bind_def) apply (rule conjI) prefer 2 apply (erule no_failD [OF f]) apply clarsimp apply (drule (1) use_valid [OF _ v]) apply (drule no_failD [OF g]) apply simp done lemma hoare_lift_Pf2: assumes P: "\x. \Q x\ m \\_. P x\" assumes f: "\P. \\s. P (f s)\ m \\_ s. P (f s)\" shows "\\s. Q (f s) s\ m \\_ s. P (f s) s\" apply (clarsimp simp add: valid_def) apply (frule (1) use_valid [OF _ P], drule (2) use_valid [OF _ f]) done lemma hoare_lift_Pf3: assumes P: "\x. \Q x\ m \P x\" assumes f: "\P. \\s. P (f s)\ m \\_ s. P (f s)\" shows "\\s. Q (f s) s\ m \\rv s. P (f s) rv s\" apply (clarsimp simp add: valid_def) apply (frule (1) use_valid [OF _ P], drule (2) use_valid [OF _ f]) done lemma no_fail_select_f [wp]: "no_fail (\s. \snd S) (select_f S)" by (simp add: select_f_def no_fail_def) lemma hoare_lift_Pf: assumes P: "\x. \P x\ m \\_. P x\" assumes f: "\P. \\s. P (f s)\ m \\_ s. P (f s)\" shows "\\s. P (f s) s\ m \\_ s. P (f s) s\" apply (clarsimp simp add: valid_def) apply (frule (1) use_valid [OF _ P], drule (2) use_valid [OF _ f]) done lemma assert_def2: "assert v = assert_opt (if v then Some () else None)" by (cases v, simp_all add: assert_def assert_opt_def) lemma hoare_if_r_and: "\P\ f \\r. if R r then Q r else Q' r\ = \P\ f \\r s. (R r \ Q r s) \ (\R r \ Q' r s)\" by (fastforce simp: valid_def) lemma no_fail_liftM [wp]: "no_fail P m \ no_fail P (liftM f m)" by (simp) lemma no_fail_pre_and: "no_fail P f \ no_fail (P and Q) f" by (erule no_fail_pre) simp lemma hoare_convert_imp: "\ \\s. \ P s\ f \\rv s. \ Q s\; \R\ f \S\ \ \ \\s. P s \ R s\ f \\rv s. Q s \ S rv s\" apply (simp only: imp_conv_disj) apply (erule(1) hoare_vcg_disj_lift) done lemma hoare_vcg_ex_lift_R: "\ \v. \P v\ f \Q v\,- \ \ \\s. \v. P v s\ f \\rv s. \v. Q v rv s\,-" apply (simp add: validE_R_def validE_def) apply (rule hoare_strengthen_post, erule hoare_vcg_ex_lift) apply (auto split: sum.split) done lemma hoare_case_option_wpR: "\\P\ f None \Q\,-; \x. \P' x\ f (Some x) \Q' x\,-\ \ \case_option P P' v\ f v \\rv. case v of None \ Q rv | Some x \ Q' x rv\,-" by (cases v) auto lemma hoare_vcg_conj_liftE_R: "\ \P\ f \P'\,-; \Q\ f \Q'\,- \ \ \P and Q\ f \\rv s. P' rv s \ Q' rv s\, -" apply (simp add: validE_R_def validE_def valid_def split: sum.splits) apply blast done lemma zipWithM_x_inv: assumes x: "\x y. \P\ m x y \\rv. P\" shows "length xs = length ys \ \P\ zipWithM_x m xs ys \\rv. P\" proof (induct xs ys rule: list_induct2) case Nil show ?case by (simp add: zipWithM_x_def sequence_x_def zipWith_def) next case (Cons a as b bs) have zipWithM_x_Cons: "\m x xs y ys. zipWithM_x m (x # xs) (y # ys) = do m x y; zipWithM_x m xs ys od" by (simp add: zipWithM_x_def sequence_x_def zipWith_def) have IH: "\P\ zipWithM_x m as bs \\rv. P\" by fact show ?case by (simp add: zipWithM_x_Cons) (wp IH x) qed lemma K_valid[wp]: "\K P\ f \\_. K P\" by (simp add: valid_def) lemma mapME_wp: assumes x: "\x. x \ S \ \P\ f x \\_. P\, \\_. E\" shows "set xs \ S \ \P\ mapME f xs \\_. P\, \\_. E\" apply (induct xs) apply (simp add: mapME_def sequenceE_def) apply wp apply simp apply (simp add: mapME_Cons) apply (wp x|simp)+ done lemmas mapME_wp' = mapME_wp [OF _ subset_refl] lemma sequence_x_Cons: "\x xs. sequence_x (x # xs) = (x >>= (\_. sequence_x xs))" by (simp add: sequence_x_def) lemma mapM_Cons: "mapM m (x # xs) = (do y \ m x; ys \ (mapM m xs); return (y # ys) od)" by (simp add: mapM_def sequence_def Let_def) lemma mapM_simps: "mapM m [] = return []" "mapM m (x#xs) = do r \ m x; rs \ (mapM m xs); return (r#rs) od" by (simp_all add: mapM_def sequence_def) lemma zipWithM_x_mapM: "zipWithM_x f as bs = (mapM (split f) (zip as bs) >>= (\_. return ()))" apply (simp add: zipWithM_x_def zipWith_def) apply (induct ("zip as bs")) apply (simp add: sequence_x_def mapM_def sequence_def) apply (simp add: sequence_x_Cons mapM_Cons bind_assoc) done (* zipWithM_x and mapM_ *) lemma mapM_wp: assumes x: "\x. x \ S \ \P\ f x \\rv. P\" shows "set xs \ S \ \P\ mapM f xs \\rv. P\" apply (induct xs) apply (simp add: mapM_def sequence_def) apply (simp add: mapM_Cons) apply wp apply (rule x, clarsimp) apply simp done lemma mapM_x_mapM: "mapM_x m l = (mapM m l >>= (\x. return ()))" apply (simp add: mapM_x_def sequence_x_def mapM_def sequence_def) apply (induct l, simp_all add: Let_def bind_assoc) done lemma mapM_x_wp: assumes x: "\x. x \ S \ \P\ f x \\rv. P\" shows "set xs \ S \ \P\ mapM_x f xs \\rv. P\" by (subst mapM_x_mapM) (wp mapM_wp x) lemma mapM_x_Nil: "mapM_x f [] = return ()" unfolding mapM_x_def sequence_x_def by simp lemma sequence_xappend1: "sequence_x (xs @ [x]) = (sequence_x xs >>= (\_. x))" by (induct xs) (simp add: sequence_x_def, simp add: sequence_x_Cons bind_assoc) lemma mapM_append_single: "mapM_x f (xs @ [y]) = (mapM_x f xs >>= (\_. f y))" unfolding mapM_x_def by (simp add: sequence_xappend1) lemma mapM_x_Cons: "mapM_x m (x # xs) = (do m x; mapM_x m xs od)" by (simp add: mapM_x_def sequence_x_def) lemma mapM_x_inv_wp2: assumes post: "\s. \ I s; V [] s \ \ Q s" and hr: "\a as. suffix (a # as) xs \ \\s. I s \ V (a # as) s\ m a \\r s. I s \ V as s\" shows "\I and V xs\ mapM_x m xs \\rv. Q\" proof (induct xs rule: list_induct_suffix) case Nil thus ?case apply (simp add: mapM_x_Nil) apply wp apply (clarsimp intro!: post) done next case (Cons x xs) thus ?case apply (simp add: mapM_x_Cons) apply wp apply (wp hr) apply assumption done qed lemma zipWithM_x_mapM_x: "zipWithM_x f as bs = mapM_x (\(x, y). f x y) (zip as bs)" apply (subst zipWithM_x_mapM) apply (subst mapM_x_mapM) apply (rule refl) done lemma zipWithM_x_append1: fixes f :: "'b \ 'c \ ('a, unit) nondet_monad" assumes ls: "length xs = length ys" shows "(zipWithM_x f (xs @ [x]) (ys @ [y])) = (zipWithM_x f xs ys >>= (\_. f x y))" unfolding zipWithM_x_def zipWith_def by (subst zip_append [OF ls], simp, rule sequence_xappend1) lemma zipWithM_x_Cons: assumes ls: "length xs = length ys" shows "(zipWithM_x f (x # xs) (y # ys)) = (f x y >>= (\_. zipWithM_x f xs ys))" unfolding zipWithM_x_def zipWith_def by (simp, rule sequence_x_Cons) lemma mapM_x_inv_wp3: fixes m :: "'b \ ('a, unit) nondet_monad" assumes hr: "\a as bs. xs = as @ [a] @ bs \ \\s. I s \ V as s\ m a \\r s. I s \ V (as @ [a]) s\" shows "\\s. I s \ V [] s\ mapM_x m xs \\rv s. I s \ V xs s\" using hr proof (induct xs rule: rev_induct) case Nil thus ?case apply (simp add: mapM_x_Nil) done next case (snoc x xs) show ?case apply (simp add: mapM_append_single) apply (wp snoc.prems) apply simp apply (rule snoc.hyps [OF snoc.prems]) apply simp apply assumption done qed lemma mapME_x_map_simp: "mapME_x m (map f xs) = mapME_x (m o f) xs" by (simp add: mapME_x_def sequenceE_x_def) lemma mapM_return: "mapM (\x. return (f x)) xs = return (map f xs)" apply (induct xs) apply (simp add: mapM_def sequence_def) apply (simp add: mapM_Cons) done lemma mapME_x_inv_wp: assumes x: "\x. \P\ f x \\rv. P\,\E\" shows "\P\ mapME_x f xs \\rv. P\,\E\" apply (induct xs) apply (simp add: mapME_x_def sequenceE_x_def) apply wp apply (simp add: mapME_x_def sequenceE_x_def) apply (fold mapME_x_def sequenceE_x_def) apply wp apply (rule x) apply assumption done lemma liftM_return [simp]: "liftM f (return x) = return (f x)" by (simp add: liftM_def) lemma mapM_x_return : "mapM_x (\_. return v) xs = return v" by (induct xs) (auto simp: mapM_x_Nil mapM_x_Cons) lemma hoare_imp_eq_substR: "\P\ f \Q\,- \ \P\ f \\rv s. rv = x \ Q x s\,-" by (fastforce simp add: valid_def validE_R_def validE_def split: sum.splits) lemma hoare_split_bind_case_sum: assumes x: "\rv. \R rv\ g rv \Q\" "\rv. \S rv\ h rv \Q\" assumes y: "\P\ f \S\,\R\" shows "\P\ f >>= case_sum g h \Q\" apply (rule hoare_seq_ext [OF _ y[unfolded validE_def]]) apply (case_tac x, simp_all add: x) done lemma hoare_split_bind_case_sumE: assumes x: "\rv. \R rv\ g rv \Q\,\E\" "\rv. \S rv\ h rv \Q\,\E\" assumes y: "\P\ f \S\,\R\" shows "\P\ f >>= case_sum g h \Q\,\E\" apply (unfold validE_def) apply (rule hoare_seq_ext [OF _ y[unfolded validE_def]]) apply (case_tac x, simp_all add: x [unfolded validE_def]) done lemma bind_comm_mapM_comm: assumes bind_comm: "\n z. do x \ a; y \ b z; (n x y :: ('a, 's) nondet_monad) od = do y \ b z; x \ a; n x y od" shows "\n'. do x \ a; ys \ mapM b zs; (n' x ys :: ('a, 's) nondet_monad) od = do ys \ mapM b zs; x \ a; n' x ys od" proof (induct zs) case Nil thus ?case by (simp add: mapM_def sequence_def) next case (Cons z zs') thus ?case by (clarsimp simp: mapM_Cons bind_assoc bind_comm intro!: bind_cong [OF refl]) qed lemma liftE_handle : "(liftE f g) = liftE f" by (simp add: handleE_def handleE'_def liftE_def) lemma mapM_empty: "mapM f [] = return []" unfolding mapM_def by (simp add: sequence_def) lemma mapM_append: "mapM f (xs @ ys) = (do x \ mapM f xs; y \ mapM f ys; return (x @ y) od)" proof (induct xs) case Nil thus ?case by (simp add: mapM_empty) next case (Cons x xs) show ?case by (simp add: mapM_Cons bind_assoc Cons.hyps) qed lemma mapM_x_append: "mapM_x f (xs @ ys) = (do x \ mapM_x f xs; y \ mapM_x f ys; return () od)" by (simp add: mapM_x_mapM mapM_append bind_assoc) lemma mapM_singleton: "mapM f [x] = (do r \ f x; return [r] od)" by (simp add: mapM_def sequence_def) lemma mapM_x_singleton: "mapM_x f [x] = f x" by (simp add: mapM_x_mapM mapM_singleton) lemma return_returnOk: "return (Inr x) = returnOk x" unfolding returnOk_def by simp lemma mapME_x_sequenceE: "mapME_x f xs \ doE _ \ sequenceE (map f xs); returnOk () odE" apply (induct xs, simp_all add: mapME_x_def sequenceE_def sequenceE_x_def) apply (simp add: Let_def bindE_assoc) done lemma sequenceE_Cons: "sequenceE (x # xs) = (doE v \ x; vs \ sequenceE xs; returnOk (v # vs) odE)" by (simp add: sequenceE_def Let_def) lemma snd_return [monad_eq]: "\ snd (return a b)" unfolding return_def by simp lemma snd_throwError [monad_eq]: "\ snd (throwError e s)" unfolding throwError_def by (simp add: snd_return) lemma snd_lift_Inr [monad_eq]: "snd (lift b (Inr r) t) = snd (b r t)" unfolding lift_def by simp lemma snd_lift_Inl [monad_eq]: "\ snd (lift b (Inl r) t)" unfolding lift_def by (simp add: snd_throwError) lemma snd_fail [monad_eq]: "snd (fail s)" apply (clarsimp simp: fail_def) done lemma not_snd_bindD: "\ \ snd ((a >>= b) s); (rv, s') \ fst (a s) \ \ \ snd (a s) \ \ snd (b rv s')" by (fastforce simp: bind_def) lemma whenE_bindE_throwError_to_if: "whenE P (throwError e) >>=E (\_. b) = (if P then (throwError e) else b)" unfolding whenE_def bindE_def by (auto simp: NonDetMonad.lift_def throwError_def returnOk_def) lemma not_snd_bindI1: "\ snd ((a >>= b) s) \ \ snd (a s)" by (fastforce simp: bind_def) lemma not_snd_bindI2: "\ \ snd ((a >>= b) s); (rv, s') \ fst (a s) \ \ \ snd (b rv s')" by (fastforce simp: bind_def) lemma empty_fail_not_snd: "\ \ snd (m s); empty_fail m \ \ \v. v \ fst (m s)" by (fastforce simp: empty_fail_def) lemma mapM_Nil: "mapM f [] = return []" by (simp add: mapM_def sequence_def) lemma hoare_vcg_exI: "\P\ f \Q x\ \ \P\ f \\rv s. \x. Q x rv s\" apply (simp add: valid_def split_def) apply blast done lemma hoare_exI_tuple: "\P\ f \\(rv,rv') s. Q x rv rv' s\ \ \P\ f \\(rv,rv') s. \x. Q x rv rv' s\" by (fastforce simp: valid_def) lemma hoare_ex_all: "(\x. \P x\ f \Q\) = \\s. \x. P x s\ f \Q\" apply (rule iffI) apply (fastforce simp: valid_def)+ done lemma empty_fail_bindE: "\ empty_fail f; \rv. empty_fail (g rv) \ \ empty_fail (f >>=E g)" apply (simp add: bindE_def) apply (erule empty_fail_bind) apply (simp add: lift_def throwError_def split: sum.split) done lemma empty_fail_error_bits: "empty_fail (returnOk v)" "empty_fail (throwError v)" "empty_fail (liftE f) = empty_fail f" apply (simp_all add: returnOk_def throwError_def) apply (rule iffI, simp_all add: liftE_def) apply (simp add: empty_fail_def bind_def return_def) apply (erule allEI) apply clarsimp done lemma mapM_upd: assumes "\x rv s s'. (rv,s') \ fst (f x s) \ x \ set xs \ (rv, g s') \ fst (f x (g s))" shows "(rv,s') \ fst (mapM f xs s) \ (rv, g s') \ fst (mapM f xs (g s))" using assms proof (induct xs arbitrary: rv s s') case Nil thus ?case by (simp add: mapM_Nil return_def) next case (Cons z zs) from Cons.prems show ?case apply (clarsimp simp: mapM_Cons in_monad) apply (drule Cons.prems, simp) apply (rule exI, erule conjI) apply (erule Cons.hyps) apply (erule Cons.prems) apply simp done qed definition cutMon :: "('s \ bool) \ ('s, 'a) nondet_monad \ ('s, 'a) nondet_monad" where "cutMon P f \ \s. if P s then f s else fail s" lemma cutMon_walk_bind: "(cutMon (op = s) (f >>= g)) = (cutMon (op = s) f >>= (\rv. cutMon (\s'. (rv, s') \ fst (f s)) (g rv)))" apply (rule ext, simp add: cutMon_def bind_def fail_def) apply (auto simp: split_def) done lemma cutMon_walk_bindE: "(cutMon (op = s) (f >>=E g)) = (cutMon (op = s) f >>=E (\rv. cutMon (\s'. (Inr rv, s') \ fst (f s)) (g rv)))" apply (simp add: bindE_def cutMon_walk_bind) apply (rule bind_cong, rule refl) apply (simp add: cutMon_def lift_def fail_def split: if_split_asm) apply (clarsimp split: sum.split) done lemma cutMon_walk_if: "cutMon (op = s) (if P then f else g) = (if P then cutMon (op = s) f else cutMon (op = s) g)" by (simp add: cutMon_def) lemma cutMon_valid_drop: "\P\ f \Q\ \ \P\ cutMon R f \Q\" by (simp add: cutMon_def valid_def fail_def) lemma cutMon_validE_drop: "\P\ f \Q\,\E\ \ \P\ cutMon R f \Q\,\E\" by (simp add: validE_def cutMon_valid_drop) lemma assertE_assert: "assertE F = liftE (assert F)" by (clarsimp simp: assertE_def assert_def liftE_def returnOk_def split: if_split) lemma snd_cutMon: "snd (cutMon P f s) = (P s \ snd (f s))" by (simp add: cutMon_def fail_def split: if_split) lemma exec_modify: "(modify f >>= g) s = g () (f s)" by (simp add: bind_def simpler_modify_def) lemma no_fail_spec: "\ \s. no_fail ((op = s) and P) f \ \ no_fail P f" by (simp add: no_fail_def) lemma no_fail_assertE [wp]: "no_fail (\_. P) (assertE P)" by (simp add: assertE_def split: if_split) lemma no_fail_spec_pre: "\ no_fail ((op = s) and P') f; \s. P s \ P' s \ \ no_fail ((op = s) and P) f" by (erule no_fail_pre, simp) lemma no_fail_whenE [wp]: "\ G \ no_fail P f \ \ no_fail (\s. G \ P s) (whenE G f)" by (simp add: whenE_def split: if_split) lemma no_fail_unlessE [wp]: "\ \ G \ no_fail P f \ \ no_fail (\s. \ G \ P s) (unlessE G f)" by (simp add: unlessE_def split: if_split) lemma no_fail_throwError [wp]: "no_fail \ (throwError e)" by (simp add: throwError_def) lemma no_fail_liftE [wp]: "no_fail P f \ no_fail P (liftE f)" unfolding liftE_def by wpsimp lemma bind_return_eq: "(a >>= return) = (b >>= return) \ a = b" apply (clarsimp simp:bind_def) apply (rule ext) apply (drule_tac x= x in fun_cong) apply (auto simp:return_def split_def) done lemma bindE_bind_linearise: "((f >>=E g) >>= h) = (f >>= case_sum (h o Inl) (\rv. g rv >>= h))" apply (simp add: bindE_def bind_assoc) apply (rule ext, rule bind_apply_cong, rule refl) apply (simp add: lift_def throwError_def split: sum.split) done lemma throwError_bind: "(throwError e >>= f) = (f (Inl e))" by (simp add: throwError_def) lemma bind_bindE_assoc: "((f >>= g) >>=E h) = f >>= (\rv. g rv >>=E h)" by (simp add: bindE_def bind_assoc) lemma returnOk_bind: "returnOk v >>= f = (f (Inr v))" by (simp add: returnOk_def) lemma liftE_bind: "(liftE m >>= m') = (m >>= (\rv. m' (Inr rv)))" by (simp add: liftE_def) lemma catch_throwError: "catch (throwError ft) g = g ft" by (simp add: catch_def throwError_bind) lemma select_bind_eq2: "\ v = v'; \x. x \ fst v \ f x s = g x s' \ \ (select_f v >>= f) s = (select_f v' >>= g) s'" by (simp add: select_f_def bind_def split_def cart_singleton_image image_image cong: image_cong) lemmas select_bind_eq = select_bind_eq2[OF refl] lemma select_f_singleton_return: "select_f ({v}, False) = return v" by (simp add: select_f_def return_def) lemma select_f_returns: "select_f (return v s) = return (v, s)" "select_f (get s) = return (s, s)" "select_f (gets f s) = return (f s, s)" "select_f (modify g s) = return ((), g s)" by (simp add: select_f_def return_def get_def simpler_gets_def simpler_modify_def)+ lemma select_eq_select_f: "select S = select_f (S, False)" by (simp add: select_def select_f_def) lemma select_f_select_f: "select_f (select_f v s) = liftM (swp Pair s) (select_f v)" apply (rule ext) apply (simp add: select_f_def liftM_def swp_def bind_def return_def split_def image_image image_constant_conv) apply fastforce done lemma select_f_select: "select_f (select S s) = liftM (swp Pair s) (select S)" unfolding select_eq_select_f by (rule select_f_select_f) lemmas select_f_selects = select_f_select_f select_f_select lemma select_f_asserts: "select_f (fail s) = fail" "select_f (assert P s) = do assert P; return ((), s) od" "select_f (assert_opt v s) = do v' \ assert_opt v; return (v', s) od" by (simp add: select_f_def fail_def assert_def return_def bind_def assert_opt_def split: if_split option.split)+ lemma liftE_bindE_handle: "((liftE f >>=E (\x. g x)) h) = f >>= (\x. g x h)" by (simp add: liftE_bindE handleE_def handleE'_def bind_assoc) lemma in_returns [monad_eq]: "(r, s) \ fst (return r s)" "(Inr r, s) \ fst (returnOk r s)" by (simp add: in_monad)+ lemma assertE_sp: "\P\ assertE Q \\rv s. Q \ P s\,\E\" by (clarsimp simp: assertE_def) wp lemma catch_liftE: "catch (liftE g) h = g" by (simp add: catch_def liftE_def) lemma catch_liftE_bindE: "catch (liftE g >>=E (\x. f x)) h = g >>= (\x. catch (f x) h)" by (simp add: liftE_bindE catch_def bind_assoc) lemma returnOk_catch_bind: "catch (returnOk v) h >>= g = g v" by (simp add: returnOk_liftE catch_liftE) lemma alternative_left_readonly_bind: "\ \op = s\ f \\rv. op = s\; fst (f s) \ {} \ \ alternative (f >>= (\x. g x)) h s = (f >>= (\x. alternative (g x) h)) s" apply (subgoal_tac "\x \ fst (f s). snd x = s") apply (clarsimp simp: alternative_def bind_def split_def) apply fastforce apply clarsimp apply (drule(1) use_valid, simp_all) done lemma liftE_bindE_assoc: "(liftE f >>=E g) >>= h = f >>= (\x. g x >>= h)" by (simp add: liftE_bindE bind_assoc) lemma empty_fail_use_cutMon: "\ \s. empty_fail (cutMon (op = s) f) \ \ empty_fail f" apply (clarsimp simp add: empty_fail_def cutMon_def) apply (fastforce split: if_split_asm) done lemma empty_fail_drop_cutMon: "empty_fail f \ empty_fail (cutMon P f)" by (simp add: empty_fail_def fail_def cutMon_def split: if_split) lemma empty_fail_cutMon: "\ \s. P s \ empty_fail (cutMon (op = s) f) \ \ empty_fail (cutMon P f)" apply (clarsimp simp: empty_fail_def cutMon_def fail_def split: if_split) apply (fastforce split: if_split_asm) done lemma empty_fail_If: "\ P \ empty_fail f; \ P \ empty_fail g \ \ empty_fail (if P then f else g)" by (simp split: if_split) lemmas empty_fail_cutMon_intros = cutMon_walk_bind[THEN arg_cong[where f=empty_fail], THEN iffD2, OF empty_fail_bind, OF _ empty_fail_cutMon] cutMon_walk_bindE[THEN arg_cong[where f=empty_fail], THEN iffD2, OF empty_fail_bindE, OF _ empty_fail_cutMon] cutMon_walk_if[THEN arg_cong[where f=empty_fail], THEN iffD2, OF empty_fail_If] lemma empty_fail_whenEs: "empty_fail f \ empty_fail (whenE P f)" "empty_fail f \ empty_fail (unlessE P f)" by (auto simp add: whenE_def unlessE_def empty_fail_error_bits split: if_split) lemma empty_fail_assertE: "empty_fail (assertE P)" by (simp add: assertE_def empty_fail_error_bits split: if_split) lemma unlessE_throw_catch_If: "catch (unlessE P (throwError e) >>=E f) g = (if P then catch (f ()) g else g e)" by (simp add: unlessE_def catch_throwError split: if_split) lemma gets_the_return: "(return x = gets_the f) = (\s. f s = Some x)" apply (subst fun_eq_iff) apply (simp add: return_def gets_the_def exec_gets assert_opt_def fail_def split: option.split) apply auto done lemma gets_the_returns[unfolded K_def]: "(return x = gets_the f) = (\s. f s = Some x)" "(returnOk x = gets_the g) = (\s. g s = Some (Inr x))" "(throwError x = gets_the h) = (\s. h s = Some (Inl x))" by (simp_all add: returnOk_def throwError_def gets_the_return) lemma all_rv_choice_fn_eq: "\ \rv. \fn. f rv = g fn \ \ \fn. f = (\rv. g (fn rv))" using all_rv_choice_fn_eq_pred[where f=f and g=g and P=\] by (simp add: fun_eq_iff) lemma cutMon_assert_opt: "cutMon P (gets_the f >>= g) = gets_the (\s. if P s then f s else None) >>= g" by (simp add: cutMon_def gets_the_def exec_gets bind_assoc fun_eq_iff assert_opt_def split: if_split) lemma gets_the_eq_bind: "\ \fn. f = gets_the (fn o fn'); \rv. \fn. g rv = gets_the (fn o fn') \ \ \fn. (f >>= g) = gets_the (fn o fn')" apply (clarsimp dest!: all_rv_choice_fn_eq) apply (rule_tac x="\s. case (fn s) of None \ None | Some v \ fna v s" in exI) apply (simp add: gets_the_def bind_assoc exec_gets assert_opt_def fun_eq_iff split: option.split) done lemma gets_the_eq_bindE: "\ \fn. f = gets_the (fn o fn'); \rv. \fn. g rv = gets_the (fn o fn') \ \ \fn. (f >>=E g) = gets_the (fn o fn')" apply (simp add: bindE_def) apply (erule gets_the_eq_bind) apply (simp add: lift_def gets_the_returns split: sum.split) apply fastforce done lemma gets_the_fail: "(fail = gets_the f) = (\s. f s = None)" by (simp add: gets_the_def exec_gets assert_opt_def fail_def return_def fun_eq_iff split: option.split) lemma gets_the_asserts: "(fail = gets_the f) = (\s. f s = None)" "(assert P = gets_the g) = (\s. g s = (if P then Some () else None))" "(assertE P = gets_the h) = (\s. h s = (if P then Some (Inr ()) else None))" by (simp add: assert_def assertE_def gets_the_fail gets_the_returns split: if_split)+ lemma gets_the_condsE: "(\fn. whenE P f = gets_the (fn o fn')) = (P \ (\fn. f = gets_the (fn o fn')))" "(\fn. unlessE P g = gets_the (fn o fn')) = (\ P \ (\fn. g = gets_the (fn o fn')))" by (simp add: whenE_def unlessE_def gets_the_returns ex_const_function split: if_split)+ lemma no_fail_gets_the [wp]: "no_fail (\s. f s \ None) (gets_the f)" apply (simp add: gets_the_def) apply (rule no_fail_pre, wp) apply simp done lemma gets_the_validE_R_wp: "\\s. f s \ None \ isRight (the (f s)) \ Q (theRight (the (f s))) s\ gets_the f \Q\,-" apply (simp add: gets_the_def validE_R_def validE_def) apply (wp | wpc | simp add: assert_opt_def)+ apply (clarsimp split: split: sum.splits) done lemma return_bindE: "isRight v \ return v >>=E f = f (theRight v)" by (clarsimp simp: isRight_def return_returnOk) lemma assert_opt_If: "assert_opt v = If (v = None) fail (return (the v))" by (simp_all add: assert_opt_def split: option.split) lemma if_to_top_of_bind: "(bind (If P x y) z) = If P (bind x z) (bind y z)" by (simp split: if_split) lemma if_to_top_of_bindE: "(bindE (If P x y) z) = If P (bindE x z) (bindE y z)" by (simp split: if_split) lemma alternative_bind: "((a \ b) >>= c) = ((a >>= c) \ (b >>= c))" apply (rule ext, simp add: alternative_def bind_def split_def) apply blast done lemma alternative_refl: "(a \ a) = a" by (rule ext, simp add: alternative_def) lemma alternative_com: "(f \ g) = (g \ f)" apply (rule ext) apply (auto simp: alternative_def) done lemma liftE_alternative: "liftE (a \ b) = (liftE a \ liftE b)" by (simp add: liftE_def alternative_bind) lemma fst_return: "fst (return v s) = {(v, s)}" by (simp add: return_def) (* FIXME: move *) lemma in_bind_split [monad_eq]: "(rv \ fst ((f >>= g) s)) = (\rv'. rv' \ fst (f s) \ rv \ fst (g (fst rv') (snd rv')))" apply (cases rv) apply (fastforce simp add: in_bind) done lemma no_fail_mapM_wp: assumes "\x. x \ set xs \ no_fail (P x) (f x)" assumes "\x y. \ x \ set xs; y \ set xs \ \ \P x\ f y \\_. P x\" shows "no_fail (\s. \x \ set xs. P x s) (mapM f xs)" using assms proof (induct xs) case Nil thus ?case by (simp add: mapM_empty) next case (Cons z zs) show ?case apply (clarsimp simp: mapM_Cons) apply (wp Cons.prems Cons.hyps hoare_vcg_const_Ball_lift|simp)+ done qed lemma zipWithM_Nil [simp]: "zipWithM f xs [] = return []" by (simp add: zipWithM_def zipWith_def sequence_def) lemma zipWithM_One: "zipWithM f (x#xs) [a] = (do z \ f x a; return [z] od)" by (simp add: zipWithM_def zipWith_def sequence_def) lemma zipWithM_x_Nil: "zipWithM_x f xs [] = return ()" by (simp add: zipWithM_x_def zipWith_def sequence_x_def) lemma zipWithM_x_One: "zipWithM_x f (x#xs) [a] = f x a" by (simp add: zipWithM_x_def zipWith_def sequence_x_def) lemma list_case_return: "(case xs of [] \ return v | y # ys \ return (f y ys)) = return (case xs of [] \ v | y # ys \ f y ys)" by (simp split: list.split) lemma gets_exs_valid: "\op = s\ gets f \\\r. op = s\" apply (clarsimp simp: exs_valid_def split_def) apply (rule bexI [where x = "(f s, s)"]) apply simp apply (simp add: in_monad) done lemma empty_fail_get: "empty_fail get" by (simp add: empty_fail_def get_def) lemma alternative_liftE_returnOk: "(liftE m \ returnOk v) = liftE (m \ return v)" by (simp add: liftE_def alternative_def returnOk_def bind_def return_def) lemma bind_inv_inv_comm_weak: "\ \s. \op = s\ f \\_. op = s\; \s. \op = s\ g \\_. op = s\; empty_fail f; empty_fail g \ \ do x \ f; y \ g; n od = do y \ g; x \ f; n od" apply (rule ext) apply (fastforce simp: bind_def valid_def empty_fail_def split_def image_def) done lemma mapM_last_Cons: "\ xs = [] \ g v = y; xs \ [] \ do x \ f (last xs); return (g x) od = do x \ f (last xs); return y od \ \ do ys \ mapM f xs; return (g (last (v # ys))) od = do mapM_x f xs; return y od" apply (cases "xs = []") apply (simp add: mapM_x_Nil mapM_Nil) apply (simp add: mapM_x_mapM) apply (subst append_butlast_last_id[symmetric], assumption, subst mapM_append)+ apply (simp add: bind_assoc mapM_Cons mapM_Nil) done lemma mapM_length_cong: "\ length xs = length ys; \x y. (x, y) \ set (zip xs ys) \ f x = g y \ \ mapM f xs = mapM g ys" by (simp add: mapM_def map_length_cong) (* FIXME: duplicate *) lemma zipWithM_mapM: "zipWithM f xs ys = mapM (split f) (zip xs ys)" by (simp add: zipWithM_def zipWith_def mapM_def) lemma zipWithM_If_cut: "zipWithM (\a b. if a < n then f a b else g a b) [0 ..< m] xs = do ys \ zipWithM f [0 ..< min n m] xs; zs \ zipWithM g [n ..< m] (drop n xs); return (ys @ zs) od" apply (cases "n < m") apply (cut_tac i=0 and j=n and k="m - n" in upt_add_eq_append) apply simp apply (simp add: min.absorb1 zipWithM_mapM) apply (simp add: zip_append1 mapM_append zip_take_triv2 split_def) apply (intro bind_cong bind_apply_cong refl mapM_length_cong fun_cong[OF mapM_length_cong]) apply (clarsimp simp: set_zip) apply (clarsimp simp: set_zip) apply (simp add: min.absorb2 zipWithM_mapM mapM_Nil) apply (intro mapM_length_cong refl) apply (clarsimp simp: set_zip) done lemma mapM_liftM_const: "mapM (\x. liftM (\y. f x) (g x)) xs = liftM (\ys. map f xs) (mapM g xs)" apply (induct xs) apply (simp add: mapM_Nil) apply (simp add: mapM_Cons) apply (simp add: liftM_def bind_assoc) done lemma empty_failD2: "\ empty_fail f; \ snd (f s) \ \ \v. v \ fst (f s)" by (fastforce simp add: empty_fail_def) lemma empty_failD3: "\ empty_fail f; \ snd (f s) \ \ fst (f s) \ {}" by (drule(1) empty_failD2, clarsimp) lemma bind_inv_inv_comm: "\ \P. \P\ f \\_. P\; \P. \P\ g \\_. P\; empty_fail f; empty_fail g \ \ do x \ f; y \ g; n x y od = do y \ g; x \ f; n x y od" apply (rule ext) apply (rename_tac s) apply (rule_tac s="(do (x, y) \ do x \ f; y \ (\_. g s) ; (\_. return (x, y) s) od; n x y od) s" in trans) apply (simp add: bind_assoc) apply (intro bind_apply_cong, simp_all)[1] apply (metis in_inv_by_hoareD) apply (simp add: return_def bind_def) apply (metis in_inv_by_hoareD) apply (rule_tac s="(do (x, y) \ do y \ g; x \ (\_. f s) ; (\_. return (x, y) s) od; n x y od) s" in trans[rotated]) apply (simp add: bind_assoc) apply (intro bind_apply_cong, simp_all)[1] apply (metis in_inv_by_hoareD) apply (simp add: return_def bind_def) apply (metis in_inv_by_hoareD) apply (rule bind_apply_cong, simp_all) apply (clarsimp simp: bind_def split_def return_def) apply (auto elim!: nonemptyE | drule(1) empty_failD3)+ done lemma throwErrorE_E [wp]: "\Q e\ throwError e -, \Q\" by (simp add: validE_E_def) wp lemma no_fail_mapM: "\x. no_fail \ (f x) \ no_fail \ (mapM f xs)" apply (induct xs) apply (simp add: mapM_def sequence_def) apply (simp add: mapM_Cons) apply (wp|fastforce)+ done lemma gets_inv [simp]: "\ P \ gets f \ \r. P \" by (simp add: gets_def, wp) lemma select_inv: "\ P \ select S \ \r. P \" by (simp add: select_def valid_def) lemmas return_inv = hoare_return_drop_var lemma assert_inv: "\P\ assert Q \\r. P\" unfolding assert_def by (cases Q) simp+ lemma assert_opt_inv: "\P\ assert_opt Q \\r. P\" unfolding assert_opt_def by (cases Q) simp+ lemma let_into_return: "(let f = x in m f) = (do f \ return x; m f od)" by simp lemma assert_opt_wp [wp]: "\\s. (x \ None) \ Q (the x) s\ assert_opt x \Q\" by (case_tac x, (simp add: assert_opt_def | wp)+) lemma gets_the_wp [wp]: "\\s. (f s \ None) \ Q (the (f s)) s\ gets_the f \Q\" by (unfold gets_the_def, wp) definition injection_handler :: "('a \ 'b) \ ('s, 'a + 'c) nondet_monad \ ('s, 'b + 'c) nondet_monad" where "injection_handler f m \ m (\ft. throwError (f ft))" lemma injection_wp: "\ t = injection_handler f; \P\ m \Q\,\\ft. E (f ft)\ \ \ \P\ t m \Q\,\E\" apply (simp add: injection_handler_def) apply (wp|simp)+ done lemma injection_wp_E: "\ t = injection_handler f; \P\ m \Q\,- \ \ \P\ t m \Q\,-" by (simp add: validE_R_def injection_wp) lemma injection_bindE: "\ t = injection_handler f; t2 = injection_handler f \ \ t (x >>=E y) = (t2 x) >>=E (\rv. t (y rv))" apply (simp add: injection_handler_def) apply (simp add: bindE_def handleE'_def bind_assoc) apply (rule arg_cong [where f="\y. x >>= y"]) apply (rule ext) apply (case_tac x, simp_all add: lift_def throwError_def) done lemma injection_liftE: "t = injection_handler f \ t (liftE x) = liftE x" apply (simp add: injection_handler_def handleE'_def liftE_def) done lemma id_injection: "id = injection_handler id" proof - have P: "case_sum throwError (\v. return (Inr v)) = return" by (auto simp: throwError_def split: sum.splits) show ?thesis by (auto simp: injection_handler_def handleE'_def P) qed lemma case_options_weak_wp: "\ \P\ f \Q\; \x. \P'\ g x \Q\ \ \ \P and P'\ case opt of None \ f | Some x \ g x \Q\" apply (cases opt) apply (clarsimp elim!: hoare_weaken_pre) apply (rule hoare_weaken_pre [where Q=P']) apply simp+ done lemma list_cases_weak_wp: assumes "\P_A\ a \Q\" assumes "\x xs. \P_B\ b x xs \Q\" shows "\P_A and P_B\ case ts of [] \ a | x#xs \ b x xs \Q\" apply (cases ts) apply (simp, rule hoare_weaken_pre, rule assms, simp)+ done lemmas hoare_FalseE_R = hoare_FalseE[where E="\\", folded validE_R_def] lemma hoare_vcg_if_lift2: "\R\ f \\rv s. (P rv s \ X rv s) \ (\ P rv s \ Y rv s)\ \ \R\ f \\rv s. if P rv s then X rv s else Y rv s\" "\R\ f \\rv s. (P' rv \ X rv s) \ (\ P' rv \ Y rv s)\ \ \R\ f \\rv. if P' rv then X rv else Y rv\" by (auto simp: valid_def split_def) lemma hoare_vcg_if_lift_ER: (* Required because of lack of rv in lifting rules *) "\R\ f \\rv s. (P rv s \ X rv s) \ (\ P rv s \ Y rv s)\, - \ \R\ f \\rv s. if P rv s then X rv s else Y rv s\, -" "\R\ f \\rv s. (P' rv \ X rv s) \ (\ P' rv \ Y rv s)\, - \ \R\ f \\rv. if P' rv then X rv else Y rv\, -" by (auto simp: valid_def validE_R_def validE_def split_def) lemma liftME_return: "liftME f (returnOk v) = returnOk (f v)" by (simp add: liftME_def) lemma lifted_if_collapse: "(if P then \ else f) = (\s. \P \ f s)" by auto lemma undefined_valid: "\\\ undefined \Q\" by (rule hoare_pre_cont) lemma Inr_in_liftE_simp [monad_eq]: "((Inr rv, x) \ fst (liftE fn s)) = ((rv, x) \ fst (fn s))" by (simp add: in_monad) lemma assertE_wp: "\\s. F \ Q () s\ assertE F \Q\,\E\" apply (rule hoare_pre) apply (unfold assertE_def) apply wp apply simp done lemma doesn't_grow_proof: assumes y: "\s. finite (S s)" assumes x: "\x. \\s. x \ S s \ P s\ f \\rv s. x \ S s\" shows "\\s. card (S s) < n \ P s\ f \\rv s. card (S s) < n\" apply (clarsimp simp: valid_def) apply (subgoal_tac "S b \ S s") apply (drule card_mono [OF y], simp) apply clarsimp apply (rule ccontr) apply (subgoal_tac "x \ S b", simp) apply (erule use_valid [OF _ x]) apply simp done lemma fold_bindE_into_list_case: "(doE v \ f; case_list (g v) (h v) x odE) = (case_list (doE v \ f; g v odE) (\x xs. doE v \ f; h v x xs odE) x)" by (simp split: list.split) lemma hoare_vcg_propE_R: "\\s. P\ f \\rv s. P\, -" by (simp add: validE_R_def validE_def valid_def split_def split: sum.split) lemma set_preserved_proof: assumes y: "\x. \\s. Q s \ x \ S s\ f \\rv s. x \ S s\" assumes x: "\x. \\s. Q s \ x \ S s\ f \\rv s. x \ S s\" shows "\\s. Q s \ P (S s)\ f \\rv s. P (S s)\" apply (clarsimp simp: valid_def) by (metis (mono_tags, lifting) equalityI post_by_hoare subsetI x y) lemma set_shrink_proof: assumes x: "\x. \\s. x \ S s\ f \\rv s. x \ S s\" shows "\\s. \S'. S' \ S s \ P S'\ f \\rv s. P (S s)\" apply (clarsimp simp: valid_def) apply (drule spec, erule mp) apply (clarsimp simp: subset_iff) apply (rule ccontr) apply (drule(1) use_valid [OF _ x]) apply simp done lemma shrinks_proof: assumes y: "\s. finite (S s)" assumes x: "\x. \\s. x \ S s \ P s\ f \\rv s. x \ S s\" assumes z: "\P\ f \\rv s. x \ S s\" assumes w: "\s. P s \ x \ S s" shows "\\s. card (S s) \ n \ P s\ f \\rv s. card (S s) < n\" apply (clarsimp simp: valid_def) apply (subgoal_tac "S b \ S s") apply (drule psubset_card_mono [OF y], simp) apply (rule psubsetI) apply clarsimp apply (rule ccontr) apply (subgoal_tac "x \ S b", simp) apply (erule use_valid [OF _ x]) apply simp apply (rule not_sym) apply (rule set_neqI[where x=x]) apply (erule w) apply (erule(1) use_valid [OF _ z]) done lemma unlessE_wp : "(\P \ \Q\ f \R\,\E\) \ \if P then R () else Q\ unlessE P f \R\,\E\" apply (clarsimp simp: unlessE_whenE whenE_def) apply wp done lemma use_validE_R: "\ (Inr r, s') \ fst (f s); \P\ f \Q\,-; P s \ \ Q r s'" unfolding validE_R_def validE_def by (frule(2) use_valid, simp) lemma valid_preservation_ex: assumes x: "\x P. \\s. P (f s x :: 'b)\ m \\rv s. P (f s x)\" shows "\\s. P (f s :: 'a \ 'b)\ m \\rv s. P (f s)\" apply (clarsimp simp: valid_def) apply (erule rsubst[where P=P]) apply (rule ext) apply (erule use_valid [OF _ x]) apply simp done lemmas valid_prove_more' = valid_prove_more[where Q="\rv. Q" for Q] lemma whenE_inv: assumes a: "\P\ f \\_. P\" shows "\P\ whenE Q f \\_. P\" unfolding whenE_def apply (cases Q) apply simp apply (wp a) apply simp apply wp done lemma whenE_liftE: "whenE P (liftE f) = liftE (when P f)" by (simp add: whenE_def when_def returnOk_liftE) lemma whenE_throwError_wp: "\\s. \ P \ Q s\ whenE P (throwError e) \\_. Q\, \\\\" unfolding whenE_def apply (cases P) apply simp apply wp apply simp apply wp done lemma whenE_whenE_body: "whenE P (throwError f) >>=E (\_. whenE Q (throwError f) >>=E r) = whenE (P \ Q) (throwError f) >>=E r" apply (cases P) apply (simp add: whenE_def) apply simp done lemma whenE_whenE_same: "whenE P (throwError f) >>=E (\_. whenE P (throwError g) >>=E r) = whenE P (throwError f) >>=E r" apply (cases P) apply (simp add: whenE_def) apply simp done lemma gets_the_inv: "\P\ gets_the V \\rv. P\" by wpsimp lemma select_f_inv: "\P\ select_f S \\_. P\" by (simp add: select_f_def valid_def) lemmas state_unchanged = in_inv_by_hoareD [THEN sym] lemma validI: assumes rl: "\s r s'. \ P s; (r, s') \ fst (S s) \ \ Q r s'" shows "\P\ S \Q\" unfolding valid_def using rl by safe lemma opt_return_pres_lift: assumes x: "\v. \P\ f v \\rv. P\" shows "\P\ case x of None \ return () | Some v \ f v \\rv. P\" by (rule hoare_pre, (wpcw; wp x), simp) lemma exec_select_f_singleton: "(select_f ({v},False) >>= f) = f v" by (simp add: select_f_def bind_def) lemma mapM_discarded: "mapM f xs >>= (\ys. g) = mapM_x f xs >>= (\_. g)" by (simp add: mapM_x_mapM) lemma valid_return_unit: "\P\ f >>= (\_. return ()) \\r. Q\ \ \P\ f \\r. Q\" apply (rule validI) apply (fastforce simp: valid_def return_def bind_def split_def) done lemma mapM_x_map: "mapM_x f (map g xs) = mapM_x (\x. f (g x)) xs" by (simp add: mapM_x_def o_def) lemma maybe_fail_bind_fail: "unless P fail >>= (\_. fail) = fail" "when P fail >>= (\_. fail) = fail" by (clarsimp simp: bind_def fail_def return_def unless_def when_def)+ lemma unless_False: "unless False f = f" by (simp add: unless_def when_def) lemma unless_True: "unless True f = return ()" by (simp add: unless_def when_def) lemma filterM_preserved: "\ \x. x \ set xs \ \P\ m x \\rv. P\ \ \ \P\ filterM m xs \\rv. P\" apply (induct xs) apply (wp | simp | erule meta_mp | drule meta_spec)+ done lemma filterM_append: "filterM f (xs @ ys) = (do xs' \ filterM f xs; ys' \ filterM f ys; return (xs' @ ys') od)" apply (induct xs) apply simp apply (simp add: bind_assoc) apply (rule ext bind_apply_cong [OF refl])+ apply simp done lemma filterM_distinct1: "\\ and K (P \ distinct xs)\ filterM m xs \\rv s. (P \ distinct rv) \ set rv \ set xs\" apply (rule hoare_gen_asm, erule rev_mp) apply (rule rev_induct [where xs=xs]) apply (clarsimp | wp)+ apply (simp add: filterM_append) apply (erule hoare_seq_ext[rotated]) apply (rule hoare_seq_ext[rotated], rule hoare_vcg_prop) apply (wp, clarsimp) apply blast done lemma filterM_subset: "\\\ filterM m xs \\rv s. set rv \ set xs\" by (rule hoare_chain, rule filterM_distinct1[where P=False], simp_all) lemma filterM_all: "\ \x y. \ x \ set xs; y \ set xs \ \ \P y\ m x \\rv. P y\ \ \ \\s. \x \ set xs. P x s\ filterM m xs \\rv s. \x \ set rv. P x s\" apply (rule_tac Q="\rv s. set rv \ set xs \ (\x \ set xs. P x s)" in hoare_strengthen_post) apply (wp filterM_subset hoare_vcg_const_Ball_lift filterM_preserved) apply simp+ apply blast done lemma filterM_distinct: "\K (distinct xs)\ filterM m xs \\rv s. distinct rv\" by (rule hoare_chain, rule filterM_distinct1[where P=True], simp_all) lemma filterM_mapM: "filterM f xs = (do ys \ mapM (\x. do v \ f x; return (x, v) od) xs; return (map fst (filter snd ys)) od)" apply (induct xs) apply (simp add: mapM_def sequence_def) apply (simp add: mapM_Cons bind_assoc) apply (rule bind_cong [OF refl] bind_apply_cong[OF refl])+ apply simp done lemma if_return_closure: "(if P then return x else return y) = (return (if P then x else y))" by simp lemma select_singleton: "select {x} = return x" by (fastforce simp add: fun_eq_iff select_def return_def) lemma static_imp_wp: "\Q\ m \R\ \ \\s. P \ Q s\ m \\rv s. P \ R rv s\" by (cases P, simp_all add: valid_def) lemma static_imp_conj_wp: "\ \Q\ m \Q'\; \R\ m \R'\ \ \ \\s. (P \ Q s) \ R s\ m \\rv s. (P \ Q' rv s) \ R' rv s\" apply (rule hoare_vcg_conj_lift) apply (rule static_imp_wp) apply assumption+ done lemma hoare_eq_P: assumes "\P. \P\ f \\_. P\" shows "\op = s\ f \\_. op = s\" by (rule assms) lemma hoare_validE_R_conj: "\\P\ f \Q\, -; \P\ f \R\, -\ \ \P\ f \Q And R\, -" by (simp add: valid_def validE_def validE_R_def Let_def split_def split: sum.splits) lemma hoare_vcg_const_imp_lift_R: "\P\ f \Q\,- \ \\s. F \ P s\ f \\rv s. F \ Q rv s\,-" by (cases F, simp_all) lemmas throwError_validE_R = throwError_wp [where E="\\", folded validE_R_def] lemma valid_case_option_post_wp: "(\x. \P x\ f \\rv. Q x\) \ \\s. case ep of Some x \ P x s | _ \ True\ f \\rv s. case ep of Some x \ Q x s | _ \ True\" by (cases ep, simp_all add: hoare_vcg_prop) lemma P_bool_lift: assumes t: "\Q\ f \\r. Q\" assumes f: "\\s. \Q s\ f \\r s. \Q s\" shows "\\s. P (Q s)\ f \\r s. P (Q s)\" apply (clarsimp simp: valid_def) apply (subgoal_tac "Q b = Q s") apply simp apply (rule iffI) apply (rule classical) apply (drule (1) use_valid [OF _ f]) apply simp apply (erule (1) use_valid [OF _ t]) done lemma fail_inv : "\ P \ fail \ \r. P \" by wp lemma gets_sp: "\P\ gets f \\rv. P and (\s. f s = rv)\" by (wp, simp) lemma post_by_hoare2: "\ \P\ f \Q\; (r, s') \ fst (f s); P s \ \ Q r s'" by (rule post_by_hoare, assumption+) lemma hoare_Ball_helper: assumes x: "\x. \P x\ f \Q x\" assumes y: "\P. \\s. P (S s)\ f \\rv s. P (S s)\" shows "\\s. \x \ S s. P x s\ f \\rv s. \x \ S s. Q x rv s\" apply (clarsimp simp: valid_def) apply (subgoal_tac "S b = S s") apply (erule post_by_hoare2 [OF x]) apply (clarsimp simp: Ball_def) apply (erule_tac P1="\x. x = S s" in post_by_hoare2 [OF y]) apply (rule refl) done lemma hoare_gets_post: "\ P \ gets f \ \r s. r = f s \ P s \" by (simp add: valid_def get_def gets_def bind_def return_def) lemma hoare_return_post: "\ P \ return x \ \r s. r = x \ P s \" by (simp add: valid_def return_def) lemma mapM_wp': assumes x: "\x. x \ set xs \ \P\ f x \\rv. P\" shows "\P\ mapM f xs \\rv. P\" apply (rule mapM_wp) apply (erule x) apply simp done lemma mapM_set: assumes "\x. x \ set xs \ \P\ f x \\_. P\" assumes "\x y. x \ set xs \ \P\ f x \\_. Q x\" assumes "\x y. \ x \ set xs; y \ set xs \ \ \P and Q y\ f x \\_. Q y\" shows "\P\ mapM f xs \\_ s. \x \ set xs. Q x s\" using assms proof (induct xs) case Nil show ?case by (simp add: mapM_def sequence_def) wp next case (Cons y ys) have PQ_inv: "\x. x \ set ys \ \P and Q y\ f x \\_. P and Q y\" apply (simp add: pred_conj_def) apply (rule hoare_pre) apply (wp Cons|simp)+ done show ?case apply (simp add: mapM_Cons) apply wp apply (rule hoare_vcg_conj_lift) apply (rule hoare_strengthen_post) apply (rule mapM_wp') apply (erule PQ_inv) apply simp apply (wp Cons|simp)+ done qed lemma case_option_fail_return_val: "(fst (case_option fail return v s) = {(rv, s')}) = (v = Some rv \ s = s')" by (cases v, simp_all add: fail_def return_def) lemma return_expanded_inv: "\P\ \s. ({(x, s)}, False) \\rv. P\" by (simp add: valid_def) lemma empty_fail_assert : "empty_fail (assert P)" unfolding assert_def by simp lemma no_fail_mapM': assumes rl: "\x. no_fail (\_. P x) (f x)" shows "no_fail (\_. \x \ set xs. P x) (mapM f xs)" proof (induct xs) case Nil thus ?case by (simp add: mapM_def sequence_def) next case (Cons x xs) have nf: "no_fail (\_. P x) (f x)" by (rule rl) have ih: "no_fail (\_. \x \ set xs. P x) (mapM f xs)" by (rule Cons) show ?case apply (simp add: mapM_Cons) apply (rule no_fail_pre) apply (wp nf ih) apply simp done qed lemma handy_prop_divs: assumes x: "\P. \\s. P (Q s) \ S s\ f \\rv s. P (Q' rv s)\" "\P. \\s. P (R s) \ S s\ f \\rv s. P (R' rv s)\" shows "\\s. P (Q s \ R s) \ S s\ f \\rv s. P (Q' rv s \ R' rv s)\" "\\s. P (Q s \ R s) \ S s\ f \\rv s. P (Q' rv s \ R' rv s)\" apply (clarsimp simp: valid_def elim!: rsubst[where P=P]) apply (rule use_valid [OF _ x(1)], assumption) apply (rule use_valid [OF _ x(2)], assumption) apply simp apply (clarsimp simp: valid_def elim!: rsubst[where P=P]) apply (rule use_valid [OF _ x(1)], assumption) apply (rule use_valid [OF _ x(2)], assumption) apply simp done lemma hoare_as_subst: "\ \P. \\s. P (fn s)\ f \\rv s. P (fn s)\; \v :: 'a. \P v\ f \Q v\ \ \ \\s. P (fn s) s\ f \\rv s. Q (fn s) rv s\" apply (rule_tac Q="\rv s. \v. v = fn s \ Q v rv s" in hoare_chain) apply (rule hoare_vcg_ex_lift) apply (rule hoare_vcg_conj_lift) apply assumption+ apply simp apply simp done lemmas hoare_vcg_ball_lift = hoare_vcg_const_Ball_lift lemma select_singleton_is_return : "select {A} = return A" unfolding select_def return_def by (simp add: Sigma_def) lemma assert_opt_eq_singleton: "(assert_opt f s = ({(v, s')},False)) = (s = s' \ f = Some v)" by (case_tac f, simp_all add: assert_opt_def fail_def return_def conj_comms) lemma hoare_set_preserved: assumes x: "\x. \fn' x\ m \\rv. fn x\" shows "\\s. set xs \ {x. fn' x s}\ m \\rv s. set xs \ {x. fn x s}\" apply (induct xs) apply simp apply wp apply simp apply (rule hoare_vcg_conj_lift) apply (rule x) apply assumption done lemma return_modify: "return () = modify id" by (simp add: return_def simpler_modify_def) lemma liftE_liftM_liftME: "liftE (liftM f m) = liftME f (liftE m)" by (simp add: liftE_liftM liftME_liftM liftM_def) lemma assert_opt_member: "(x, s') \ fst (assert_opt y s) = (y = Some x \ s' = s)" apply (case_tac y, simp_all add: assert_opt_def fail_def return_def) apply safe done lemma bind_return_unit: "f = (f >>= (\x. return ()))" by simp lemma det_mapM: assumes x: "\x. x \ S \ det (f x)" shows "set xs \ S \ det (mapM f xs)" apply (induct xs) apply (simp add: mapM_def sequence_def) apply (simp add: mapM_Cons x) done lemma det_zipWithM_x: assumes x: "\x y. (x, y) \ set (zip xs ys) \ det (f x y)" shows "det (zipWithM_x f xs ys)" apply (simp add: zipWithM_x_mapM) apply (rule bind_detI) apply (rule det_mapM [where S="set (zip xs ys)"]) apply (clarsimp simp add: x) apply simp apply simp done lemma empty_fail_sequence_x : assumes "\m. m \ set ms \ empty_fail m" shows "empty_fail (sequence_x ms)" using assms by (induct ms) (auto simp: sequence_x_def) lemma gets_the_member: "(x, s') \ fst (gets_the f s) = (f s = Some x \ s' = s)" by (case_tac "f s", simp_all add: gets_the_def simpler_gets_def bind_def assert_opt_member) lemma hoare_ex_wp: assumes x: "\x. \P x\ f \Q x\" shows "\\s. \x. P x s\ f \\rv s. \x. Q x rv s\" apply (clarsimp simp: valid_def) apply (rule exI) apply (rule post_by_hoare [OF x]) apply assumption+ done lemma hoare_ex_pre: (* safe, unlike hoare_ex_wp *) "(\x. \P x\ f \Q\) \ \\s. \x. P x s\ f \Q\" by (fastforce simp: valid_def) lemma hoare_ex_pre_conj: "(\x. \\s. P x s \ P' s\ f \Q\) \ \\s. (\x. P x s) \ P' s\ f \Q\" by (fastforce simp: valid_def) lemma hoare_conj_lift_inv: "\\P\ f \Q\; \\s. P' s \ I s\ f \\rv. I\; \s. P s \ P' s\ \ \\s. P s \ I s\ f \\rv s. Q rv s \ I s\" by (fastforce simp: valid_def) lemma hoare_in_monad_post : assumes x: "\P. \P\ f \\x. P\" shows "\\\ f \\rv s. (rv, s) \ fst (f s)\" apply (clarsimp simp: valid_def) apply (subgoal_tac "s = b", simp) apply (simp add: state_unchanged [OF x]) done lemma mapM_gets: assumes P: "\x. m x = gets (f x)" shows "mapM m xs = gets (\s. map (\x. f x s) xs)" proof (induct xs) case Nil show ?case by (simp add: mapM_def sequence_def gets_def get_def bind_def) next case (Cons y ys) thus ?case by (simp add: mapM_Cons P simpler_gets_def return_def bind_def) qed lemma mapM_map_simp: "mapM m (map f xs) = mapM (m \ f) xs" apply (induct xs) apply (simp add: mapM_def sequence_def) apply (simp add: mapM_Cons) done lemma modify_id_return: "modify id = return ()" by (simp add: simpler_modify_def return_def) definition oblivious :: "('a \ 'a) \ ('a, 'b) nondet_monad \ bool" where "oblivious f m \ \s. (\(rv, s') \ fst (m s). (rv, f s') \ fst (m (f s))) \ (\(rv, s') \ fst (m (f s)). \s''. (rv, s'') \ fst (m s) \ s' = f s'') \ snd (m (f s)) = snd (m s)" lemma oblivious_return [simp]: "oblivious f (return x)" by (simp add: oblivious_def return_def) lemma oblivious_fail [simp]: "oblivious f fail" by (simp add: oblivious_def fail_def) lemma oblivious_assert [simp]: "oblivious f (assert x)" by (simp add: assert_def) lemma oblivious_assert_opt [simp]: "oblivious f (assert_opt fn)" by (simp add: assert_opt_def split: option.splits) lemma oblivious_bind: "\ oblivious f m; \rv. oblivious f (m' rv) \ \ oblivious f (m >>= m')" apply atomize apply (simp add: oblivious_def) apply (erule allEI) apply (intro conjI) apply (drule conjunct1) apply (clarsimp simp: in_monad) apply fastforce apply (drule conjunct2, drule conjunct1) apply (clarsimp simp: in_monad) apply fastforce apply (clarsimp simp: bind_def disj_commute) apply (rule disj_cong [OF refl]) apply (rule iffI) apply (clarsimp simp: split_def) apply fastforce apply clarsimp apply (drule(1) bspec) apply (clarsimp simp: split_def) apply (erule (1) my_BallE) apply (rule bexI [rotated], assumption) apply clarsimp done lemma oblivious_gets [simp]: "oblivious f (gets f') = (\s. f' (f s) = f' s)" by (fastforce simp add: oblivious_def simpler_gets_def) lemma oblivious_liftM: "oblivious f m \ oblivious f (liftM g m)" by (simp add: liftM_def oblivious_bind) lemma oblivious_modify [simp]: "oblivious f (modify f') = (\s. f' (f s) = f (f' s))" apply (simp add: oblivious_def simpler_modify_def) apply (rule ball_cong[where A=UNIV, OF refl, simplified]) apply fastforce done lemma oblivious_modify_swap: "\ oblivious f m \ \ (modify f >>= (\rv. m)) = (m >>= (\rv. modify f))" apply (clarsimp simp: bind_def simpler_modify_def) apply (rule ext)+ apply (case_tac "m (f s)", clarsimp) apply (simp add: oblivious_def) apply (drule_tac x=s in spec) apply (rule conjI) apply (rule set_eqI) apply (rule iffI) apply (drule conjunct2, drule conjunct1) apply (drule_tac x=x in bspec, simp) apply clarsimp apply (rule_tac x="((), s'')" in bexI) apply simp apply simp apply (drule conjunct1) apply fastforce apply (drule conjunct2)+ apply fastforce done lemma univ_wp: "\\s. \(rv, s') \ fst (f s). Q rv s'\ f \Q\" by (simp add: valid_def) lemma univ_get_wp: assumes x: "\P. \P\ f \\rv. P\" shows "\\s. \(rv, s') \ fst (f s). s = s' \ Q rv s'\ f \Q\" apply (rule hoare_pre_imp [OF _ univ_wp]) apply clarsimp apply (erule my_BallE, assumption, simp) apply (subgoal_tac "s = b", simp) apply (simp add: state_unchanged [OF x]) done lemma result_in_set_wp : assumes x: "\P. \P\ fn \\rv. P\" shows "\\s. True\ fn \\v s'. (v, s') \ fst (fn s')\" by (rule hoare_pre_imp [OF _ univ_get_wp], simp_all add: x split_def) clarsimp lemma other_result_in_set_wp: assumes x: "\P. \P\ fn \\rv. P\" shows "\\s. \(v, s) \ fst (fn s). F v = v\ fn \\v s'. (F v, s') \ fst (fn s')\" proof - have P: "\v s. (F v = v) \ (v, s) \ fst (fn s) \ (F v, s) \ fst (fn s)" by simp show ?thesis apply (rule hoare_post_imp [OF P], assumption) apply (rule hoare_pre_imp) defer apply (rule hoare_vcg_conj_lift) apply (rule univ_get_wp [OF x]) apply (rule result_in_set_wp [OF x]) apply clarsimp apply (erule my_BallE, assumption, simp) done qed lemma weak_if_wp: "\ \P\ f \Q\; \P'\ f \Q'\ \ \ \P and P'\ f \\r. if C r then Q r else Q' r\" by (auto simp add: valid_def split_def) lemma zipWithM_x_modify: "zipWithM_x (\a b. modify (f a b)) as bs = modify (\s. foldl (\s (a, b). f a b s) s (zip as bs))" apply (simp add: zipWithM_x_def zipWith_def sequence_x_def) apply (induct ("zip as bs")) apply (simp add: simpler_modify_def return_def) apply (rule ext) apply (simp add: simpler_modify_def bind_def split_def) done lemma bindE_split_recursive_asm: assumes x: "\x s'. \ (Inr x, s') \ fst (f s) \ \ \\s. B x s \ s = s'\ g x \C\, \E\" shows "\A\ f \B\, \E\ \ \\st. A st \ st = s\ f >>=E g \C\, \E\" apply (clarsimp simp: validE_def valid_def bindE_def bind_def lift_def) apply (erule allE, erule(1) impE) apply (erule(1) my_BallE, simp) apply (case_tac a, simp_all add: throwError_def return_def) apply (drule x) apply (clarsimp simp: validE_def valid_def) apply (erule(1) my_BallE, simp) done lemma bind_known_operation_eq: "\ no_fail P f; \Q\ f \\rv s. rv = x \ s = t\; P s; Q s; empty_fail f \ \ (f >>= g) s = g x t" apply (drule(1) no_failD) apply (subgoal_tac "fst (f s) = {(x, t)}") apply (clarsimp simp: bind_def) apply (rule not_psubset_eq) apply (drule(1) empty_failD2, clarsimp) apply fastforce apply clarsimp apply (drule(1) use_valid, simp+) done lemma gets_the_bind_eq: "\ f s = Some x; g x s = h s \ \ (gets_the f >>= g) s = h s" by (simp add: gets_the_def bind_assoc exec_gets assert_opt_def) lemma hoare_const_imp_R: "\Q\ f \R\,- \ \\s. P \ Q s\ f \\rv s. P \ R rv s\,-" by (cases P, simp_all) lemma hoare_vcg_imp_lift_R: "\ \P'\ f \\rv s. \ P rv s\, -; \Q'\ f \Q\, - \ \ \\s. P' s \ Q' s\ f \\rv s. P rv s \ Q rv s\, -" by (auto simp add: valid_def validE_R_def validE_def split_def split: sum.splits) lemma hoare_disj_division: "\ P \ Q; P \ \R\ f \S\; Q \ \T\ f \S\ \ \ \\s. (P \ R s) \ (Q \ T s)\ f \S\" apply safe apply (rule hoare_pre_imp) prefer 2 apply simp apply simp apply (rule hoare_pre_imp) prefer 2 apply simp apply simp done lemma hoare_grab_asm: "\ G \ \P\ f \Q\ \ \ \\s. G \ P s\ f \Q\" by (cases G, simp+) lemma hoare_grab_asm2: "(P' \ \\s. P s \ R s\ f \Q\) \ \\s. P s \ P' \ R s\ f \Q\" by (fastforce simp: valid_def) lemma hoare_grab_exs: assumes x: "\x. P x \ \P'\ f \Q\" shows "\\s. \x. P x \ P' s\ f \Q\" apply (clarsimp simp: valid_def) apply (erule(2) use_valid [OF _ x]) done lemma hoare_prop_E: "\\rv. P\ f -,\\rv s. P\" unfolding validE_E_def by (rule hoare_pre, wp, simp) lemma hoare_vcg_conj_lift_R: "\ \P\ f \Q\,-; \R\ f \S\,- \ \ \\s. P s \ R s\ f \\rv s. Q rv s \ S rv s\,-" apply (simp add: validE_R_def validE_def) apply (drule(1) hoare_vcg_conj_lift) apply (erule hoare_strengthen_post) apply (clarsimp split: sum.splits) done lemma hoare_walk_assmsE: assumes x: "\P\ f \\rv. P\" and y: "\s. P s \ Q s" and z: "\P\ g \\rv. Q\" shows "\P\ doE x \ f; g odE \\rv. Q\" apply (wp z) apply (simp add: validE_def) apply (rule hoare_strengthen_post [OF x]) apply (case_tac r, simp_all add: y) done lemma mapME_set: assumes est: "\x. \R\ f x \P\, -" and invp: "\x y. \R and P x\ f y \\_. P x\, -" and invr: "\x. \R\ f x \\_. R\, -" shows "\R\ mapME f xs \\rv s. \x \ set rv. P x s\, -" proof (rule hoare_post_imp_R [where Q' = "\rv s. R s \ (\x \ set rv. P x s)"], induct xs) case Nil thus ?case by (simp add: mapME_Nil | wp returnOKE_R_wp)+ next case (Cons y ys) have minvp: "\x. \R and P x\ mapME f ys \\_. P x\, -" apply (rule hoare_pre) apply (rule_tac Q' = "\_ s. R s \ P x s" in hoare_post_imp_R) apply (wp mapME_wp) apply (fold validE_R_def) apply (wp hoare_vcg_R_conj [OF invr invp]) apply simp apply (rule subset_refl) apply simp apply simp done show ?case apply (simp add: mapME_Cons) apply (wp) apply (rule_tac Q' = "\xs s. (R s \ (\x \ set xs. P x s)) \ P x s" in hoare_post_imp_R) apply (wp Cons.hyps minvp) apply simp apply (fold validE_R_def) apply simp apply (wp invr est) apply simp done qed clarsimp lemma unlessE_throwError_returnOk: "(if P then returnOk v else throwError x) = (unlessE P (throwError x) >>=E (\_. returnOk v))" by (cases P, simp_all add: unlessE_def) lemma weaker_hoare_ifE: assumes x: "\P \ a \Q\,\E\" assumes y: "\P'\ b \Q\,\E\" shows "\P and P'\ if test then a else b \Q\,\E\" apply (rule hoare_vcg_precond_impE) apply (wp x y) apply simp done lemma wp_split_const_if: assumes x: "\P\ f \Q\" assumes y: "\P'\ f \Q'\" shows "\\s. (G \ P s) \ (\ G \ P' s)\ f \\rv s. (G \ Q rv s) \ (\ G \ Q' rv s)\" by (case_tac G, simp_all add: x y) lemma wp_split_const_if_R: assumes x: "\P\ f \Q\,-" assumes y: "\P'\ f \Q'\,-" shows "\\s. (G \ P s) \ (\ G \ P' s)\ f \\rv s. (G \ Q rv s) \ (\ G \ Q' rv s)\,-" by (case_tac G, simp_all add: x y) lemma wp_throw_const_imp: assumes x: "\P\ f \Q\" shows "\\s. G \ P s\ f \\rv s. G \ Q rv s\" by (case_tac G, simp_all add: x hoare_vcg_prop) lemma wp_throw_const_impE: assumes x: "\P\ f \Q\,\E\" shows "\\s. G \ P s\ f \\rv s. G \ Q rv s\,\\rv s. G \ E rv s\" apply (case_tac G, simp_all add: x) apply wp done lemma distinct_tuple_helper: "\P\ f \\rv s. distinct (x # xs rv s)\ \ \P\ f \\rv s. distinct (x # (map fst (zip (xs rv s) (ys rv s))))\" apply (erule hoare_strengthen_post) apply (erule distinct_prefix) apply (simp add: map_fst_zip_prefix) done lemma hoare_whenE_wp: "(P \ \Q\ f \R\, \E\) \ \if P then Q else R ()\ whenE P f \R\, \E\" unfolding whenE_def apply clarsimp apply wp done lemma list_case_throw_validE_R: "\ \y ys. xs = y # ys \ \P\ f y ys \Q\,- \ \ \P\ case xs of [] \ throwError e | x # xs \ f x xs \Q\,-" apply (case_tac xs, simp_all) apply wp done lemma validE_R_sp: assumes x: "\P\ f \Q\,-" assumes y: "\x. \Q x\ g x \R\,-" shows "\P\ f >>=E (\x. g x) \R\,-" by (rule hoare_pre, wp x y, simp) lemma valid_set_take_helper: "\P\ f \\rv s. \x \ set (xs rv s). Q x rv s\ \ \P\ f \\rv s. \x \ set (take (n rv s) (xs rv s)). Q x rv s\" apply (erule hoare_strengthen_post) apply (clarsimp dest!: in_set_takeD) done lemma whenE_throwError_sp: "\P\ whenE Q (throwError e) \\rv s. \ Q \ P s\, -" apply (simp add: whenE_def validE_R_def) apply (intro conjI impI; wp) done lemma no_fail_bindE [wp]: "\ no_fail P f; \rv. no_fail (R rv) (g rv); \Q\ f \R\,- \ \ no_fail (P and Q) (f >>=E g)" apply (simp add: bindE_def) apply (erule no_fail_bind) apply (simp add: lift_def) apply wpc apply (simp add: throwError_def) apply wp apply assumption apply (simp add: validE_R_def validE_def) apply (erule hoare_strengthen_post) apply clarsimp done lemma empty_fail_mapM_x [simp]: "(\x. empty_fail (a x)) \ empty_fail (mapM_x a xs)" apply (induct_tac xs) apply (clarsimp simp: mapM_x_Nil) apply (clarsimp simp: mapM_x_Cons) done lemma fst_throwError_returnOk: "fst (throwError e s) = {(Inl e, s)}" "fst (returnOk v s) = {(Inr v, s)}" by (simp add: throwError_def returnOk_def return_def)+ lemma liftE_bind_return_bindE_returnOk: "liftE (v >>= (\rv. return (f rv))) = (liftE v >>=E (\rv. returnOk (f rv)))" by (simp add: liftE_bindE, simp add: liftE_def returnOk_def) lemma bind_eqI: "g = g' \ f >>= g = f >>= g'" by simp lemma not_snd_bindE_I1: "\ snd ((a >>=E b) s) \ \ snd (a s)" unfolding bindE_def by (erule not_snd_bindI1) lemma snd_assert [monad_eq]: "snd (assert P s) = (\ P)" apply (clarsimp simp: fail_def return_def assert_def) done lemma not_snd_assert : "(\ snd (assert P s)) = P" by (metis snd_assert) lemma snd_assert_opt [monad_eq]: "snd (assert_opt f s) = (f = None)" by (monad_eq simp: assert_opt_def split: option.splits) declare in_assert_opt [monad_eq] lemma empty_fail_catch: "\ empty_fail f; \x. empty_fail (g x) \ \ empty_fail (catch f g)" apply (simp add: catch_def) apply (erule empty_fail_bind) apply (simp split: sum.split) done lemma oblivious_returnOk [simp]: "oblivious f (returnOk e)" by (simp add: returnOk_def) lemma oblivious_assertE [simp]: "oblivious f (assertE P)" by (simp add: assertE_def split: if_split) lemma oblivious_throwError [simp]: "oblivious f (throwError e)" by (simp add: throwError_def) lemma oblivious_bindE: "\ oblivious u f; \v. oblivious u (g v) \ \ oblivious u (f >>=E (\v. g v))" apply (simp add: bindE_def) apply (erule oblivious_bind) apply (simp add: lift_def split: sum.split) done lemma oblivious_catch: "\ oblivious u f; \v. oblivious u (g v) \ \ oblivious u (catch f g)" apply (simp add: catch_def) apply (erule oblivious_bind) apply (simp split: sum.split) done lemma oblivious_when [simp]: "oblivious f (when P m) = (P \ oblivious f m)" by (simp add: when_def split: if_split) lemma oblivious_whenE [simp]: "oblivious f (whenE P g) = (P \ oblivious f g)" by (simp add: whenE_def split: if_split) lemma select_f_oblivious [simp]: "oblivious f (select_f v)" by (simp add: oblivious_def select_f_def) lemma oblivious_select: "oblivious f (select S)" by (simp add: oblivious_def select_def) lemma validE_R_abstract_rv: "\P\ f \\rv s. \rv'. Q rv' s\,- \ \P\ f \Q\,-" by (erule hoare_post_imp_R, simp) lemma validE_cases_valid: "\P\ f \\rv s. Q (Inr rv) s\,\\rv s. Q (Inl rv) s\ \ \P\ f \Q\" apply (simp add: validE_def) apply (erule hoare_strengthen_post) apply (simp split: sum.split_asm) done lemma valid_isRight_theRight_split: "\P\ f \\rv s. Q True rv s\,\\e s. \v. Q False v s\ \ \P\ f \\rv s. Q (isRight rv) (theRight rv) s\" apply (simp add: validE_def) apply (erule hoare_strengthen_post) apply (simp add: isRight_def split: sum.split_asm) done lemma bind_return_subst: assumes r: "\r. \\s. P x = r\ f x \\rv s. Q rv = r\" shows "do a \ f x; g (Q a) od = do _ \ f x; g (P x) od" proof - have "do a \ f x; return (Q a) od = do _ \ f x; return (P x) od" using r apply (subst fun_eq_iff) apply (fastforce simp: bind_def valid_def return_def) done hence "do a \ f x; return (Q a) od >>= g = do _ \ f x; return (P x) od >>= g" by (rule bind_cong, simp) thus ?thesis by simp qed lemma zipWithM_x_Nil2 : "zipWithM_x f xs [] = return ()" by (simp add: zipWithM_x_mapM mapM_Nil) lemma assert2: "(do v1 \ assert P; v2 \ assert Q; c od) = (do v \ assert (P \ Q); c od)" by (simp add: assert_def split: if_split) lemma assert_opt_def2: "assert_opt v = (do assert (v \ None); return (the v) od)" by (simp add: assert_opt_def split: option.split) lemma filterM_voodoo: "\ys. P ys (do zs \ filterM m xs; return (ys @ zs) od) \ P [] (filterM m xs)" by (drule spec[where x=Nil], simp) lemma gets_assert: "(do v1 \ assert v; v2 \ gets f; c v1 v2 od) = (do v2 \ gets f; v1 \ assert v; c v1 v2 od)" by (simp add: simpler_gets_def return_def assert_def fail_def bind_def split: if_split) lemma list_case_return2: "(case x of [] \ return v | y # ys \ return (f y ys)) = return (case x of [] \ v | y # ys \ f y ys)" by (simp split: list.split) lemma modify_assert: "(do v2 \ modify f; v1 \ assert v; c v1 od) = (do v1 \ assert v; v2 \ modify f; c v1 od)" by (simp add: simpler_modify_def return_def assert_def fail_def bind_def split: if_split) lemma gets_fold_into_modify: "do x \ gets f; modify (g x) od = modify (\s. g (f s) s)" "do x \ gets f; _ \ modify (g x); h od = do modify (\s. g (f s) s); h od" by (simp_all add: fun_eq_iff modify_def bind_assoc exec_gets exec_get exec_put) lemma bind_assoc2: "(do x \ a; _ \ b; c x od) = (do x \ (do x' \ a; _ \ b; return x' od); c x od)" by (simp add: bind_assoc) lemma liftM_pre: assumes rl: "\\s. \ P s \ a \ \_ _. False \" shows "\\s. \ P s \ liftM f a \ \_ _. False \" unfolding liftM_def apply (rule seq) apply (rule rl) apply wp apply simp done lemma not_snd_bindD': "\\ snd ((a >>= b) s); \ snd (a s) \ (rv, s') \ fst (a s)\ \ \ snd (a s) \ \ snd (b rv s')" apply (frule not_snd_bindI1) apply (erule not_snd_bindD) apply simp done lemma snd_bind [monad_eq]: "snd ((a >>= b) s) = (snd (a s) \ (\r s'. (r, s') \ fst (a s) \ snd (b r s')))" apply (clarsimp simp add: bind_def Bex_def image_def) apply (subst surjective_pairing, subst prod.inject, force) done lemma in_lift [monad_eq]: "(rv, s') \ fst (lift M v s) = (case v of Inl x \ rv = Inl x \ s' = s | Inr x \ (rv, s') \ fst (M x s))" apply (clarsimp simp: lift_def throwError_def return_def split: sum.splits) done lemma snd_lift [monad_eq]: "snd (lift M a b) = (\x. a = Inr x \ snd (M x b))" apply (clarsimp simp: lift_def throwError_def return_def split: sum.splits) done lemma snd_bindE [monad_eq]: "snd ((a >>=E b) s) = (snd (a s) \ (\r s'. (r, s') \ fst (a s) \ (\a. r = Inr a \ snd (b a s'))))" apply (clarsimp simp: bindE_def) apply monad_eq done lemma snd_get [monad_eq]: shows "(snd (get s)) = False" by (simp add: get_def) lemma snd_gets [monad_eq]: shows "(snd (gets f s)) = False" by (simp add: gets_def snd_bind snd_get snd_return) lemma mapME_x_Cons: "mapME_x f (x # xs) = (doE f x; mapME_x f xs odE)" by (simp add: mapME_x_def sequenceE_x_def) lemma liftME_map_mapME: "liftME (map f) (mapME m xs) = mapME (liftME f o m) xs" apply (rule sym) apply (induct xs) apply (simp add: liftME_def mapME_Nil) apply (simp add: mapME_Cons liftME_def bindE_assoc) done lemma mapM_upd_inv: assumes f: "\x rv. (rv,s) \ fst (f x s) \ x \ set xs \ (rv, g s) \ fst (f x (g s))" assumes inv: "\x. \op = s\ f x \\_. op = s\" shows "(rv,s) \ fst (mapM f xs s) \ (rv, g s) \ fst (mapM f xs (g s))" using f inv proof (induct xs arbitrary: rv s) case Nil thus ?case by (simp add: mapM_Nil return_def) next case (Cons z zs) from Cons.prems show ?case apply (clarsimp simp: mapM_Cons in_monad) apply (frule use_valid, assumption, rule refl) apply clarsimp apply (drule Cons.prems, simp) apply (rule exI, erule conjI) apply (drule Cons.hyps) apply simp apply assumption apply simp done qed (* FUXME: duplicate *) lemma mapM_x_append2: "mapM_x f (xs @ ys) = do mapM_x f xs; mapM_x f ys od" apply (simp add: mapM_x_def sequence_x_def) apply (induct xs) apply simp apply (simp add: bind_assoc) done lemma mapM_x_split_append: "mapM_x f xs = do _ \ mapM_x f (take n xs); mapM_x f (drop n xs) od" using mapM_x_append[where f=f and xs="take n xs" and ys="drop n xs"] by simp lemma hoare_gen_asm': "(P \ \P'\ f \Q\) \ \P' and (\_. P)\ f \Q\" apply (auto intro: hoare_assume_pre) done lemma hoare_gen_asm_conj: "(P \ \P'\ f \Q\) \ \\s. P' s \ P\ f \Q\" by (fastforce simp: valid_def) lemma hoare_add_K: "\P\ f \Q\ \ \\s. P s \ I\ f \\rv s. Q rv s \ I\" by (fastforce simp: valid_def) lemma valid_rv_lift: "\P'\ f \\rv s. rv \ Q rv s\ \ \\s. P \ P' s\ f \\rv s. rv \ P \ Q rv s\" by (fastforce simp: valid_def) lemma valid_imp_ex: "\P\ f \\rv s. \x. rv \ Q rv s x\ \ \P\ f \\rv s. rv \ (\x. Q rv s x)\" by (fastforce simp: valid_def) lemma valid_rv_split: "\\P\ f \\rv s. rv \ Q s\; \P\ f \\rv s. \rv \ Q' s\\ \ \P\ f \\rv s. if rv then Q s else Q' s\" by (fastforce simp: valid_def) lemma hoare_rv_split: "\\P\ f \\rv s. rv \ (Q rv s)\; \P\ f \\rv s. (\rv) \ (Q rv s)\\ \ \P\ f \Q\" apply (clarsimp simp: valid_def) apply (case_tac a, fastforce+) done lemma case_option_find_give_me_a_map: "case_option a return (find f xs) = liftM theLeft (mapME (\x. if (f x) then throwError x else returnOk ()) xs >>=E (\x. assert (\x \ set xs. \ f x) >>= (\_. liftM (Inl :: 'a \ 'a + unit) a)))" apply (induct xs) apply simp apply (simp add: liftM_def mapME_Nil) apply (simp add: mapME_Cons split: if_split) apply (clarsimp simp add: throwError_def bindE_def bind_assoc liftM_def) apply (rule bind_cong [OF refl]) apply (simp add: lift_def throwError_def returnOk_def split: sum.split) done lemma if_bind: "(if P then (a >>= (\_. b)) else return ()) = (if P then a else return ()) >>= (\_. if P then b else return ())" apply (cases P) apply simp apply simp done lemma in_handleE' [monad_eq]: "((rv, s') \ fst ((f g) s)) = ((\ex. rv = Inr ex \ (Inr ex, s') \ fst (f s)) \ (\rv' s''. (rv, s') \ fst (g rv' s'') \ (Inl rv', s'') \ fst (f s)))" apply (clarsimp simp: handleE'_def) apply (rule iffI) apply (subst (asm) in_bind_split) apply (clarsimp simp: return_def split: sum.splits) apply (case_tac a) apply (erule allE, erule (1) impE) apply clarsimp apply (erule allE, erule (1) impE) apply clarsimp apply (subst in_bind_split) apply (clarsimp simp: return_def split: sum.splits) apply blast done lemma in_handleE [monad_eq]: "(a, b) \ fst ((A B) s) = ((\x. a = Inr x \ (Inr x, b) \ fst (A s)) \ (\r t. ((Inl r, t) \ fst (A s)) \ ((a, b) \ fst (B r t))))" apply (unfold handleE_def) apply (monad_eq split: sum.splits) apply blast done lemma snd_handleE' [monad_eq]: "snd ((A B) s) = (snd (A s) \ (\r s'. (r, s')\fst (A s) \ (\a. r = Inl a \ snd (B a s'))))" apply (clarsimp simp: handleE'_def) apply (monad_eq simp: Bex_def split: sum.splits) apply (metis sum.sel(1) sum.distinct(1) sumE) done lemma snd_handleE [monad_eq]: "snd ((A B) s) = (snd (A s) \ (\r s'. (r, s')\fst (A s) \ (\a. r = Inl a \ snd (B a s'))))" apply (unfold handleE_def) apply (rule snd_handleE') done declare in_liftE [monad_eq] lemma snd_liftE [monad_eq]: "snd ((liftE x) s) = snd (x s)" apply (clarsimp simp: liftE_def snd_bind snd_return) done lemma snd_returnOk [monad_eq]: "\ snd (returnOk x s)" apply (clarsimp simp: returnOk_def return_def) done lemma snd_when [monad_eq]: "snd (when P M s) = (P \ snd (M s))" apply (clarsimp simp: when_def return_def) done lemma bind_liftE_distrib: "(liftE (A >>= (\x. B x))) = (liftE A >>=E (\x. liftE (\s. B x s)))" apply (clarsimp simp: liftE_def bindE_def lift_def bind_assoc) done lemma in_condition [monad_eq]: "((a, b) \ fst (condition C L R s)) = ((C s \ (a, b) \ fst (L s)) \ (\ C s \ (a, b) \ fst (R s)))" by (rule condition_split) lemma snd_condition [monad_eq]: "(snd (condition C L R s)) = ((C s \ snd (L s)) \ (\ C s \ snd (R s)))" by (rule condition_split) lemma condition_apply_cong: "\ c s = c' s'; s = s'; \s. c' s \ l s = l' s ; \s. \ c' s \ r s = r' s \ \ condition c l r s = condition c' l' r' s'" apply (clarsimp split: condition_splits) done lemma condition_cong [cong, fundef_cong]: "\ c = c'; \s. c' s \ l s = l' s; \s. \ c' s \ r s = r' s \ \ condition c l r = condition c' l' r'" apply (rule ext) apply (clarsimp split: condition_splits) done (* Alternative definition of no_throw; easier to work with than unfolding validE. *) lemma no_throw_def': "no_throw P A = (\s. P s \ (\(r, t) \ fst (A s). (\x. r = Inr x)))" apply (clarsimp simp: no_throw_def validE_def2 split_def split: sum.splits) done lemma no_throw_returnOk [simp]: "no_throw P (returnOk a)" apply (clarsimp simp: no_throw_def) apply wp done lemma no_throw_liftE [simp]: "no_throw P (liftE x)" apply (clarsimp simp: liftE_def no_throw_def validE_def) apply (wp | simp)+ done lemma no_throw_bindE: "\ no_throw A X; \a. no_throw B (Y a); \ A \ X \ \_. B \,\ \_ _. True \ \ \ no_throw A (X >>=E Y)" apply atomize apply (clarsimp simp: no_throw_def) apply (rule seqE [rotated]) apply force apply (rule hoare_validE_cases) apply simp apply simp done lemma no_throw_bindE_simple: "\ no_throw \ L; \x. no_throw \ (R x) \ \ no_throw \ (L >>=E R)" apply (erule no_throw_bindE) apply assumption apply wp done lemma no_throw_handleE_simple: notes hoare_pre [wp_pre del] shows "\ \x. no_throw \ L \ no_throw \ (R x) \ \ no_throw \ (L R)" apply (clarsimp simp: no_throw_def) apply atomize apply clarsimp apply (erule disjE) apply wp apply simp apply (rule handleE_wp) apply (erule_tac x=x in allE) apply assumption apply wp done lemma no_throw_handle2: "\ \a. no_throw Y (B a); \ X \ A \ \_ _. True \,\ \_. Y \ \ \ no_throw X (A B)" apply atomize apply (clarsimp simp: no_throw_def' handleE'_def bind_def) apply (clarsimp simp: validE_def valid_def) apply (erule allE, erule (1) impE) apply (erule (1) my_BallE) apply (clarsimp simp: return_def split: sum.splits) apply force done lemma no_throw_handle: "\ \a. no_throw Y (B a); \ X \ A \ \_ _. True \,\ \_. Y \ \ \ no_throw X (A B)" apply (unfold handleE_def) apply (erule (1) no_throw_handle2) done lemma no_throw_fail [simp]: "no_throw P fail" apply (clarsimp simp: no_throw_def) done lemma lift_Inr [simp]: "(lift X (Inr r)) = (X r)" apply (rule ext)+ apply (clarsimp simp: lift_def bind_def split_def image_def) done lemma lift_Inl [simp]: "lift C (Inl a) = throwError a" apply (clarsimp simp: lift_def throwError_def) done lemma returnOk_def2: "returnOk a = return (Inr a)" apply (clarsimp simp: returnOk_def return_def) done lemma empty_fail_spec [simp]: "empty_fail (state_select F)" apply (clarsimp simp: state_select_def empty_fail_def) done declare snd_fail [simp] lemma empty_fail_select [simp]: "empty_fail (select V) = (V \ {})" apply (clarsimp simp: select_def empty_fail_def) done lemma bind_fail_propagates: "\ empty_fail A \ \ A >>= (\_. fail) = fail" apply (monad_eq simp: empty_fail_def) by fastforce lemma bindE_fail_propagates: "\ no_throw \ A; empty_fail A \ \ A >>=E (\_. fail) = fail" apply (rule ext) apply (clarsimp simp: empty_fail_def) apply (clarsimp simp: no_throw_def validE_def valid_def bind_def bindE_def split_def fail_def NonDetMonad.lift_def throwError_def split:sum.splits) apply fastforce done lemma empty_fail_liftE [simp]: "empty_fail (liftE X) = empty_fail X" apply (simp add: empty_fail_error_bits) done declare snd_returnOk [simp, monad_eq] lemma liftE_fail [simp]: "liftE fail = fail" apply monad_eq done lemma in_catch [monad_eq]: "(r, t) \ fst ((M E) s) = ((Inr r, t) \ fst (M s) \ (\r' s'. ((Inl r', s') \ fst (M s)) \ (r, t) \ fst (E r' s')))" apply (rule iffI) apply (clarsimp simp: catch_def in_bind in_return split: sum.splits) apply (metis sumE) apply (clarsimp simp: catch_def in_bind in_return split: sum.splits) apply (metis sum.sel(1) sum.distinct(1) sum.inject(2)) done lemma snd_catch [monad_eq]: "snd ((M E) s) = (snd (M s) \ (\r' s'. ((Inl r', s') \ fst (M s)) \ snd (E r' s')))" apply (rule iffI) apply (clarsimp simp: catch_def snd_bind snd_return split: sum.splits) apply (clarsimp simp: catch_def snd_bind snd_return split: sum.splits) apply force done lemma in_get [monad_eq]: "(r, s') \ fst (get s) = (r = s \ s' = s)" apply (clarsimp simp: get_def) done lemma returnOk_cong: "\ \s. B a s = B' a s \ \ ((returnOk a) >>=E B) = ((returnOk a) >>=E B')" apply monad_eq done lemma in_state_assert [monad_eq, simp]: "(rv, s') \ fst (state_assert P s) = (rv = () \ s' = s \ P s)" apply (monad_eq simp: state_assert_def) apply metis done lemma snd_state_assert [monad_eq]: "snd (state_assert P s) = (\ P s)" apply (monad_eq simp: state_assert_def Bex_def) done lemma state_assert_false [simp]: "state_assert (\_. False) = fail" by monad_eq lemma no_fail_state_assert [wp]: "no_fail P (state_assert P)" by (monad_eq simp: no_fail_def state_assert_def) lemma no_fail_modify [wp]: "no_fail \ (modify M)" by (metis non_fail_modify) lemma combine_validE: "\ \ P \ x \ Q \,\ E \; \ P' \ x \ Q' \,\ E' \ \ \ \ P and P' \ x \ \r. (Q r) and (Q' r) \,\\r. (E r) and (E' r) \" apply (clarsimp simp: validE_def valid_def split: sum.splits) apply (erule allE, erule (1) impE)+ apply (erule (1) my_BallE)+ apply clarsimp done lemma condition_swap: "(condition C A B) = (condition (\s. \ C s) B A)" apply (rule ext) apply (clarsimp split: condition_splits) done lemma condition_fail_rhs: "(condition C X fail) = (state_assert C >>= (\_. X))" apply (rule ext) apply (monad_eq simp: Bex_def) done lemma condition_fail_lhs: "(condition C fail X) = (state_assert (\s. \ C s) >>= (\_. X))" apply (metis condition_fail_rhs condition_swap) done lemma condition_bind_fail [simp]: "(condition C A B >>= (\_. fail)) = condition C (A >>= (\_. fail)) (B >>= (\_. fail))" apply monad_eq apply blast done lemma no_throw_Inr: "\ x \ fst (A s); no_throw P A; P s \ \ \y. fst x = Inr y" apply (clarsimp simp: no_throw_def' split: sum.splits) apply (erule allE, erule (1) impE, erule (1) my_BallE) apply clarsimp done lemma no_throw_handleE': "no_throw \ A \ (A B) = A" apply (rule monad_eqI) apply monad_eq apply (fastforce dest: no_throw_Inr) apply monad_eq apply (metis (lifting) fst_conv no_throw_Inr) apply monad_eq apply (fastforce dest: no_throw_Inr) done lemma no_throw_handleE: "no_throw \ A \ (A B) = A" apply (unfold handleE_def) apply (subst no_throw_handleE') apply auto done lemma whileLoopE_nothrow: "\ \x. no_throw \ (B x) \ \ no_throw \ (whileLoopE C B x)" apply atomize apply (clarsimp simp: no_throw_def) apply (rule validE_whileLoopE [where I="\_ _. True"]) apply simp apply (rule validE_weaken) apply force apply simp apply simp apply simp apply simp done lemma handleE'_nothrow_lhs: "\ no_throw \ L \ \ no_throw \ (L R)" apply (clarsimp simp: no_throw_def) apply (rule handleE'_wp [rotated]) apply assumption apply simp done lemma handleE'_nothrow_rhs: "\ \x. no_throw \ (R x) \ \ no_throw \ (L R)" apply atomize apply (clarsimp simp: no_throw_def) apply (rule handleE'_wp) apply (erule allE) apply assumption apply (rule hoareE_TrueI) done lemma handleE_nothrow_lhs: "\ no_throw \ L \ \ no_throw \ (L R)" by (metis handleE'_nothrow_lhs handleE_def) lemma handleE_nothrow_rhs: "\ \x. no_throw \ (R x) \ \ no_throw \ (L R)" by (metis no_throw_handleE_simple) lemma condition_nothrow: "\ no_throw \ L; no_throw \ R \ \ no_throw \ (condition C L R)" apply (clarsimp simp: condition_def no_throw_def validE_def2) done lemma empty_fail_guard [simp]: "empty_fail (state_assert G)" apply (clarsimp simp: state_assert_def assert_def empty_fail_def get_def return_def bind_def) done lemma simple_bind_fail [simp]: "(state_assert X >>= (\_. fail)) = fail" "(modify M >>= (\_. fail)) = fail" "(return X >>= (\_. fail)) = fail" "(gets X >>= (\_. fail)) = fail" apply (auto intro!: bind_fail_propagates) done lemma valid_case_prod: "\ \x y. valid (P x y) (f x y) Q \ \ valid (case_prod P v) (case_prod (\x y. f x y) v) Q" by (simp add: split_def) lemma validE_case_prod: "\ \x y. validE (P x y) (f x y) Q E \ \ validE (case_prod P v) (case_prod (\x y. f x y) v) Q E" by (simp add: split_def) lemma in_select [monad_eq]: "(rv, s') \ fst (select S s) = (s' = s \ rv \ S)" apply (clarsimp simp: select_def) apply blast done lemma snd_select [monad_eq]: "\ snd (select S s)" by (clarsimp simp: select_def) lemma bindE_handleE_join: "no_throw \ A \ (A >>=E (\x. (B x) C)) = ((A >>=E B C))" apply (monad_eq simp: Bex_def Ball_def no_throw_def') apply blast done lemma snd_put [monad_eq]: "\ snd (put t s)" by (clarsimp simp: put_def) lemma snd_modify [monad_eq]: "\ snd (modify t s)" by (clarsimp simp: modify_def put_def get_def bind_def) lemma no_fail_False [simp]: "no_fail (\_. False) X" by (clarsimp simp: no_fail_def) lemma valid_pre_satisfies_post: "\ \s r' s'. P s \ Q r' s' \ \ \ P \ m \ Q \" by (clarsimp simp: valid_def) lemma validE_pre_satisfies_post: "\ \s r' s'. P s \ Q r' s'; \s r' s'. P s \ R r' s' \ \ \ P \ m \ Q \,\ R \" by (clarsimp simp: validE_def2 split: sum.splits) lemma snd_gets_the [monad_eq]: "snd (gets_the X s) = (X s = None)" by (monad_eq simp: gets_the_def gets_def get_def) lemma validE_K_bind [wp_comb]: "\ P \ x \ Q \, \ E \ \ \ P \ K_bind x f \ Q \, \ E \" by simp lemma liftE_K_bind: "liftE ((K_bind (\s. A s)) x) = K_bind (liftE (\s. A s)) x" by clarsimp lemma hoare_assume_preNF: "(\s. P s \ \P\ f \Q\!) \ \P\ f \Q\!" by (metis validNF_alt_def) lemma bexEI: "\\x\S. Q x; \x. \x \ S; Q x\ \ P x\ \ \x\S. P x" by blast lemma monad_eq_split: assumes "\r s. Q r s \ f r s = f' r s" "\P\ g \\r s. Q r s\" "P s" shows "(g >>= f) s = (g >>= f') s" proof - have pre: "\rv s'. \(rv, s') \ fst (g s)\ \ f rv s' = f' rv s'" using assms unfolding valid_def by (erule_tac x=s in allE) auto show ?thesis apply (simp add: bind_def image_def) apply (intro conjI) apply (rule set_eqI) apply (clarsimp simp: Union_eq) apply (rule iffI; elim exEI conjE; simp; elim exEI bexEI; clarsimp simp: pre) apply (rule iffI; cases "snd (g s)"; simp; elim exEI bexEI; clarsimp simp: pre) done qed lemma monad_eq_split2: assumes eq: " g' s = g s" assumes tail:"\r s. Q r s \ f r s = f' r s" and hoare: "\P\g\\r s. Q r s\" "P s" shows "(g>>=f) s = (g'>>= f') s" proof - have pre: "\aa bb. \(aa, bb) \ fst (g s)\ \ Q aa bb" using hoare by (auto simp: valid_def) show ?thesis apply (simp add:bind_def eq split_def image_def) apply (rule conjI) apply (rule set_eqI) apply (clarsimp simp:Union_eq) apply (metis pre surjective_pairing tail) apply (metis pre surjective_pairing tail) done qed lemma monad_eq_split_tail: "\f = g;a s = b s\ \ (a >>= f) s = ((b >>= g) s)" by (simp add:bind_def) lemma double_gets_drop_regets: "(do x \ gets f; xa \ gets f; m xa x od) = (do xa \ gets f; m xa xa od)" by (simp add: gets_def get_def bind_def return_def) definition monad_commute where "monad_commute P a b \ (\s. (P s \ ((do x\a;y\b; return (x, y) od) s) = ((do y\b;x\a; return (x, y) od) s)))" lemma monad_eq: "a s = b s \ (a >>= g) s = (b >>= g) s" by (auto simp:bind_def) lemma monad_commute_simple: "\monad_commute P a b;P s\ \ ((do x\a;y\b; g x y od) s) = ((do y\b;x\a; g x y od) s)" apply (clarsimp simp:monad_commute_def) apply (drule spec) apply (erule(1) impE) apply (drule_tac g = "(\t. g (fst t) (snd t))" in monad_eq) apply (simp add:bind_assoc) done lemma commute_commute: "monad_commute P f h \ monad_commute P h f" apply (simp (no_asm) add: monad_commute_def) apply (clarsimp) apply (erule monad_commute_simple[symmetric]) apply simp done lemma assert_commute: "monad_commute (K G) (assert G) f" by (clarsimp simp:assert_def monad_commute_def) lemma cond_fail_commute: "monad_commute (K (\G)) (when G fail) f" by (clarsimp simp:when_def fail_def monad_commute_def) lemma return_commute: "monad_commute \ (return a) f" by (clarsimp simp: return_def bind_def monad_commute_def) lemma monad_commute_guard_imp: "\monad_commute P f h; \s. Q s \ P s \ \ monad_commute Q f h" by (clarsimp simp:monad_commute_def) lemma monad_commute_split: "\\r. monad_commute (Q r) f (g r); monad_commute P f h; \P'\ h \\r. Q r\\ \ monad_commute (P and P') f (h>>=g)" apply (simp (no_asm) add:monad_commute_def) apply (clarsimp simp:bind_assoc) apply (subst monad_commute_simple) apply simp+ apply (rule_tac Q = "(\x. Q x)" in monad_eq_split) apply (subst monad_commute_simple[where a = f]) apply assumption apply simp+ done lemma monad_commute_get: assumes hf: "\P. \P\ f \\r. P\" and hg: "\P. \P\ g \\r. P\" and eptyf: "empty_fail f" "empty_fail g" shows "monad_commute \ f g" proof - have fsame: "\a b s. (a,b) \ fst (f s) \ b = s" by (drule use_valid[OF _ hf],auto) have gsame: "\a b s. (a,b) \ fst (g s) \ b = s" by (drule use_valid[OF _ hg],auto) note ef = empty_fail_not_snd[OF _ eptyf(1)] note eg = empty_fail_not_snd[OF _ eptyf(2)] show ?thesis apply (simp add:monad_commute_def) apply (clarsimp simp:bind_def split_def return_def) apply (intro conjI) apply (rule set_eqI) apply (rule iffI) apply (clarsimp simp:Union_eq dest!: singletonD) apply (frule fsame) apply clarsimp apply (frule gsame) apply (metis fst_conv snd_conv) apply (clarsimp simp:Union_eq dest!: singletonD) apply (frule gsame) apply clarsimp apply (frule fsame) apply clarsimp apply (metis fst_conv snd_conv) apply (rule iffI) apply (erule disjE) apply (clarsimp simp:image_def) apply (metis fsame) apply (clarsimp simp:image_def) apply (drule eg) apply clarsimp apply (rule bexI [rotated], assumption) apply (frule gsame) apply clarsimp apply (erule disjE) apply (clarsimp simp:image_def dest!:gsame) apply (clarsimp simp:image_def) apply (drule ef) apply clarsimp apply (frule fsame) apply (erule bexI[rotated]) apply simp done qed lemma mapM_x_commute: assumes commute: "\r. monad_commute (P r) a (b r)" and single: "\r x. \P r and K (f r \ f x) and P x\ b x \\v. P r \" shows "monad_commute (\s. (distinct (map f list)) \ (\r\ set list. P r s)) a (mapM_x b list)" apply (induct list) apply (clarsimp simp:mapM_x_Nil return_def bind_def monad_commute_def) apply (clarsimp simp:mapM_x_Cons) apply (rule monad_commute_guard_imp) apply (rule monad_commute_split) apply assumption apply (rule monad_commute_guard_imp[OF commute]) apply assumption apply (wp hoare_vcg_ball_lift) apply (rule single) apply (clarsimp simp: image_def) apply auto done lemma commute_name_pre_state: assumes "\s. P s \ monad_commute (op = s) f g" shows "monad_commute P f g" using assms by (clarsimp simp:monad_commute_def) lemma commute_rewrite: assumes rewrite: "\s. Q s \ f s = t s" and hold : "\P\ g \\x. Q\" shows "monad_commute R t g \ monad_commute (P and Q and R) f g" apply (clarsimp simp:monad_commute_def bind_def split_def return_def) apply (drule_tac x = s in spec) apply (clarsimp simp:rewrite[symmetric]) apply (intro conjI) apply (rule set_eqI) apply (rule iffI) apply clarsimp apply (rule bexI[rotated],assumption) apply (subst rewrite) apply (rule use_valid[OF _ hold]) apply simp+ apply (erule bexI[rotated],simp) apply clarsimp apply (rule bexI[rotated],assumption) apply (subst rewrite[symmetric]) apply (rule use_valid[OF _ hold]) apply simp+ apply (erule bexI[rotated],simp) apply (intro iffI) apply clarsimp apply (rule bexI[rotated],assumption) apply simp apply (subst rewrite) apply (erule(1) use_valid[OF _ hold]) apply simp apply (clarsimp) apply (drule bspec,assumption) apply clarsimp apply (metis rewrite use_valid[OF _ hold]) done lemma commute_grab_asm: "(F \ monad_commute P f g) \ (monad_commute (P and (K F)) f g)" by (clarsimp simp: monad_commute_def) end