(* * * Copyright 2017, Data61, CSIRO * * 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(DATA61_BSD) *) theory CorresK_Lemmas imports "Lib.Corres_Method" "ExecSpec.Syscall_H" "ASpec.Syscall_A" begin lemma corres_throwError_str [corres_concrete_rER]: "corres_underlyingK sr nf nf' (r (Inl a) (Inl b)) r \ \ (throwError a) (throw b)" "corres_underlyingK sr nf nf' (r (Inl a) (Inl b)) r \ \ (throwError a) (throwError b)" by (simp add: corres_underlyingK_def)+ lemma corresK_use_guard: "(F \ corres_underlyingK sr nf nf' F r Q Q' f g) \ corres_underlyingK sr nf nf' F r Q Q' f g" by (simp add: corres_underlyingK_def) lemma mapME_x_corresK_inv: assumes x: "\x. corres_underlyingK sr nf nf' (F x) (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'\,-" shows "corres_underlyingK sr nf nf' (xs = ys \ (\x \ set xs. F x)) (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 (rule corresK_use_guard, elim conjE) subgoal premises prems unfolding \xs = ys\ proof (induct ys) case Nil show ?case by (simp add: mapME_x_def sequenceE_x_def returnOk_def corres_underlyingK_def) next case (Cons z zs) from Cons have IH: "corres_underlyingK sr nf nf' (\x\set zs. F x) (f \ dc) (\s. \x\set zs. P x s) (\s. \y\set zs. P' y s) (mapME_x m zs) (mapME_x m' zs)" by (auto simp add: dc_def) show ?case apply (simp add: mapME_x_def sequenceE_x_def) apply (fold mapME_x_def sequenceE_x_def dc_def) apply (corressimp corresK: x IH wp: y) done qed done lemma corresK_mapM: assumes S: "set (zip xs ys) \ S" assumes z: "\x y. corres_protect ((x, y) \ S) \ corres_underlyingK R nf nf' (F x y) r' P P' (f x) (f' y)" assumes w: "\x y. (x, y) \ S \ \P\ f x \\rv. P\" "\x y. (x, y) \ S \ \P'\ f' y \\rv. P'\" shows "corres_underlyingK R nf nf' ((length xs = length ys) \ (r [] []) \ (\x xs y ys. r xs ys \ r' x y \ r (x # xs) (y # ys)) \ (\(x,y)\S. F x y)) r P P' (mapM f xs) (mapM f' ys)" unfolding corres_underlyingK_def apply (rule impI, rule corres_mapM[of r r' S]) using assms unfolding corres_underlyingK_def by (auto simp: corres_protect_def) definition "F_all2 F xs ys = (\F'. (\x y xs ys. (F' x y \ list_all2 F' xs ys \ F x y xs ys)) \ list_all2 F' xs ys)" lemma F_all2_pointwise[simp]: "F_all2 (\x y _ _. F x y) xs ys = list_all2 F xs ys" apply (rule iffI) apply (clarsimp simp: F_all2_def) subgoal for F' apply (rule list_all2_induct_suffixeq[where Q=F']; simp) apply (drule_tac x=x in spec) apply (drule_tac x=y in spec) apply fastforce done apply (clarsimp simp:F_all2_def) apply (rule_tac x=F in exI) apply clarsimp done lemma F_all2_list: "F xs ys \ \F'. (\xs ys. (F xs ys = list_all2 F' xs ys)) \ F_all2 (\_ _ xs ys. F xs ys) xs ys" apply (clarsimp simp: F_all2_def) apply (rule_tac x=F' in exI) by auto lemma list_all2_conjD: "list_all2 (\x y. Q x y \ P x y) xs ys \ list_all2 Q xs ys \ list_all2 P xs ys" by (induct rule: list_all2_induct; simp) lemma list_all2_to_list_all: "list_all2 P xs xs = list_all (\x. P x x) xs" by (induct xs;simp) lemma list_all_mem_subset: "list_all (\y. y \ set xs) ys = (set ys \ set xs)" by (induct ys; simp) lemma F_all2_eq: "(\x xs'. x \ set xs \ set xs' \ set xs \ F x x xs' xs') \ F_all2 F xs xs" apply (simp add: F_all2_def) apply (rule_tac x="\x y. x \ set xs \ x = y" in exI) apply (intro conjI impI allI) apply (drule list_all2_conjD) apply (simp add: list.rel_eq) apply clarsimp apply (simp add: list_all2_to_list_all list_all_mem_subset) apply (rule list.rel_refl_strong;simp) done lemma F_all2_conjI: "F_all2 F xs ys \ F_all2 F' xs ys \ F_all2 (\x y xs ys. F x y xs ys \ F' x y xs ys) xs ys" apply (clarsimp simp: F_all2_def) apply (rule_tac x="\x y. F'a x y \ F'aa x y" in exI) by (auto dest: list_all2_conjD intro: list_all2_conj) lemma corresK_mapM_list_all2: assumes x: "\x y xs ys. corres_underlyingK sr nf nf' (F x y xs ys) r' (I (x#xs)) (I' (y#ys)) (f x) (f' y)" assumes "\x y xs. \ S x y; suffix (x#xs) as \ \ \ I (x#xs) \ f x \ \rv. I xs \" assumes "\x y ys. \ S x y; suffix (y#ys) cs \ \ \ I' (y#ys) \ f' y \ \rv. I' ys \" shows "corres_underlyingK sr nf nf' (r [] [] \ (\ x y xs ys. r' x y \ r xs ys \ r (x # xs) (y # ys)) \ list_all2 S as cs \ F_all2 F as cs) r (I as) (I' cs) (mapM f as) (mapM f' cs)" unfolding corres_underlyingK_def apply (clarsimp simp: F_all2_def) subgoal for F' apply (rule corres_mapM_list_all2[of r r' "\x y. S x y \ F' x y"]; (rule assms | assumption | clarsimp)?) apply (rule x[THEN corresK_unlift]) apply (drule list_all2_conjD) apply (clarsimp simp: assms | assumption)+ apply (rule list_all2_conj; simp) done done lemma corresK_discard_rv: assumes A[corresK]: "corres_underlyingK sr nf nf' F r' P P' a c" shows "corres_underlyingK sr nf nf' F dc P P' (do x \ a; return () od) (do x \ c; return () od)" by corressimp lemma corresK_mapM_mapM_x: assumes "corres_underlyingK sr nf nf' F r' P P' (mapM f as) (mapM f' cs)" shows "corres_underlyingK sr nf nf' F dc P P' (mapM_x f as) (mapM_x f' cs)" unfolding mapM_x_mapM by (rule corresK_discard_rv, rule assms) lemmas corresK_mapM_x_list_all2 = corresK_mapM_list_all2[where r'=dc, THEN corresK_mapM_mapM_x[where r'=dc], simplified] lemmas corresK_mapM_x = corresK_mapM[where r'=dc, THEN corresK_mapM_mapM_x[where r'=dc], simplified] lemma corresK_subst_both: "g' = f' \ g = f \ corres_underlyingK sr nf nf' F r P P' f f' \ corres_underlyingK sr nf nf' F r P P' g g'" by simp lemma if_fun_true: "(if A then B else (\_. True)) = (\s. (A \ B s))" by simp lemmas corresK_whenE [corres_splits] = corresK_if[THEN corresK_subst_both[OF whenE_def[THEN meta_eq_to_obj_eq] whenE_def[THEN meta_eq_to_obj_eq]], OF _ corresK_returnOk[where r="f \ dc" for f], simplified, simplified if_fun_true] lemmas corresK_head_splits[corres_splits] = corresK_split[where d="return", simplified] corresK_splitE[where d="returnOk", simplified] corresK_split[where b="return", simplified] corresK_splitE[where b="returnOk", simplified] lemmas corresK_modify = corres_modify[atomized, @corresK_convert] lemmas corresK_modifyT = corresK_modify[where P=\ and P'=\, simplified] lemma corresK_Id: "(nf' \ no_fail P' g) \ corres_underlyingK Id nf nf' (f = g \ (\rv. r rv rv)) r (\_. True) P' f g" apply (rule corresK_assume_guard) apply (rule corres_Id;simp) done lemmas [corresK] = corresK_Id[where nf'=False and r="(=)",simplified] corresK_Id[where nf'=False,simplified] corresK_Id[where nf'=True and r="(=)", simplified] corresK_Id[where nf'=True, simplified] lemma corresK_unit_rv_eq_any[corres_concrete_r]: "corres_underlyingK sr nf nf' F r P P' f f' \ corres_underlyingK sr nf nf' F (\(x :: unit) (y :: unit). x = y) P P' f f'" apply (clarsimp simp add: corres_underlyingK_def) apply (erule corres_rel_imp) apply simp done lemma corresK_unit_rv_dc_any[corres_concrete_r]: "corres_underlyingK sr nf nf' F r P P' f f' \ corres_underlyingK sr nf nf' F (\(x :: unit) (y :: unit). dc x y) P P' f f'" apply (clarsimp simp add: corres_underlyingK_def) apply (erule corres_rel_imp) apply simp done end