| author | desharna |
| Thu, 08 Oct 2020 16:36:00 +0200 | |
| changeset 72399 | f8900a5ad4a7 |
| parent 72388 | 633d14bd1e59 |
| permissions | -rw-r--r-- |
| 72000 | 1 |
(* Author: Jeremy Dawson, NICTA |
2 |
*) |
|
3 |
||
4 |
section \<open>Operation variant for the most significant bit\<close> |
|
5 |
||
6 |
theory Misc_msb |
|
|
72088
a36db1c8238e
separation of reversed bit lists from other material
haftmann
parents:
72000
diff
changeset
|
7 |
imports |
|
a36db1c8238e
separation of reversed bit lists from other material
haftmann
parents:
72000
diff
changeset
|
8 |
Word |
|
a36db1c8238e
separation of reversed bit lists from other material
haftmann
parents:
72000
diff
changeset
|
9 |
Reversed_Bit_Lists |
| 72000 | 10 |
begin |
11 |
||
12 |
class msb = |
|
13 |
fixes msb :: \<open>'a \<Rightarrow> bool\<close> |
|
14 |
||
15 |
instantiation int :: msb |
|
16 |
begin |
|
17 |
||
18 |
definition \<open>msb x \<longleftrightarrow> x < 0\<close> for x :: int |
|
19 |
||
20 |
instance .. |
|
21 |
||
22 |
end |
|
23 |
||
24 |
lemma msb_conv_bin_sign: "msb x \<longleftrightarrow> bin_sign x = -1" |
|
25 |
by(simp add: bin_sign_def not_le msb_int_def) |
|
26 |
||
27 |
lemma msb_bin_rest [simp]: "msb (bin_rest x) = msb x" |
|
28 |
by(simp add: msb_int_def) |
|
29 |
||
30 |
lemma int_msb_and [simp]: "msb ((x :: int) AND y) \<longleftrightarrow> msb x \<and> msb y" |
|
31 |
by(simp add: msb_int_def) |
|
32 |
||
33 |
lemma int_msb_or [simp]: "msb ((x :: int) OR y) \<longleftrightarrow> msb x \<or> msb y" |
|
34 |
by(simp add: msb_int_def) |
|
35 |
||
36 |
lemma int_msb_xor [simp]: "msb ((x :: int) XOR y) \<longleftrightarrow> msb x \<noteq> msb y" |
|
37 |
by(simp add: msb_int_def) |
|
38 |
||
39 |
lemma int_msb_not [simp]: "msb (NOT (x :: int)) \<longleftrightarrow> \<not> msb x" |
|
40 |
by(simp add: msb_int_def not_less) |
|
41 |
||
42 |
lemma msb_shiftl [simp]: "msb ((x :: int) << n) \<longleftrightarrow> msb x" |
|
43 |
by(simp add: msb_int_def) |
|
44 |
||
45 |
lemma msb_shiftr [simp]: "msb ((x :: int) >> r) \<longleftrightarrow> msb x" |
|
46 |
by(simp add: msb_int_def) |
|
47 |
||
48 |
lemma msb_bin_sc [simp]: "msb (bin_sc n b x) \<longleftrightarrow> msb x" |
|
49 |
by(simp add: msb_conv_bin_sign) |
|
50 |
||
51 |
lemma msb_0 [simp]: "msb (0 :: int) = False" |
|
52 |
by(simp add: msb_int_def) |
|
53 |
||
54 |
lemma msb_1 [simp]: "msb (1 :: int) = False" |
|
55 |
by(simp add: msb_int_def) |
|
56 |
||
57 |
lemma msb_numeral [simp]: |
|
58 |
"msb (numeral n :: int) = False" |
|
59 |
"msb (- numeral n :: int) = True" |
|
60 |
by(simp_all add: msb_int_def) |
|
61 |
||
62 |
instantiation word :: (len) msb |
|
63 |
begin |
|
64 |
||
65 |
definition msb_word :: \<open>'a word \<Rightarrow> bool\<close> |
|
66 |
where \<open>msb a \<longleftrightarrow> bin_sign (sbintrunc (LENGTH('a) - 1) (uint a)) = - 1\<close>
|
|
67 |
||
68 |
lemma msb_word_eq: |
|
69 |
\<open>msb w \<longleftrightarrow> bit w (LENGTH('a) - 1)\<close> for w :: \<open>'a::len word\<close>
|
|
70 |
by (simp add: msb_word_def bin_sign_lem bit_uint_iff) |
|
71 |
||
72 |
instance .. |
|
73 |
||
74 |
end |
|
75 |
||
| 72388 | 76 |
lemma msb_word_iff_bit: |
77 |
\<open>msb w \<longleftrightarrow> bit w (LENGTH('a) - Suc 0)\<close>
|
|
78 |
for w :: \<open>'a::len word\<close> |
|
79 |
by (simp add: msb_word_def bin_sign_def bit_uint_iff) |
|
80 |
||
| 72000 | 81 |
lemma word_msb_def: |
82 |
"msb a \<longleftrightarrow> bin_sign (sint a) = - 1" |
|
83 |
by (simp add: msb_word_def sint_uint) |
|
84 |
||
85 |
lemma word_msb_sint: "msb w \<longleftrightarrow> sint w < 0" |
|
86 |
by (simp add: msb_word_eq bit_last_iff) |
|
87 |
||
| 72388 | 88 |
lemma msb_word_iff_sless_0: |
89 |
\<open>msb w \<longleftrightarrow> w <s 0\<close> |
|
90 |
by (simp add: word_msb_sint word_sless_alt) |
|
91 |
||
| 72000 | 92 |
lemma msb_word_of_int: "msb (word_of_int x::'a::len word) = bin_nth x (LENGTH('a) - 1)"
|
93 |
by (simp add: word_msb_def word_sbin.eq_norm bin_sign_lem) |
|
94 |
||
95 |
lemma word_msb_numeral [simp]: |
|
96 |
"msb (numeral w::'a::len word) = bin_nth (numeral w) (LENGTH('a) - 1)"
|
|
97 |
unfolding word_numeral_alt by (rule msb_word_of_int) |
|
98 |
||
99 |
lemma word_msb_neg_numeral [simp]: |
|
100 |
"msb (- numeral w::'a::len word) = bin_nth (- numeral w) (LENGTH('a) - 1)"
|
|
101 |
unfolding word_neg_numeral_alt by (rule msb_word_of_int) |
|
102 |
||
103 |
lemma word_msb_0 [simp]: "\<not> msb (0::'a::len word)" |
|
104 |
by (simp add: word_msb_def bin_sign_def sint_uint sbintrunc_eq_take_bit) |
|
105 |
||
106 |
lemma word_msb_1 [simp]: "msb (1::'a::len word) \<longleftrightarrow> LENGTH('a) = 1"
|
|
107 |
unfolding word_1_wi msb_word_of_int eq_iff [where 'a=nat] |
|
108 |
by (simp add: Suc_le_eq) |
|
109 |
||
110 |
lemma word_msb_nth: "msb w = bin_nth (uint w) (LENGTH('a) - 1)"
|
|
111 |
for w :: "'a::len word" |
|
112 |
by (simp add: word_msb_def sint_uint bin_sign_lem) |
|
113 |
||
114 |
lemma word_msb_alt: "msb w \<longleftrightarrow> hd (to_bl w)" |
|
115 |
for w :: "'a::len word" |
|
116 |
apply (simp add: msb_word_eq) |
|
117 |
apply (subst hd_conv_nth) |
|
118 |
apply simp |
|
119 |
apply (subst nth_to_bl) |
|
120 |
apply simp |
|
121 |
apply simp |
|
122 |
done |
|
123 |
||
124 |
lemma msb_nth: "msb w = w !! (LENGTH('a) - 1)"
|
|
125 |
for w :: "'a::len word" |
|
126 |
by (simp add: word_msb_nth word_test_bit_def) |
|
127 |
||
128 |
lemma word_msb_n1 [simp]: "msb (-1::'a::len word)" |
|
129 |
by (simp add: msb_word_eq exp_eq_zero_iff not_le) |
|
130 |
||
131 |
lemma msb_shift: "msb w \<longleftrightarrow> w >> (LENGTH('a) - 1) \<noteq> 0"
|
|
132 |
for w :: "'a::len word" |
|
133 |
by (simp add: msb_word_eq shiftr_word_eq bit_iff_odd_drop_bit drop_bit_eq_zero_iff_not_bit_last) |
|
134 |
||
135 |
lemmas word_ops_msb = msb1 [unfolded msb_nth [symmetric, unfolded One_nat_def]] |
|
136 |
||
137 |
end |