cparser+crefine: move h_t_array_valid_array_assertion to cparser session

Signed-off-by: Michael McInerney <michael.mcinerney@proofcraft.systems>
This commit is contained in:
Michael McInerney 2023-01-30 12:55:41 +10:30 committed by michaelmcinerney
parent a13db04598
commit 3c322eab1d
2 changed files with 7 additions and 7 deletions

View File

@ -1913,13 +1913,6 @@ lemmas ccorres_move_c_guards =
ccorres_move_Guard[OF abs_c_guard_from_abs_h_t_valid]
ccorres_move_Guard
lemma h_t_array_valid_array_assertion:
"h_t_array_valid htd ptr n \<Longrightarrow> 0 < n
\<Longrightarrow> array_assertion ptr n htd"
apply (simp add: array_assertion_def)
apply (fastforce intro: exI[where x=0])
done
lemma array_assertion_abs_to_const:
"\<forall>s s'. (s, s') \<in> rf_sr \<and> P s \<and> P' s'
\<longrightarrow> (Suc 0 = 0 \<or> array_assertion (ptr s s') (n s s') (htd s s'))

View File

@ -106,6 +106,13 @@ lemma array_ptr_valid_array_assertionI:
by (auto dest: array_ptr_valid_array_assertionD
simp: array_assertion_shrink_right)
lemma h_t_array_valid_array_assertion:
"h_t_array_valid htd ptr n \<Longrightarrow> 0 < n
\<Longrightarrow> array_assertion ptr n htd"
apply (simp add: array_assertion_def)
apply (fastforce intro: exI[where x=0])
done
text \<open>Derived from array_assertion, an appropriate assertion for performing
a pointer addition, or for dereferencing a pointer addition (the strong case).