lib+proof: rename bind_assoc_reverse to bind_assoc_return_reverse

This also improves the style of this lemma

Signed-off-by: Michael McInerney <michael.mcinerney@proofcraft.systems>
This commit is contained in:
Michael McInerney 2023-03-26 16:50:56 +10:30 committed by michaelmcinerney
parent 3981e9a60e
commit 27d838af86
6 changed files with 16 additions and 8 deletions

View File

@ -388,9 +388,17 @@ lemma bind_assoc2:
"(do x \<leftarrow> a; _ \<leftarrow> b; c x od) = (do x \<leftarrow> (do x' \<leftarrow> a; _ \<leftarrow> b; return x' od); c x od)"
by (simp add: bind_assoc)
lemma bind_assoc_reverse:
"(do x \<leftarrow> A; _ \<leftarrow> B x; C x od) =
(do x \<leftarrow> do x \<leftarrow> A; _ \<leftarrow> B x; return x od; C x od)"
lemma bind_assoc_return_reverse:
"do x \<leftarrow> f;
_ \<leftarrow> g x;
h x
od =
do x \<leftarrow> do x \<leftarrow> f;
_ \<leftarrow> g x;
return x
od;
h x
od"
by (simp only: bind_assoc return_bind)
lemma if_bind:

View File

@ -827,7 +827,7 @@ text \<open>
state.
The extra return statement on the Haskell side allows us to establish a nontrivial return relation
between the values set on the concrete and abstract side. The @{thm bind_assoc_reverse} rule
between the values set on the concrete and abstract side. The @{thm bind_assoc_return_reverse} rule
may assist with rewriting statements to add the extra return needed by this rule\<close>
lemma ccorres_call_getter_setter:
assumes cul: "ccorresG sr \<Gamma> (=) xf' P (i ` P') [] getter (Call f)"

View File

@ -5329,7 +5329,7 @@ lemma corres_retype_region_createNewCaps:
\<comment> \<open>CapTable\<close>
apply (subst retype_region2_extra_ext_trivial)
apply (simp add: APIType_map2_def)
apply (subst bind_assoc_reverse[of "createObjects y n (KOCTE makeObject) us"])
apply (subst bind_assoc_return_reverse[of "createObjects y n (KOCTE makeObject) us"])
apply (subst liftM_def
[of "map (\<lambda>addr. capability.CNodeCap addr us 0 0)", symmetric])
apply simp

View File

@ -5381,7 +5381,7 @@ lemma corres_retype_region_createNewCaps:
\<comment> \<open>CapTable\<close>
apply (subst retype_region2_extra_ext_trivial)
apply (simp add: APIType_map2_def)
apply (subst bind_assoc_reverse[of "createObjects y n (KOCTE makeObject) us"])
apply (subst bind_assoc_return_reverse[of "createObjects y n (KOCTE makeObject) us"])
apply (subst liftM_def
[of "map (\<lambda>addr. capability.CNodeCap addr us 0 0)", symmetric])
apply simp

View File

@ -5162,7 +5162,7 @@ lemma corres_retype_region_createNewCaps:
\<comment> \<open>CapTable\<close>
apply (subst retype_region2_extra_ext_trivial)
apply (simp add: APIType_map2_def)
apply (subst bind_assoc_reverse[of "createObjects y n (KOCTE makeObject) us"])
apply (subst bind_assoc_return_reverse[of "createObjects y n (KOCTE makeObject) us"])
apply (subst liftM_def[of "map (\<lambda>addr. capability.CNodeCap addr us 0 0)", symmetric])
apply simp
apply (rule corres_rel_imp)

View File

@ -5436,7 +5436,7 @@ lemma corres_retype_region_createNewCaps:
\<comment> \<open>CapTable\<close>
apply (subst retype_region2_extra_ext_trivial)
apply (simp add: APIType_map2_def)
apply (subst bind_assoc_reverse[of "createObjects y n (KOCTE makeObject) us"])
apply (subst bind_assoc_return_reverse[of "createObjects y n (KOCTE makeObject) us"])
apply (subst liftM_def[of "map (\<lambda>addr. capability.CNodeCap addr us 0 0)", symmetric])
apply simp
apply (rule corres_rel_imp)