isabelle2021-1 word_lib: add lemmas from l4v

Signed-off-by: Gerwin Klein <gerwin.klein@proofcraft.systems>
This commit is contained in:
Gerwin Klein 2022-02-13 19:52:54 +11:00 committed by Gerwin Klein
parent b29a3433ef
commit 65fbeb5b01
4 changed files with 32 additions and 2 deletions

View File

@ -122,8 +122,9 @@ lemma word_ge_min:
\<open>- (2 ^ (word_bits - 1)) \<le> sint x\<close> for x :: machine_word
using sint_ge [of x] by (simp add: word_bits_def)
(* We want the rhs unfolded here for later matching *)
lemma word_rsplit_0:
"word_rsplit (0 :: machine_word) = replicate (word_bits div 8) (0 :: 8 word)"
"word_rsplit (0::machine_word) = [0, 0, 0, (0::8 word)]"
by (simp add: word_rsplit_def bin_rsplit_def word_bits_def word_size_def Cons_replicate_eq)
lemma x_less_2_0_1:

View File

@ -122,8 +122,9 @@ lemma word_ge_min:
\<open>- (2 ^ (word_bits - 1)) \<le> sint x\<close> for x :: machine_word
using sint_ge [of x] by (simp add: word_bits_def)
(* We want the rhs unfolded here for later matching *)
lemma word_rsplit_0:
"word_rsplit (0 :: machine_word) = replicate (word_bits div 8) (0 :: 8 word)"
"word_rsplit (0::machine_word) = [0, 0, 0, 0, 0, 0, 0, (0::8 word)]"
by (simp add: word_rsplit_def bin_rsplit_def word_bits_def word_size_def Cons_replicate_eq)
lemma x_less_2_0_1:

View File

@ -783,6 +783,10 @@ lemma from_bool_eqI:
unfolding from_bool_def
by (auto split: bool.splits)
lemma from_bool_odd_eq_and:
"from_bool (odd w) = w AND 1"
unfolding from_bool_def by (simp add: word_and_1)
lemma neg_mask_in_mask_range:
"is_aligned ptr bits \<Longrightarrow> (ptr' AND NOT(mask bits) = ptr) = (ptr' \<in> mask_range ptr bits)"
apply (erule is_aligned_get_word_bits)

View File

@ -19,6 +19,8 @@ theory Word_Lemmas
Word_EqI
begin
lemmas take_bit_Suc_numeral[simp] = take_bit_Suc[where a="numeral w" for w]
context
includes bit_operations_syntax
begin
@ -1886,6 +1888,28 @@ lemma nasty_split_less:
apply simp+
done
lemma is_aligned_shiftr_add:
"\<lbrakk>is_aligned a n; is_aligned b m; b < 2^n; m \<le> n; n < LENGTH('a)\<rbrakk>
\<Longrightarrow> a + b >> m = (a >> m) + (b >> m)" for a :: "'a::len word"
apply(simp add: shiftr_div_2n_w word_size)
apply (rule word_unat_eq_iff[THEN iffD2])
apply (subst unat_plus_simple[THEN iffD1])
apply (subst shiftr_div_2n_w[symmetric])+
apply (rule is_aligned_no_wrap')
apply (rule is_aligned_shiftr[where n = "n - m"])
apply simp
apply (rule shiftr_less_t2n)
apply simp
apply (simp add:unat_div)
apply (subst unat_plus_simple[THEN iffD1])
apply (erule is_aligned_no_wrap')
apply simp
by (meson div_plus_div_distrib_dvd_left is_aligned_iff_dvd_nat is_aligned_weaken)
lemma shiftr_eq_neg_mask_eq:
"a >> b = c >> b \<Longrightarrow> a AND NOT (mask b) = c AND NOT (mask b)" for a :: "'a::len word"
by word_eqI (metis less_eqE)
end
end