(* * 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 Corres_UL imports Crunch_Instances_NonDet "Monad_WP/wp/WPEx" "Monad_WP/wp/WPFix" HaskellLemmaBucket begin text \Definition of correspondence\ definition corres_underlying :: "(('s \ 't) set) \ bool \ bool \ ('a \ 'b \ bool) \ ('s \ bool) \ ('t \ bool) \ ('s, 'a) nondet_monad \ ('t, 'b) nondet_monad \ bool" where "corres_underlying srel nf nf' rrel G G' \ \m m'. \(s, s') \ srel. G s \ G' s' \ (nf \ \ snd (m s)) \ (\(r', t') \ fst (m' s'). \(r, t) \ fst (m s). (t, t') \ srel \ rrel r r') \ (nf' \ \ snd (m' s'))" text \Base case facts about correspondence\ lemma corres_underlyingD: "\ corres_underlying R nf nf' rs P P' f f'; (s,s') \ R; P s; P' s'; nf \ \ snd (f s) \ \ (\(r',t')\fst (f' s'). \(r,t)\fst (f s). (t, t') \ R \ rs r r') \ (nf' \ \ snd (f' s'))" by (fastforce simp: corres_underlying_def) lemma corres_underlyingD2: "\ corres_underlying R nf nf' rs P P' f f'; (s,s') \ R; P s; P' s'; (r',t')\fst (f' s'); nf \ \ snd (f s) \ \ \(r,t)\fst (f s). (t, t') \ R \ rs r r'" by (fastforce dest: corres_underlyingD) lemma propagate_no_fail: "\ corres_underlying S nf True R P P' f f'; no_fail P f; \s'. P' s' \ (\s. P s \ (s,s') \ S) \ \ no_fail P' f'" apply (clarsimp simp: corres_underlying_def no_fail_def) apply (erule allE, erule (1) impE) apply clarsimp apply (drule (1) bspec, clarsimp) done lemma corres_underlying_serial: "\ corres_underlying S False True rrel G G' m m'; empty_fail m' \ \ \s. (\s'. (s,s') \ S \ G s \ G' s') \ fst (m s) \ {}" apply (clarsimp simp: corres_underlying_def empty_fail_def) apply (drule_tac x="(s, s')" in bspec, simp) apply (drule_tac x=s' in spec) apply auto done (* FIXME: duplicated with HOL.iff_allI *) lemma All_eqI: assumes ass: "\x. A x = B x" shows "(\x. A x) = (\x. B x)" apply (subst ass) apply (rule refl) done lemma corres_singleton: "corres_underlying sr nf nf' r P P' (\s. ({(R s, S s)},x)) (\s. ({(R' s, S' s)},False)) = (\s s'. P s \ P' s' \ (s, s') \ sr \ (nf \ \ x) \ ((S s, S' s') \ sr \ r (R s) (R' s')))" by (auto simp: corres_underlying_def) lemma corres_return[simp]: "corres_underlying sr nf nf' r P P' (return a) (return b) = ((\s s'. P s \ P' s' \ (s, s') \ sr) \ r a b)" by (simp add: return_def corres_singleton) lemma corres_get[simp]: "corres_underlying sr nf nf' r P P' get get = (\ s s'. (s, s') \ sr \ P s \ P' s' \ r s s')" apply (simp add: get_def corres_singleton) apply (rule All_eqI)+ apply safe done lemma corres_gets[simp]: "corres_underlying sr nf nf' r P P' (gets a) (gets b) = (\ s s'. P s \ P' s' \ (s, s') \ sr \ r (a s) (b s'))" by (simp add: simpler_gets_def corres_singleton) lemma corres_throwError[simp]: "corres_underlying sr nf nf' r P P' (throwError a) (throwError b) = ((\s s'. P s \ P' s' \ (s, s') \ sr) \ r (Inl a) (Inl b))" by (simp add: throwError_def) lemma corres_no_failI_base: assumes f: "nf \ no_fail P f" assumes f': "nf' \ no_fail P' f'" assumes corres: "\(s, s') \ S. P s \ P' s' \ (\(r', t') \ fst (f' s'). \(r, t) \ fst (f s). (t, t') \ S \ R r r')" shows "corres_underlying S nf nf' R P P' f f'" using assms by (simp add: corres_underlying_def no_fail_def) (* This lemma gets the shorter name because many existing proofs want nf=False *) lemma corres_no_failI: assumes f': "nf' \ no_fail P' f'" assumes corres: "\(s, s') \ S. P s \ P' s' \ (\(r', t') \ fst (f' s'). \(r, t) \ fst (f s). (t, t') \ S \ R r r')" shows "corres_underlying S False nf' R P P' f f'" using assms by (simp add: corres_underlying_def no_fail_def) text \A congruence rule for the correspondence functions.\ lemma corres_cong: assumes P: "\s. P s = P' s" assumes Q: "\s. Q s = Q' s" assumes f: "\s. P' s \ f s = f' s" assumes g: "\s. Q' s \ g s = g' s" assumes r: "\x y s t s' t'. \ P' s; Q' t; (x, s') \ fst (f' s); (y, t') \ fst (g' t) \ \ r x y = r' x y" shows "corres_underlying sr nf nf' r P Q f g = corres_underlying sr nf nf' r' P' Q' f' g'" apply (simp add: corres_underlying_def) apply (rule ball_cong [OF refl]) apply (clarsimp simp: P Q) apply (rule imp_cong [OF refl]) apply (clarsimp simp: f g) apply (rule imp_cong [OF refl]) apply (rule conj_cong) apply (rule ball_cong [OF refl]) apply clarsimp apply (rule bex_cong [OF refl]) apply (clarsimp simp: r) apply simp done text \The guard weakening rule\ lemma stronger_corres_guard_imp: assumes x: "corres_underlying sr nf nf' r Q Q' f g" assumes y: "\s s'. \ P s; P' s'; (s, s') \ sr \ \ Q s" assumes z: "\s s'. \ P s; P' s'; (s, s') \ sr \ \ Q' s'" shows "corres_underlying sr nf nf' r P P' f g" using x by (auto simp: y z corres_underlying_def) lemma corres_guard_imp: assumes x: "corres_underlying sr nf nf' r Q Q' f g" assumes y: "\s. P s \ Q s" "\s. P' s \ Q' s" shows "corres_underlying sr nf nf' r P P' f g" apply (rule stronger_corres_guard_imp) apply (rule x) apply (simp add: y)+ done lemma corres_rel_imp: assumes x: "corres_underlying sr nf nf' r' P P' f g" assumes y: "\x y. r' x y \ r x y" shows "corres_underlying sr nf nf' r P P' f g" apply (insert x) apply (simp add: corres_underlying_def) apply clarsimp apply (drule (1) bspec, clarsimp) apply (drule (1) bspec, clarsimp) apply (blast intro: y) done text \Splitting rules for correspondence of composite monads\ lemma corres_underlying_split: assumes ac: "corres_underlying s nf nf' r' G G' a c" assumes valid: "\G\ a \P\" "\G'\ c \P'\" assumes bd: "\rv rv'. r' rv rv' \ corres_underlying s nf nf' r (P rv) (P' rv') (b rv) (d rv')" shows "corres_underlying s nf nf' r G G' (a >>= (\rv. b rv)) (c >>= (\rv'. d rv'))" using ac bd valid apply (clarsimp simp: corres_underlying_def bind_def) apply (clarsimp simp: Bex_def Ball_def valid_def) apply meson done lemma corres_split': assumes x: "corres_underlying sr nf nf' r' P P' a c" assumes y: "\rv rv'. r' rv rv' \ corres_underlying sr nf nf' r (Q rv) (Q' rv') (b rv) (d rv')" assumes "\P\ a \Q\" "\P'\ c \Q'\" shows "corres_underlying sr nf nf' r P P' (a >>= (\rv. b rv)) (c >>= (\rv'. d rv'))" by (fastforce intro!: corres_underlying_split assms) text \Derivative splitting rules\ lemma corres_split: assumes y: "\rv rv'. r' rv rv' \ corres_underlying sr nf nf' r (R rv) (R' rv') (b rv) (d rv')" assumes x: "corres_underlying sr nf nf' r' P P' a c" assumes "\Q\ a \R\" "\Q'\ c \R'\" shows "corres_underlying sr nf nf' r (P and Q) (P' and Q') (a >>= (\rv. b rv)) (c >>= (\rv'. d rv'))" using assms apply - apply (rule corres_split') apply (rule corres_guard_imp, rule x, simp_all) apply (erule y) apply (rule hoare_weaken_pre, assumption) apply simp apply (rule hoare_weaken_pre, assumption) apply simp done primrec rel_sum_comb :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a + 'c \ 'b + 'd \ bool)" (infixl "\" 95) where "(f \ g) (Inr x) y = (\y'. y = Inr y' \ (g x y'))" | "(f \ g) (Inl x) y = (\y'. y = Inl y' \ (f x y'))" lemma rel_sum_comb_r2[simp]: "(f \ g) x (Inr y) = (\x'. x = Inr x' \ g x' y)" apply (case_tac x, simp_all) done lemma rel_sum_comb_l2[simp]: "(f \ g) x (Inl y) = (\x'. x = Inl x' \ f x' y)" apply (case_tac x, simp_all) done lemma corres_splitEE: assumes y: "\rv rv'. r' rv rv' \ corres_underlying sr nf nf' (f \ r) (R rv) (R' rv') (b rv) (d rv')" assumes "corres_underlying sr nf nf' (f \ r') P P' a c" assumes x: "\Q\ a \R\,\\\\" "\Q'\ c \R'\,\\\\" shows "corres_underlying sr nf nf' (f \ r) (P and Q) (P' and Q') (a >>=E (\rv. b rv)) (c >>=E (\rv'. d rv'))" using assms apply (unfold bindE_def validE_def) apply (rule corres_split) defer apply assumption+ apply (case_tac rv) apply (clarsimp simp: lift_def y)+ done lemma corres_split_handle: assumes y: "\ft ft'. f' ft ft' \ corres_underlying sr nf nf' (f \ r) (E ft) (E' ft') (b ft) (d ft')" assumes "corres_underlying sr nf nf' (f' \ r) P P' a c" assumes x: "\Q\ a \\\\,\E\" "\Q'\ c \\\\,\E'\" shows "corres_underlying sr nf nf' (f \ r) (P and Q) (P' and Q') (a (\ft. b ft)) (c (\ft'. d ft'))" using assms apply (simp add: handleE_def handleE'_def validE_def) apply (rule corres_split) defer apply assumption+ apply (case_tac v, simp_all, safe, simp_all add: y) done lemma corres_split_catch: assumes y: "\ft ft'. f ft ft' \ corres_underlying sr nf nf' r (E ft) (E' ft') (b ft) (d ft')" assumes x: "corres_underlying sr nf nf' (f \ r) P P' a c" assumes z: "\Q\ a \\\\,\E\" "\Q'\ c \\\\,\E'\" shows "corres_underlying sr nf nf' r (P and Q) (P' and Q') (a (\ft. b ft)) (c (\ft'. d ft'))" apply (simp add: catch_def) apply (rule corres_split [OF _ x, where R="case_sum E \\" and R'="case_sum E' \\"]) apply (case_tac x) apply (clarsimp simp: y) apply clarsimp apply (insert z) apply (simp add: validE_def valid_def split_def split: sum.splits)+ done lemma corres_split_eqr: assumes y: "\rv. corres_underlying sr nf nf' r (R rv) (R' rv) (b rv) (d rv)" assumes x: "corres_underlying sr nf nf' (=) P P' a c" "\Q\ a \R\" "\Q'\ c \R'\" shows "corres_underlying sr nf nf' r (P and Q) (P' and Q') (a >>= (\rv. b rv)) (c >>= d)" apply (rule corres_split[OF _ x]) apply (simp add: y) done definition "dc \ \rv rv'. True" lemma dc_simp[simp]: "dc a b" by (simp add: dc_def) lemma dc_o_simp1[simp]: "dc \ f = dc" by (simp add: dc_def o_def) lemma dc_o_simp2[simp]: "dc x \ f = dc x" by (simp add: dc_def o_def) lemma unit_dc_is_eq: "(dc::unit\_\_) = (=)" by (fastforce simp: dc_def) lemma corres_split_nor: "\ corres_underlying sr nf nf' r R R' b d; corres_underlying sr nf nf' dc P P' a c; \Q\ a \\x. R\; \Q'\ c \\x. R'\ \ \ corres_underlying sr nf nf' r (P and Q) (P' and Q') (a >>= (\rv. b)) (c >>= (\rv. d))" apply (rule corres_split, assumption+) done lemma corres_split_norE: "\ corres_underlying sr nf nf' (f \ r) R R' b d; corres_underlying sr nf nf' (f \ dc) P P' a c; \Q\ a \\x. R\, \\\\; \Q'\ c \\x. R'\,\\\\ \ \ corres_underlying sr nf nf' (f \ r) (P and Q) (P' and Q') (a >>=E (\rv. b)) (c >>=E (\rv. d))" apply (rule corres_splitEE, assumption+) done lemma corres_split_eqrE: assumes y: "\rv. corres_underlying sr nf nf' (f \ r) (R rv) (R' rv) (b rv) (d rv)" assumes z: "corres_underlying sr nf nf' (f \ (=)) P P' a c" assumes x: "\Q\ a \R\,\\\\" "\Q'\ c \R'\,\\\\" shows "corres_underlying sr nf nf' (f \ r) (P and Q) (P' and Q') (a >>=E (\rv. b rv)) (c >>=E d)" apply (rule corres_splitEE[OF _ z x]) apply (simp add: y) done lemma corres_split_mapr: assumes x: "\rv. corres_underlying sr nf nf' r (R rv) (R' (f rv)) (b rv) (d (f rv))" assumes y: "corres_underlying sr nf nf' ((=) \ f) P P' a c" assumes z: "\Q\ a \R\" "\Q'\ c \R'\" shows "corres_underlying sr nf nf' r (P and Q) (P' and Q') (a >>= (\rv. b rv)) (c >>= d)" apply (rule corres_split[OF _ y z]) apply simp apply (drule sym) apply (simp add: x) done lemma corres_split_maprE: assumes y: "\rv. corres_underlying sr nf nf' (r' \ r) (R rv) (R' (f rv)) (b rv) (d (f rv))" assumes z: "corres_underlying sr nf nf' (r' \ ((=) \ f)) P P' a c" assumes x: "\Q\ a \R\,\\\\" "\Q'\ c \R'\,\\\\" shows "corres_underlying sr nf nf' (r' \ r) (P and Q) (P' and Q') (a >>=E (\rv. b rv)) (c >>=E d)" apply (rule corres_splitEE[OF _ z x]) apply simp apply (drule sym) apply (simp add: y) done text \Some rules for walking correspondence into basic constructs\ lemma corres_if: "\ G = G'; corres_underlying sr nf nf' r P P' a c; corres_underlying sr nf nf' r Q Q' b d \ \ corres_underlying sr nf nf' r (if G then P else Q) (if G' then P' else Q') (if G then a else b) (if G' then c else d)" by simp lemma corres_whenE: "\ G = G'; corres_underlying sr nf nf' (fr \ r) P P' f g; r () () \ \ corres_underlying sr nf nf' (fr \ r) (\s. G \ P s) (\s. G' \ P' s) (whenE G f) (whenE G' g)" by (simp add: whenE_def returnOk_def) lemmas corres_if2 = corres_if[unfolded if_apply_def2] lemmas corres_when = corres_if2[where b="return ()" and d="return ()" and Q="\" and Q'="\" and r=dc, simplified, folded when_def] lemma corres_if_r: "\ corres_underlying sr nf nf' r P P' a c; corres_underlying sr nf nf' r P Q' a d \ \ corres_underlying sr nf nf' r (P) (if G' then P' else Q') (a) (if G' then c else d)" by (simp) lemma corres_if3: "\ G = G'; G \ corres_underlying sr nf nf' r P P' a c; \ G' \ corres_underlying sr nf nf' r Q Q' b d \ \ corres_underlying sr nf nf' r (if G then P else Q) (if G' then P' else Q') (if G then a else b) (if G' then c else d)" by simp text \Some equivalences about liftM and other useful simps\ lemma snd_liftM [simp]: "snd (liftM t f s) = snd (f s)" by (auto simp: liftM_def bind_def return_def) lemma corres_liftM_simp[simp]: "(corres_underlying sr nf nf' r P P' (liftM t f) g) = (corres_underlying sr nf nf' (r \ t) P P' f g)" apply (simp add: corres_underlying_def handy_liftM_lemma Ball_def Bex_def) apply (rule All_eqI)+ apply blast done lemma corres_liftM2_simp[simp]: "corres_underlying sr nf nf' r P P' f (liftM t g) = corres_underlying sr nf nf' (\x. r x \ t) P P' f g" apply (simp add: corres_underlying_def handy_liftM_lemma Ball_def) apply (rule All_eqI)+ apply blast done lemma corres_liftE_rel_sum[simp]: "corres_underlying sr nf nf' (f \ r) P P' (liftE m) (liftE m') = corres_underlying sr nf nf' r P P' m m'" by (simp add: liftE_liftM o_def) text \Support for proving correspondence to noop with hoare triples\ lemma corres_noop: assumes P: "\s. P s \ \\s'. (s, s') \ sr \ P' s'\ f \\rv s'. (s, s') \ sr \ r x rv\" assumes nf': "\s. \ P s; nf' \ \ no_fail (\s'. (s, s') \ sr \ P' s') f" shows "corres_underlying sr nf nf' r P P' (return x) f" apply (simp add: corres_underlying_def return_def) apply clarsimp apply (frule P) apply (insert nf') apply (clarsimp simp: valid_def no_fail_def) done lemma corres_noopE: assumes P: "\s. P s \ \\s'. (s, s') \ sr \ P' s'\ f \\rv s'. (s, s') \ sr \ r x rv\,\\r s. False\" assumes nf': "\s. \ P s; nf' \ \ no_fail (\s'. (s, s') \ sr \ P' s') f" shows "corres_underlying sr nf nf' (fr \ r) P P' (returnOk x) f" proof - have Q: "\P f Q E. \P\f\Q\,\E\ \ \P\ f \\r s. case_sum (\e. E e s) (\r. Q r s) r\" by (simp add: validE_def) thus ?thesis apply (simp add: returnOk_def) apply (rule corres_noop) apply (rule hoare_post_imp) defer apply (rule Q) apply (rule P) apply assumption apply (erule(1) nf') apply (case_tac ra, simp_all) done qed (* this could be stronger in the no_fail part *) lemma corres_noop2: assumes x: "\s. P s \ \(=) s\ f \\\r. (=) s\" assumes y: "\s. P' s \ \(=) s\ g \\r. (=) s\" assumes z: "nf' \ no_fail P f" "nf' \ no_fail P' g" shows "corres_underlying sr nf nf' dc P P' f g" apply (clarsimp simp: corres_underlying_def) apply (rule conjI) apply clarsimp apply (rule use_exs_valid) apply (rule exs_hoare_post_imp) prefer 2 apply (rule x) apply assumption apply simp_all apply (subgoal_tac "ba = b") apply simp apply (rule sym) apply (rule use_valid[OF _ y], assumption+) apply simp apply (insert z) apply (clarsimp simp: no_fail_def) done text \Support for dividing correspondence along logical boundaries\ lemma corres_disj_division: "\ P \ Q; P \ corres_underlying sr nf nf' r R S x y; Q \ corres_underlying sr nf nf' r T U x y \ \ corres_underlying sr nf nf' r (\s. (P \ R s) \ (Q \ T s)) (\s. (P \ S s) \ (Q \ U s)) x y" apply safe apply (rule corres_guard_imp) apply simp apply simp apply simp apply (rule corres_guard_imp) apply simp apply simp apply simp done lemma corres_weaker_disj_division: "\ P \ Q; P \ corres_underlying sr nf nf' r R S x y; Q \ corres_underlying sr nf nf' r T U x y \ \ corres_underlying sr nf nf' r (R and T) (S and U) x y" apply (rule corres_guard_imp) apply (rule corres_disj_division) apply simp+ done lemma corres_symmetric_bool_cases: "\ P = P'; \ P; P' \ \ corres_underlying srel nf nf' r Q Q' f g; \ \ P; \ P' \ \ corres_underlying srel nf nf' r R R' f g \ \ corres_underlying srel nf nf' r (\s. (P \ Q s) \ (\ P \ R s)) (\s. (P' \ Q' s) \ (\ P' \ R' s)) f g" by (cases P, simp_all) text \Support for symbolically executing into the guards and manipulating them\ lemma corres_symb_exec_l: assumes z: "\rv. corres_underlying sr nf nf' r (Q rv) P' (x rv) y" assumes x: "\s. P s \ \(=) s\ m \\\r. (=) s\" assumes y: "\P\ m \Q\" assumes nf: "nf' \ no_fail P m" shows "corres_underlying sr nf nf' r P P' (m >>= (\rv. x rv)) y" apply (rule corres_guard_imp) apply (subst gets_bind_ign[symmetric], rule corres_split) apply (rule z) apply (rule corres_noop2) apply (erule x) apply (rule gets_wp) apply (erule nf) apply (rule non_fail_gets) apply (rule y) apply (rule gets_wp) apply simp+ done lemma corres_symb_exec_r: assumes z: "\rv. corres_underlying sr nf nf' r P (Q' rv) x (y rv)" assumes y: "\P'\ m \Q'\" assumes x: "\s. P' s \ \(=) s\ m \\r. (=) s\" assumes nf: "nf' \ no_fail P' m" shows "corres_underlying sr nf nf' r P P' x (m >>= (\rv. y rv))" apply (rule corres_guard_imp) apply (subst gets_bind_ign[symmetric], rule corres_split) apply (rule z) apply (rule corres_noop2) apply (simp add: simpler_gets_def exs_valid_def) apply (erule x) apply (rule non_fail_gets) apply (erule nf) apply (rule gets_wp) apply (rule y) apply simp+ done lemma corres_symb_exec_r_conj: assumes z: "\rv. corres_underlying sr nf nf' r Q (R' rv) x (y rv)" assumes y: "\Q'\ m \R'\" assumes x: "\s. \\s'. (s, s') \ sr \ P' s'\ m \\rv s'. (s, s') \ sr\" assumes nf: "\s. nf' \ no_fail (\s'. (s, s') \ sr \ P' s') m" shows "corres_underlying sr nf nf' r Q (P' and Q') x (m >>= (\rv. y rv))" proof - have P: "corres_underlying sr nf nf' dc \ P' (return undefined) m" apply (rule corres_noop) apply (simp add: x) apply (erule nf) done show ?thesis apply (rule corres_guard_imp) apply (subst return_bind[symmetric], rule corres_split [OF _ P]) apply (rule z) apply wp apply (rule y) apply simp+ done qed lemma corres_bind_return_r: "corres_underlying S nf nf' (\x y. r x (h y)) P Q f g \ corres_underlying S nf nf' r P Q f (do x \ g; return (h x) od)" by (fastforce simp: corres_underlying_def bind_def return_def) lemma corres_underlying_symb_exec_l: "\ corres_underlying sr nf nf' dc P P' f (return ()); \rv. corres_underlying sr nf nf' r (Q rv) P' (g rv) h; \P\ f \Q\ \ \ corres_underlying sr nf nf' r P P' (f >>= g) h" apply (drule(1) corres_underlying_split) apply (rule return_wp) apply clarsimp apply (erule meta_allE, assumption) apply simp done text \Inserting assumptions to be proved later\ lemma corres_req: assumes x: "\s s'. \ (s, s') \ sr; P s; P' s' \ \ F" assumes y: "F \ corres_underlying sr nf nf' r P P' f g" shows "corres_underlying sr nf nf' r P P' f g" apply (cases "F") apply (rule y) apply assumption apply (simp add: corres_underlying_def) apply clarsimp apply (subgoal_tac "F") apply simp apply (rule x, assumption+) done (* Insert assumption to be proved later, on the left-hand (abstract) side *) lemma corres_gen_asm: assumes x: "F \ corres_underlying sr nf nf' r P P' f g" shows "corres_underlying sr nf nf' r (P and (\s. F)) P' f g" apply (rule corres_req[where F=F]) apply simp apply (rule corres_guard_imp [OF x]) apply simp+ done (* Insert assumption to be proved later, on the right-hand (concrete) side *) lemma corres_gen_asm2: assumes x: "F \ corres_underlying sr nf nf' r P P' f g" shows "corres_underlying sr nf nf' r P (P' and (\s. F)) f g" apply (rule corres_req[where F=F]) apply simp apply (rule corres_guard_imp [OF x]) apply simp+ done lemma corres_trivial: "corres_underlying sr nf nf' r \ \ f g \ corres_underlying sr nf nf' r \ \ f g" by assumption lemma corres_assume_pre: assumes R: "\s s'. \ P s; Q s'; (s,s') \ sr \ \ corres_underlying sr nf nf' r P Q f g" shows "corres_underlying sr nf nf' r P Q f g" apply (clarsimp simp add: corres_underlying_def) apply (frule (2) R) apply (clarsimp simp add: corres_underlying_def) apply blast done lemma corres_guard_imp2: "\corres_underlying sr nf nf' r Q P' f g; \s. P s \ Q s\ \ corres_underlying sr nf nf' r P P' f g" by (blast intro: corres_guard_imp) (* FIXME: names\ (cf. corres_guard2_imp below) *) lemmas corres_guard1_imp = corres_guard_imp2 lemma corres_guard2_imp: "\corres_underlying sr nf nf' r P Q' f g; \s. P' s \ Q' s\ \ corres_underlying sr nf nf' r P P' f g" by (drule (1) corres_guard_imp[where P'=P' and Q=P], assumption+) lemma corres_initial_splitE: "\ corres_underlying sr nf nf' (f \ r') P P' a c; \rv rv'. r' rv rv' \ corres_underlying sr nf nf' (f \ r) (Q rv) (Q' rv') (b rv) (d rv'); \P\ a \Q\, \\r s. True\; \P'\ c \Q'\, \\r s. True\\ \ corres_underlying sr nf nf' (f \ r) P P' (a >>=E b) (c >>=E d)" apply (rule corres_guard_imp) apply (erule (3) corres_splitEE) apply simp apply simp done lemma corres_assert_assume: "\ P' \ corres_underlying sr nf nf' r P Q f (g ()); \s. Q s \ P' \ \ corres_underlying sr nf nf' r P Q f (assert P' >>= g)" by (auto simp: bind_def assert_def fail_def return_def corres_underlying_def) lemma corres_state_assert: "corres_underlying sr nf nf' rr P Q f (g ()) \ (\s. Q s \ R s) \ corres_underlying sr nf nf' rr P Q f (state_assert R >>= g)" by (clarsimp simp: corres_underlying_def state_assert_def get_def assert_def return_def bind_def) lemma corres_stateAssert_assume: "\ corres_underlying sr nf nf' r P Q f (g ()); \s. Q s \ P' s \ \ corres_underlying sr nf nf' r P Q f (stateAssert P' [] >>= g)" apply (clarsimp simp: bind_assoc stateAssert_def) apply (rule corres_symb_exec_r [OF _ get_sp]) apply (rule corres_assert_assume) apply (rule corres_assume_pre) apply (erule corres_guard_imp, clarsimp+) apply (wp | rule no_fail_pre)+ done lemma corres_stateAssert_implied: "\ corres_underlying sr nf nf' r P Q f (g ()); \s s'. \ (s, s') \ sr; P s; P' s; Q s' \ \ Q' s' \ \ corres_underlying sr nf nf' r (P and P') Q f (stateAssert Q' [] >>= g)" apply (clarsimp simp: bind_assoc stateAssert_def) apply (rule corres_symb_exec_r [OF _ get_sp]) apply (rule corres_assume_pre) apply (rule corres_assert_assume) apply (erule corres_guard_imp, clarsimp+) apply (wp | rule no_fail_pre)+ done lemma corres_assert: "corres_underlying sr nf nf' dc (%_. P) (%_. Q) (assert P) (assert Q)" by (clarsimp simp add: corres_underlying_def return_def) lemma corres_split2: assumes corr: "\a a' b b'. \ r a a' b b'\ \ corres_underlying sr nf nf' r1 (P1 a b) (P1' a' b') (H a b) (H' a' b')" and corr': "corres_underlying sr nf nf' (\(a, b).\(a', b'). r a a' b b') P P' (do a \ F; b \ G; return (a, b) od) (do a' \ F'; b' \ G'; return (a', b') od)" and h1: "\P\ do fx \ F; gx \ G; return (fx, gx) od \\rv. P1 (fst rv) (snd rv)\" and h2: "\P'\ do fx \ F'; gx \ G'; return (fx, gx) od \\rv. P1' (fst rv) (snd rv)\" shows "corres_underlying sr nf nf' r1 P P' (do a \ F; b \ G; H a b od) (do a' \ F'; b' \ G'; H' a' b' od)" proof - have "corres_underlying sr nf nf' r1 P P' (do a \ F; b \ G; rv \ return (a, b); H (fst rv) (snd rv) od) (do a' \ F'; b' \ G'; rv' \ return (a', b'); H' (fst rv') (snd rv') od)" by (rule corres_split' [OF corr' corr, simplified bind_assoc, OF _ h1 h2]) (simp add: split_beta split_def) thus ?thesis by simp qed lemma corres_split3: assumes corr: "\a a' b b' c c'. \ r a a' b b' c c'\ \ corres_underlying sr nf nf' r1 (P1 a b c) (P1' a' b' c') (H a b c) (H' a' b' c')" and corr': "corres_underlying sr nf nf' (\(a, b, c).\(a', b', c'). r a a' b b' c c') P P' (do a \ A; b \ B a; c \ C a b; return (a, b, c) od) (do a' \ A'; b' \ B' a'; c' \ C' a' b'; return (a', b', c') od)" and h1: "\P\ do a \ A; b \ B a; c \ C a b; return (a, b, c) od \\(a, b, c). P1 a b c\" and h2: "\P'\ do a' \ A'; b' \ B' a'; c' \ C' a' b'; return (a', b', c') od \\(a', b', c'). P1' a' b' c'\" shows "corres_underlying sr nf nf' r1 P P' (do a \ A; b \ B a; c \ C a b; H a b c od) (do a' \ A'; b' \ B' a'; c' \ C' a' b'; H' a' b' c' od)" proof - have "corres_underlying sr nf nf' r1 P P' (do a \ A; b \ B a; c \ C a b; rv \ return (a, b, c); H (fst rv) (fst (snd rv)) (snd (snd rv)) od) (do a' \ A'; b' \ B' a'; c' \ C' a' b'; rv \ return (a', b', c'); H' (fst rv) (fst (snd rv)) (snd (snd rv)) od)" using h1 h2 by - (rule corres_split' [OF corr' corr, simplified bind_assoc ], simp_all add: split_beta split_def) thus ?thesis by simp qed (* A little broken --- see above *) lemma corres_split4: assumes corr: "\a a' b b' c c' d d'. \ r a a' b b' c c' d d'\ \ corres_underlying sr nf nf' r1 (P1 a b c d) (P1' a' b' c' d') (H a b c d) (H' a' b' c' d')" and corr': "corres_underlying sr nf nf' (\(a, b, c, d).\(a', b', c', d'). r a a' b b' c c' d d') P P' (do a \ A; b \ B; c \ C; d \ D; return (a, b, c, d) od) (do a' \ A'; b' \ B'; c' \ C'; d' \ D'; return (a', b', c', d') od)" and h1: "\P\ do a \ A; b \ B; c \ C; d \ D; return (a, b, c, d) od \\(a, b, c, d). P1 a b c d\" and h2: "\P'\ do a' \ A'; b' \ B'; c' \ C'; d' \ D'; return (a', b', c', d') od \\(a', b', c', d'). P1' a' b' c' d'\" shows "corres_underlying sr nf nf' r1 P P' (do a \ A; b \ B; c \ C; d \ D; H a b c d od) (do a' \ A'; b' \ B'; c' \ C'; d' \ D'; H' a' b' c' d' od)" proof - have "corres_underlying sr nf nf' r1 P P' (do a \ A; b \ B; c \ C; d \ D; rv \ return (a, b, c, d); H (fst rv) (fst (snd rv)) (fst (snd (snd rv))) (snd (snd (snd rv))) od) (do a' \ A'; b' \ B'; c' \ C'; d' \ D'; rv \ return (a', b', c', d'); H' (fst rv) (fst (snd rv)) (fst (snd (snd rv))) (snd (snd (snd rv))) od)" using h1 h2 by - (rule corres_split' [OF corr' corr, simplified bind_assoc], simp_all add: split_beta split_def) thus ?thesis by simp qed (* for instantiations *) lemma corres_inst: "corres_underlying sr nf nf' r P P' f g \ corres_underlying sr nf nf' r P P' f g" . lemma corres_assert_opt_assume: assumes "\x. P' = Some x \ corres_underlying sr nf nf' r P Q f (g x)" assumes "\s. Q s \ P' \ None" shows "corres_underlying sr nf nf' r P Q f (assert_opt P' >>= g)" using assms by (auto simp: bind_def assert_opt_def assert_def fail_def return_def corres_underlying_def split: option.splits) text \Support for proving correspondance by decomposing the state relation\ lemma corres_underlying_decomposition: assumes x: "corres_underlying {(s, s'). P s s'} nf nf' r Pr Pr' f g" and y: "\s'. \R s'\ f \\rv s. Q s s'\" and z: "\s. \P s and Q s and K (Pr s) and Pr'\ g \\rv s'. R s' s\" shows "corres_underlying {(s, s'). P s s' \ Q s s'} nf nf' r Pr Pr' f g" using x apply (clarsimp simp: corres_underlying_def) apply (elim allE, drule(1) mp, clarsimp) apply (drule(1) bspec) apply clarsimp apply (rule rev_bexI, assumption) apply simp apply (erule use_valid [OF _ y]) apply (erule use_valid [OF _ z]) apply simp done lemma corres_stronger_no_failI: assumes f': "nf' \ no_fail (\s'. \s. P s \ (s,s') \ S \ P' s') f'" assumes corres: "\(s, s') \ S. P s \ P' s' \ (\(r', t') \ fst (f' s'). \(r, t) \ fst (f s). (t, t') \ S \ R r r')" shows "corres_underlying S nf nf' R P P' f f'" using assms apply (simp add: corres_underlying_def no_fail_def) apply clarsimp apply (rule conjI) apply clarsimp apply blast apply clarsimp apply blast done lemma corres_fail: assumes no_fail: "\s s'. \ (s,s') \ sr; P s; P' s'; nf' \ \ False" shows "corres_underlying sr nf nf' R P P' f fail" using no_fail by (auto simp add: corres_underlying_def fail_def) lemma corres_returnOk: "(\s s'. \ (s,s') \ sr; P s; P' s' \ \ r x y) \ corres_underlying sr nf nf' (r' \ r) P P' (returnOk x) (returnOk y)" apply (rule corres_noopE) apply wp apply clarsimp apply wp done lemmas corres_returnOkTT = corres_trivial [OF corres_returnOk] lemma corres_False [simp]: "corres_underlying sr nf nf' r P \ f f'" by (simp add: corres_underlying_def) lemma corres_liftME[simp]: "corres_underlying sr nf nf' (f \ r) P P' (liftME fn m) m' = corres_underlying sr nf nf' (f \ (r \ fn)) P P' m m'" apply (simp add: liftME_liftM) apply (rule corres_cong [OF refl refl refl refl]) apply (case_tac x, simp_all) done lemma corres_liftME2[simp]: "corres_underlying sr nf nf' (f \ r) P P' m (liftME fn m') = corres_underlying sr nf nf' (f \ (\x. r x \ fn)) P P' m m'" apply (simp add: liftME_liftM) apply (rule corres_cong [OF refl refl refl refl]) apply (case_tac y, simp_all) done lemma corres_assertE_assume: "\\s. P s \ P'; \s'. Q s' \ Q'\ \ corres_underlying sr nf nf' (f \ (=)) P Q (assertE P') (assertE Q')" apply (simp add: corres_underlying_def assertE_def returnOk_def fail_def return_def) by blast definition rel_prod :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c \ 'b \ 'd \ bool)" (infix "\" 97) where "rel_prod \ \f g (a,b) (c,d). f a c \ g b d" lemma rel_prod_apply [simp]: "(f \ g) (a,b) (c,d) = (f a c \ g b d)" by (simp add: rel_prod_def) lemma mapME_x_corres_inv: assumes x: "\x. corres_underlying sr nf nf' (f \ dc) (P x) (P' x) (m x) (m' x)" assumes y: "\x P. \P\ m x \\x. P\,-" "\x P'. \P'\ m' x \\x. P'\,-" assumes z: "xs = ys" shows "corres_underlying sr nf nf' (f \ dc) (\s. \x \ set xs. P x s) (\s. \y \ set ys. P' y s) (mapME_x m xs) (mapME_x m' ys)" unfolding z proof (induct ys) case Nil show ?case by (simp add: mapME_x_def sequenceE_x_def returnOk_def) next case (Cons z zs) from Cons have IH: "corres_underlying sr nf nf' (f \ dc) (\s. \x\set zs. P x s) (\s. \y\set zs. P' y s) (mapME_x m zs) (mapME_x m' zs)" . show ?case apply (simp add: mapME_x_def sequenceE_x_def) apply (fold mapME_x_def sequenceE_x_def dc_def) apply (rule corres_guard_imp) apply (rule corres_splitEE) apply (rule IH) apply (rule x) apply (fold validE_R_def) apply (rule y)+ apply simp+ done qed lemma select_corres_eq: "corres_underlying sr nf nf' (=) \ \ (select UNIV) (select UNIV)" by (simp add: corres_underlying_def select_def) lemma corres_cases: "\ R \ corres_underlying sr nf nf' r P P' f g; \R \ corres_underlying sr nf nf' r Q Q' f g \ \ corres_underlying sr nf nf' r (P and Q) (P' and Q') f g" by (simp add: corres_underlying_def) blast lemma corres_alternate1: "corres_underlying sr nf nf' r P P' a c \ corres_underlying sr nf nf' r P P' (a OR b) c" apply (simp add: corres_underlying_def alternative_def) apply clarsimp apply (drule (1) bspec, clarsimp)+ apply (rule rev_bexI) apply (rule UnI1) apply assumption apply simp done lemma corres_alternate2: "corres_underlying sr nf nf' r P P' b c \ corres_underlying sr nf nf' r P P' (a OR b) c" apply (simp add: corres_underlying_def alternative_def) apply clarsimp apply (drule (1) bspec, clarsimp)+ apply (rule rev_bexI) apply (rule UnI2) apply assumption apply simp done lemma corres_False': "corres_underlying sr nf nf' r \ P' f g" by (simp add: corres_underlying_def) lemma corres_symb_exec_l_Ex: assumes x: "\rv. corres_underlying sr nf nf' r (Q rv) P' (g rv) h" shows "corres_underlying sr nf nf' r (\s. \rv. Q rv s \ (rv, s) \ fst (f s)) P' (do rv \ f; g rv od) h" apply (clarsimp simp add: corres_underlying_def) apply (cut_tac rv=rv in x) apply (clarsimp simp add: corres_underlying_def) apply (drule(1) bspec, clarsimp) apply (case_tac nf) apply (clarsimp simp: bind_def') apply blast apply clarsimp apply (drule(1) bspec, clarsimp) apply (clarsimp simp: bind_def | erule rev_bexI)+ done lemma corres_symb_exec_r_All: assumes nf: "\rv. nf' \ no_fail (Q' rv) g" assumes x: "\rv. corres_underlying sr nf nf' r P (Q' rv) f (h rv)" shows "corres_underlying sr nf nf' r P (\s. (\p \ fst (g s). snd p = s \ Q' (fst p) s) \ (\rv. Q' rv s)) f (do rv \ g; h rv od)" apply (clarsimp simp add: corres_underlying_def bind_def) apply (rule conjI) apply clarsimp apply (cut_tac rv=aa in x) apply (clarsimp simp add: corres_underlying_def bind_def) apply (drule(1) bspec, clarsimp)+ apply (insert nf) apply (clarsimp simp: no_fail_def) apply (erule (1) my_BallE) apply (cut_tac rv="aa" in x) apply (clarsimp simp add: corres_underlying_def bind_def) apply (drule(1) bspec, clarsimp)+ done lemma corres_split_bind_case_sum: assumes x: "corres_underlying sr nf nf' (lr \ rr) P P' a d" assumes y: "\rv rv'. lr rv rv' \ corres_underlying sr nf nf' r (R rv) (R' rv') (b rv) (e rv')" assumes z: "\rv rv'. rr rv rv' \ corres_underlying sr nf nf' r (S rv) (S' rv') (c rv) (f rv')" assumes w: "\Q\ a \S\,\R\" "\Q'\ d \S'\,\R'\" shows "corres_underlying sr nf nf' r (P and Q) (P' and Q') (a >>= (\rv. case_sum b c rv)) (d >>= (\rv'. case_sum e f rv'))" apply (rule corres_split [OF _ x]) defer apply (insert w)[2] apply (simp add: validE_def)+ apply (case_tac rv) apply (clarsimp simp: y) apply (clarsimp simp: z) done lemma whenE_throwError_corres_initial: assumes P: "frel f f'" assumes Q: "P = P'" assumes R: "\ P \ corres_underlying sr nf nf' (frel \ rvr) Q Q' m m'" shows "corres_underlying sr nf nf' (frel \ rvr) Q Q' (whenE P (throwError f ) >>=E (\_. m )) (whenE P' (throwError f') >>=E (\_. m'))" unfolding whenE_def apply (cases P) apply (simp add: P Q) apply (simp add: Q) apply (rule R) apply (simp add: Q) done lemma whenE_throwError_corres: assumes P: "frel f f'" assumes Q: "P = P'" assumes R: "\ P \ corres_underlying sr nf nf' (frel \ rvr) Q Q' m m'" shows "corres_underlying sr nf nf' (frel \ rvr) (\s. \ P \ Q s) (\s. \ P' \ Q' s) (whenE P (throwError f ) >>=E (\_. m )) (whenE P' (throwError f') >>=E (\_. m'))" apply (rule whenE_throwError_corres_initial) apply (simp_all add: P Q R) done lemma corres_move_asm: "\ corres_underlying sr nf nf' r P Q f g; \s s'. \(s,s') \ sr; P s; P' s'\ \ Q s'\ \ corres_underlying sr nf nf' r P P' f g" by (fastforce simp: corres_underlying_def) lemmas corres_cross_over_guard = corres_move_asm[rotated] lemma corres_weak_cong: "\\s. P s \ f s = f' s; \s. Q s \ g s = g' s\ \ corres_underlying sr nf nf' r P Q f g = corres_underlying sr nf nf' r P Q f' g'" by (simp cong: corres_cong) lemma corres_either_alternate: "\ corres_underlying sr nf nf' r P Pa' a c; corres_underlying sr nf nf' r P Pb' b c \ \ corres_underlying sr nf nf' r P (Pa' or Pb') (a \ b) c" apply (simp add: corres_underlying_def alternative_def) apply clarsimp apply (drule (1) bspec, clarsimp)+ apply (erule disjE, clarsimp) apply (drule(1) bspec, clarsimp) apply (rule rev_bexI) apply (erule UnI1) apply simp apply (clarsimp, drule(1) bspec, clarsimp) apply (rule rev_bexI) apply (erule UnI2) apply simp done lemma corres_either_alternate2: "\ corres_underlying sr nf nf' r P R a c; corres_underlying sr nf nf' r Q R b c \ \ corres_underlying sr nf nf' r (P or Q) R (a \ b) c" apply (simp add: corres_underlying_def alternative_def) apply clarsimp apply (drule (1) bspec, clarsimp)+ apply (erule disjE) apply clarsimp apply (drule(1) bspec, clarsimp) apply (rule rev_bexI) apply (erule UnI1) apply simp apply clarsimp apply (drule(1) bspec, clarsimp) apply (rule rev_bexI) apply (erule UnI2) apply simp done lemma option_corres: assumes None: "\ x = None; x' = None \ \ corres_underlying sr nf nf' r P P' (A None) (C None)" assumes Some: "\z z'. \ x = Some z; x' = Some z' \ \ corres_underlying sr nf nf' r (Q z) (Q' z') (A (Some z)) (C (Some z'))" assumes None_eq: "(x = None) = (x' = None)" shows "corres_underlying sr nf nf' r (\s. (x = None \ P s) \ (\z. x = Some z \ Q z s)) (\s. (x' = None \ P' s) \ (\z. x' = Some z \ Q' z s)) (A x) (C x')" apply (cases x; cases x'; simp add: assms) apply (simp add: None flip: None_eq) apply (simp flip: None_eq) done lemma corres_bind_return: "corres_underlying sr nf nf' r P P' (f >>= return) g \ corres_underlying sr nf nf' r P P' f g" by (simp add: corres_underlying_def) lemma corres_bind_return2: "corres_underlying sr nf nf' r P P' f (g >>= return) \ corres_underlying sr nf nf' r P P' f g" by simp lemma corres_stateAssert_implied2: assumes c: "corres_underlying sr nf nf' r P Q f g" assumes sr: "\s s'. \(s, s') \ sr; R s; R' s'\ \ Q' s'" assumes f: "\P\ f \\_. R\" assumes g: "\Q\ g \\_. R'\" shows "corres_underlying sr nf nf' dc P Q f (g >>= (\_. stateAssert Q' []))" apply (subst bind_return[symmetric]) apply (rule corres_guard_imp) apply (rule corres_split) prefer 2 apply (rule c) apply (clarsimp simp: corres_underlying_def return_def stateAssert_def bind_def get_def assert_def fail_def) apply (drule (2) sr) apply simp apply (rule f) apply (rule g) apply simp apply simp done lemma corres_add_noop_lhs: "corres_underlying sr nf nf' r P P' (return () >>= (\_. f)) g \ corres_underlying sr nf nf' r P P' f g" by simp lemma corres_add_noop_lhs2: "corres_underlying sr nf nf' r P P' (f >>= (\_. return ())) g \ corres_underlying sr nf nf' r P P' f g" by simp lemmas corres_split_noop_rhs = corres_split_nor[THEN corres_add_noop_lhs, OF _ _ return_wp] lemmas corres_split_noop_rhs2 = corres_split_nor[THEN corres_add_noop_lhs2] lemmas corres_split_dc = corres_split[where r'=dc, simplified] lemma isLeft_case_sum: "isLeft v \ (case v of Inl v' \ f v' | Inr v' \ g v') = f (theLeft v)" by (clarsimp simp: isLeft_def) lemma corres_symb_exec_catch_r: "\ \rv. corres_underlying sr nf nf' r P (Q' rv) f (h rv); \P'\ g \\\\, \Q'\; \s. \(=) s\ g \\r. (=) s\; nf' \ no_fail P' g \ \ corres_underlying sr nf nf' r P P' f (g h)" apply (simp add: catch_def) apply (rule corres_symb_exec_r, simp_all) apply (rule_tac F="isLeft x" in corres_gen_asm2) apply (simp add: isLeft_case_sum) apply assumption apply (simp add: validE_def) apply (erule hoare_chain, simp_all)[1] apply (simp add: isLeft_def split: sum.split_asm) done lemma corres_return_eq_same: "a = b \ corres_underlying srel nf' nf (=) \ \ (return a) (return b)" apply (simp add: corres_underlying_def return_def) done lemmas corres_discard_r = corres_symb_exec_r [where P'=P' and Q'="\_. P'" for P', simplified] lemmas corres_returnTT = corres_return[where P=\ and P'=\, THEN iffD2] lemma corres_assert_gen_asm: "\ F \ corres_underlying sr nf nf' r P Q f (g ()) \ \ corres_underlying sr nf nf' r (P and (\_. F)) Q f (assert F >>= g)" by (simp add: corres_gen_asm) lemma corres_assert_gen_asm_l: "\ F \ corres_underlying sr nf nf' r P Q (f ()) g \ \ corres_underlying sr nf nf' r (P and (\_. F)) Q (assert F >>= f) g" by (simp add: corres_gen_asm) lemma corres_assert_gen_asm2: "\ F \ corres_underlying sr nf nf' r P Q f (g ()) \ \ corres_underlying sr nf nf' r P (Q and (\_. F)) f (assert F >>= g)" by (simp add: corres_gen_asm2) lemma corres_assert_gen_asm_l2: "\ F \ corres_underlying sr nf nf' r P Q (f ()) g \ \ corres_underlying sr nf nf' r P (Q and (\_. F)) (assert F >>= f) g" by (simp add: corres_gen_asm2) lemma corres_add_guard: "\\s s'. \Q s; Q' s'; (s, s') \ sr\ \ P s \ P' s'; corres_underlying sr nf nf' r (Q and P) (Q' and P') f g\ \ corres_underlying sr nf nf' r Q Q' f g" by (auto simp: corres_underlying_def) (* safer non-rewrite version of corres_gets *) lemma corres_gets_trivial: "\\s s'. (s,s') \ sr \ f s = f' s' \ \ corres_underlying sr nf nf' (=) \ \ (gets f) (gets f')" unfolding corres_underlying_def gets_def get_def return_def bind_def by clarsimp text \Some setup of specialised methods.\ lemma (in strengthen_implementation) wpfix_strengthen_corres_guard_imp: "(\s. st (\ F) (\) (P s) (Q s)) \ (\s. st (\ F) (\) (P' s) (Q' s)) \ st F (\) (corres_underlying sr nf nf' r P P' f g) (corres_underlying sr nf nf' r Q Q' f g)" by (cases F; auto elim: corres_guard_imp) lemmas wpfix_strengthen_corres_guard_imp[wp_fix_strgs] = strengthen_implementation.wpfix_strengthen_corres_guard_imp lemma corres_name_pre: "\ \s s'. \ P s; P' s'; (s, s') \ sr \ \ corres_underlying sr nf nf' r ((=) s) ((=) s') f g \ \ corres_underlying sr nf nf' r P P' f g" apply (simp add: corres_underlying_def split_def Ball_def) apply blast done lemma corres_return_trivial: "corres_underlying srel nf' nf dc \ \ (return a) (return b)" by (simp add: corres_underlying_def return_def) lemma mapME_x_corres_same_xs: assumes x: "\x. x \ set xs \ corres_underlying sr nf nf' (f \ dc) (P x) (P' x) (m x) (m' x)" assumes y: "\x. x \ set xs \ \\s. \y \ set xs. P y s\ m x \\_ s. \y \ set xs. P y s\,-" "\x. x \ set xs \ \\s. \y \ set xs. P' y s\ m' x \\_ s. \y \ set xs. P' y s\,-" assumes z: "xs = ys" shows "corres_underlying sr nf nf' (f \ dc) (\s. \x \ set xs. P x s) (\s. \y \ set ys. P' y s) (mapME_x m xs) (mapME_x m' ys)" apply (subgoal_tac "set ys \ set xs \ corres_underlying sr nf nf' (f \ dc) (\s. \x \ set xs. P x s) (\s. \y \ set xs. P' y s) (mapME_x m ys) (mapME_x m' ys)") apply (simp add: z) proof (induct ys) case Nil show ?case by (simp add: mapME_x_def sequenceE_x_def returnOk_def) next case (Cons z zs) from Cons have IH: "corres_underlying sr nf nf' (f \ dc) (\s. \x\set xs. P x s) (\s. \y\set xs. P' y s) (mapME_x m zs) (mapME_x m' zs)" by (simp add: dc_def) from Cons have in_set: "z \ set xs" "set zs \ set xs" by auto thus ?case apply (simp add: mapME_x_def sequenceE_x_def) apply (fold mapME_x_def sequenceE_x_def dc_def) apply (rule corres_guard_imp) apply (rule corres_splitEE) apply (rule IH) apply (rule x, simp) apply (wp y | simp)+ done qed end