|
75655
|
1 |
theory Word_Lsb_Msb
|
|
|
2 |
imports "HOL-Library.Word"
|
|
|
3 |
begin
|
|
|
4 |
|
|
|
5 |
class word = ring_bit_operations +
|
|
|
6 |
fixes word_length :: \<open>'a itself \<Rightarrow> nat\<close>
|
|
|
7 |
assumes word_length_positive [simp]: \<open>0 < word_length TYPE('a)\<close>
|
|
|
8 |
and possible_bit_msb [simp]: \<open>possible_bit TYPE('a) (word_length TYPE('a) - Suc 0)\<close>
|
|
|
9 |
and not_possible_bit_length [simp]: \<open>\<not> possible_bit TYPE('a) (word_length TYPE('a))\<close>
|
|
|
10 |
begin
|
|
|
11 |
|
|
|
12 |
lemma word_length_not_0 [simp]:
|
|
|
13 |
\<open>word_length TYPE('a) \<noteq> 0\<close>
|
|
|
14 |
using word_length_positive
|
|
|
15 |
by simp
|
|
|
16 |
|
|
|
17 |
end
|
|
|
18 |
|
|
|
19 |
instantiation word :: (len) word
|
|
|
20 |
begin
|
|
|
21 |
|
|
|
22 |
definition word_length_word :: \<open>'a word itself \<Rightarrow> nat\<close>
|
|
|
23 |
where [simp, code_unfold]: \<open>word_length_word _ = LENGTH('a)\<close>
|
|
|
24 |
|
|
|
25 |
instance
|
|
|
26 |
by standard simp_all
|
|
|
27 |
|
|
|
28 |
end
|
|
|
29 |
|
|
|
30 |
context word
|
|
|
31 |
begin
|
|
|
32 |
|
|
|
33 |
context
|
|
|
34 |
includes bit_operations_syntax
|
|
|
35 |
begin
|
|
|
36 |
|
|
|
37 |
abbreviation lsb :: \<open>'a \<Rightarrow> bool\<close>
|
|
|
38 |
where \<open>lsb \<equiv> odd\<close>
|
|
|
39 |
|
|
|
40 |
definition msb :: \<open>'a \<Rightarrow> bool\<close>
|
|
|
41 |
where \<open>msb w = bit w (word_length TYPE('a) - Suc 0)\<close>
|
|
|
42 |
|
|
|
43 |
lemma not_msb_0 [simp]:
|
|
|
44 |
\<open>\<not> msb 0\<close>
|
|
|
45 |
by (simp add: msb_def)
|
|
|
46 |
|
|
|
47 |
lemma msb_minus_1 [simp]:
|
|
|
48 |
\<open>msb (- 1)\<close>
|
|
|
49 |
by (simp add: msb_def)
|
|
|
50 |
|
|
|
51 |
lemma msb_1_iff [simp]:
|
|
|
52 |
\<open>msb 1 \<longleftrightarrow> word_length TYPE('a) = 1\<close>
|
|
|
53 |
by (auto simp add: msb_def bit_simps le_less)
|
|
|
54 |
|
|
|
55 |
lemma msb_minus_iff [simp]:
|
|
|
56 |
\<open>msb (- w) \<longleftrightarrow> \<not> msb (w - 1)\<close>
|
|
|
57 |
by (simp add: msb_def bit_simps)
|
|
|
58 |
|
|
|
59 |
lemma msb_not_iff [simp]:
|
|
|
60 |
\<open>msb (NOT w) \<longleftrightarrow> \<not> msb w\<close>
|
|
|
61 |
by (simp add: msb_def bit_simps)
|
|
|
62 |
|
|
|
63 |
lemma msb_and_iff [simp]:
|
|
|
64 |
\<open>msb (v AND w) \<longleftrightarrow> msb v \<and> msb w\<close>
|
|
|
65 |
by (simp add: msb_def bit_simps)
|
|
|
66 |
|
|
|
67 |
lemma msb_or_iff [simp]:
|
|
|
68 |
\<open>msb (v OR w) \<longleftrightarrow> msb v \<or> msb w\<close>
|
|
|
69 |
by (simp add: msb_def bit_simps)
|
|
|
70 |
|
|
|
71 |
lemma msb_xor_iff [simp]:
|
|
|
72 |
\<open>msb (v XOR w) \<longleftrightarrow> \<not> (msb v \<longleftrightarrow> msb w)\<close>
|
|
|
73 |
by (simp add: msb_def bit_simps)
|
|
|
74 |
|
|
|
75 |
lemma msb_exp_iff [simp]:
|
|
|
76 |
\<open>msb (2 ^ n) \<longleftrightarrow> n = word_length TYPE('a) - Suc 0\<close>
|
|
|
77 |
by (simp add: msb_def bit_simps)
|
|
|
78 |
|
|
|
79 |
lemma msb_mask_iff [simp]:
|
|
|
80 |
\<open>msb (mask n) \<longleftrightarrow> word_length TYPE('a) \<le> n\<close>
|
|
|
81 |
by (simp add: msb_def bit_simps less_diff_conv2 Suc_le_eq less_Suc_eq_le)
|
|
|
82 |
|
|
|
83 |
lemma msb_set_bit_iff [simp]:
|
|
|
84 |
\<open>msb (set_bit n w) \<longleftrightarrow> n = word_length TYPE('a) - Suc 0 \<or> msb w\<close>
|
|
|
85 |
by (simp add: set_bit_eq_or ac_simps)
|
|
|
86 |
|
|
|
87 |
lemma msb_unset_bit_iff [simp]:
|
|
|
88 |
\<open>msb (unset_bit n w) \<longleftrightarrow> n \<noteq> word_length TYPE('a) - Suc 0 \<and> msb w\<close>
|
|
|
89 |
by (simp add: unset_bit_eq_and_not ac_simps)
|
|
|
90 |
|
|
|
91 |
lemma msb_flip_bit_iff [simp]:
|
|
|
92 |
\<open>msb (flip_bit n w) \<longleftrightarrow> (n \<noteq> word_length TYPE('a) - Suc 0 \<longleftrightarrow> msb w)\<close>
|
|
|
93 |
by (auto simp add: flip_bit_eq_xor)
|
|
|
94 |
|
|
|
95 |
lemma msb_push_bit_iff:
|
|
|
96 |
\<open>msb (push_bit n w) \<longleftrightarrow> n < word_length TYPE('a) \<and> bit w (word_length TYPE('a) - Suc n)\<close>
|
|
|
97 |
by (simp add: msb_def bit_simps le_diff_conv2 Suc_le_eq)
|
|
|
98 |
|
|
|
99 |
(*lemma msb_drop_bit_iff [simp]:
|
|
|
100 |
\<open>msb (drop_bit n w) \<longleftrightarrow> n = 0 \<and> msb w\<close>
|
|
|
101 |
apply (cases n)
|
|
|
102 |
apply simp_all
|
|
|
103 |
apply (auto simp add: msb_def bit_simps)
|
|
|
104 |
oops*)
|
|
|
105 |
|
|
|
106 |
lemma msb_take_bit_iff [simp]:
|
|
|
107 |
\<open>msb (take_bit n w) \<longleftrightarrow> word_length TYPE('a) \<le> n \<and> msb w\<close>
|
|
|
108 |
by (simp add: take_bit_eq_mask ac_simps)
|
|
|
109 |
|
|
|
110 |
(*lemma msb_signed_take_bit_iff:
|
|
|
111 |
\<open>msb (signed_take_bit n w) \<longleftrightarrow> P w n\<close>
|
|
|
112 |
unfolding signed_take_bit_def
|
|
|
113 |
apply (simp add: signed_take_bit_def not_le)
|
|
|
114 |
apply auto
|
|
|
115 |
oops*)
|
|
|
116 |
|
|
|
117 |
end
|
|
|
118 |
|
|
|
119 |
end
|
|
|
120 |
|
|
|
121 |
end
|