| 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>
 | 
| 75712 |      8 |     and possible_bit_msb: \<open>possible_bit TYPE('a) (word_length TYPE('a) - Suc 0)\<close>
 | 
|  |      9 |     and not_possible_bit_length: \<open>\<not> possible_bit TYPE('a) (word_length TYPE('a))\<close>
 | 
| 75655 |     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 | 
 | 
| 75712 |     17 | lemma possible_bit_iff_less_word_length:
 | 
|  |     18 |   \<open>possible_bit TYPE('a) n \<longleftrightarrow> n < word_length TYPE('a)\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
 | 
|  |     19 | proof
 | 
|  |     20 |   assume \<open>?P\<close>
 | 
|  |     21 |   show ?Q
 | 
|  |     22 |   proof (rule ccontr)
 | 
|  |     23 |     assume \<open>\<not> n < word_length TYPE('a)\<close>
 | 
|  |     24 |     then have \<open>word_length TYPE('a) \<le> n\<close>
 | 
|  |     25 |       by simp
 | 
|  |     26 |     with \<open>?P\<close> have \<open>possible_bit TYPE('a) (word_length TYPE('a))\<close>
 | 
|  |     27 |       by (rule possible_bit_less_imp)
 | 
|  |     28 |     with not_possible_bit_length show False ..
 | 
|  |     29 |   qed
 | 
|  |     30 | next
 | 
|  |     31 |   assume \<open>?Q\<close>
 | 
|  |     32 |   then have \<open>n \<le> word_length TYPE('a) - Suc 0\<close>
 | 
|  |     33 |     by simp
 | 
|  |     34 |   with possible_bit_msb show ?P
 | 
|  |     35 |     by (rule possible_bit_less_imp)
 | 
|  |     36 | qed
 | 
|  |     37 | 
 | 
| 75655 |     38 | end
 | 
|  |     39 | 
 | 
|  |     40 | instantiation word :: (len) word
 | 
|  |     41 | begin
 | 
|  |     42 | 
 | 
|  |     43 | definition word_length_word :: \<open>'a word itself \<Rightarrow> nat\<close>
 | 
|  |     44 |   where [simp, code_unfold]: \<open>word_length_word _ = LENGTH('a)\<close>
 | 
|  |     45 | 
 | 
|  |     46 | instance
 | 
|  |     47 |   by standard simp_all
 | 
|  |     48 | 
 | 
|  |     49 | end
 | 
|  |     50 | 
 | 
|  |     51 | context word
 | 
|  |     52 | begin
 | 
|  |     53 | 
 | 
|  |     54 | context
 | 
|  |     55 |   includes bit_operations_syntax
 | 
|  |     56 | begin
 | 
|  |     57 | 
 | 
|  |     58 | abbreviation lsb :: \<open>'a \<Rightarrow> bool\<close>
 | 
|  |     59 |   where \<open>lsb \<equiv> odd\<close>
 | 
|  |     60 | 
 | 
|  |     61 | definition msb :: \<open>'a \<Rightarrow> bool\<close>
 | 
|  |     62 |   where \<open>msb w = bit w (word_length TYPE('a) - Suc 0)\<close>
 | 
|  |     63 | 
 | 
|  |     64 | lemma not_msb_0 [simp]:
 | 
|  |     65 |   \<open>\<not> msb 0\<close>
 | 
|  |     66 |   by (simp add: msb_def)
 | 
|  |     67 | 
 | 
|  |     68 | lemma msb_minus_1 [simp]:
 | 
|  |     69 |   \<open>msb (- 1)\<close>
 | 
| 75712 |     70 |   by (simp add: msb_def possible_bit_iff_less_word_length)
 | 
| 75655 |     71 | 
 | 
|  |     72 | lemma msb_1_iff [simp]:
 | 
|  |     73 |   \<open>msb 1 \<longleftrightarrow> word_length TYPE('a) = 1\<close>
 | 
|  |     74 |   by (auto simp add: msb_def bit_simps le_less)
 | 
|  |     75 | 
 | 
|  |     76 | lemma msb_minus_iff [simp]:
 | 
|  |     77 |   \<open>msb (- w) \<longleftrightarrow> \<not> msb (w - 1)\<close>
 | 
| 75712 |     78 |   by (simp add: msb_def bit_simps possible_bit_iff_less_word_length)
 | 
| 75655 |     79 | 
 | 
|  |     80 | lemma msb_not_iff [simp]:
 | 
|  |     81 |   \<open>msb (NOT w) \<longleftrightarrow> \<not> msb w\<close>
 | 
| 75712 |     82 |   by (simp add: msb_def bit_simps possible_bit_iff_less_word_length)
 | 
| 75655 |     83 | 
 | 
|  |     84 | lemma msb_and_iff [simp]:
 | 
|  |     85 |   \<open>msb (v AND w) \<longleftrightarrow> msb v \<and> msb w\<close>
 | 
|  |     86 |   by (simp add: msb_def bit_simps)
 | 
|  |     87 | 
 | 
|  |     88 | lemma msb_or_iff [simp]:
 | 
|  |     89 |   \<open>msb (v OR w) \<longleftrightarrow> msb v \<or> msb w\<close>
 | 
|  |     90 |   by (simp add: msb_def bit_simps)
 | 
|  |     91 | 
 | 
|  |     92 | lemma msb_xor_iff [simp]:
 | 
|  |     93 |   \<open>msb (v XOR w) \<longleftrightarrow> \<not> (msb v \<longleftrightarrow> msb w)\<close>
 | 
|  |     94 |   by (simp add: msb_def bit_simps)
 | 
|  |     95 | 
 | 
|  |     96 | lemma msb_exp_iff [simp]:                                             
 | 
|  |     97 |   \<open>msb (2 ^ n) \<longleftrightarrow> n = word_length TYPE('a) - Suc 0\<close>
 | 
| 75712 |     98 |   by (simp add: msb_def bit_simps possible_bit_iff_less_word_length)
 | 
| 75655 |     99 | 
 | 
|  |    100 | lemma msb_mask_iff [simp]:
 | 
|  |    101 |   \<open>msb (mask n) \<longleftrightarrow> word_length TYPE('a) \<le> n\<close>
 | 
| 75712 |    102 |   by (simp add: msb_def bit_simps less_diff_conv2 Suc_le_eq less_Suc_eq_le possible_bit_iff_less_word_length)
 | 
| 75655 |    103 | 
 | 
|  |    104 | lemma msb_set_bit_iff [simp]:
 | 
|  |    105 |   \<open>msb (set_bit n w) \<longleftrightarrow> n = word_length TYPE('a) - Suc 0 \<or> msb w\<close>
 | 
|  |    106 |   by (simp add: set_bit_eq_or ac_simps)
 | 
|  |    107 | 
 | 
|  |    108 | lemma msb_unset_bit_iff [simp]:
 | 
|  |    109 |   \<open>msb (unset_bit n w) \<longleftrightarrow> n \<noteq> word_length TYPE('a) - Suc 0 \<and> msb w\<close>
 | 
|  |    110 |   by (simp add: unset_bit_eq_and_not ac_simps)
 | 
|  |    111 | 
 | 
|  |    112 | lemma msb_flip_bit_iff [simp]:
 | 
|  |    113 |   \<open>msb (flip_bit n w) \<longleftrightarrow> (n \<noteq> word_length TYPE('a) - Suc 0 \<longleftrightarrow> msb w)\<close>
 | 
|  |    114 |   by (auto simp add: flip_bit_eq_xor)
 | 
|  |    115 | 
 | 
|  |    116 | lemma msb_push_bit_iff:
 | 
|  |    117 |   \<open>msb (push_bit n w) \<longleftrightarrow> n < word_length TYPE('a) \<and> bit w (word_length TYPE('a) - Suc n)\<close>
 | 
| 75712 |    118 |   by (simp add: msb_def bit_simps le_diff_conv2 Suc_le_eq possible_bit_iff_less_word_length)
 | 
| 75655 |    119 | 
 | 
| 75712 |    120 | lemma msb_drop_bit_iff [simp]:
 | 
| 75655 |    121 |   \<open>msb (drop_bit n w) \<longleftrightarrow> n = 0 \<and> msb w\<close>
 | 
| 75712 |    122 |   by (cases n)
 | 
|  |    123 |     (auto simp add: msb_def bit_simps possible_bit_iff_less_word_length intro!: impossible_bit)
 | 
| 75655 |    124 | 
 | 
|  |    125 | lemma msb_take_bit_iff [simp]:
 | 
|  |    126 |   \<open>msb (take_bit n w) \<longleftrightarrow> word_length TYPE('a) \<le> n \<and> msb w\<close>
 | 
|  |    127 |   by (simp add: take_bit_eq_mask ac_simps)
 | 
|  |    128 | 
 | 
| 75712 |    129 | lemma msb_signed_take_bit_iff:
 | 
|  |    130 |   \<open>msb (signed_take_bit n w) \<longleftrightarrow> bit w (min n (word_length TYPE('a) - Suc 0))\<close>
 | 
|  |    131 |   by (simp add: msb_def bit_simps possible_bit_iff_less_word_length)
 | 
| 75655 |    132 | 
 | 
| 75713 |    133 | definition signed_drop_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
 | 
|  |    134 |   where \<open>signed_drop_bit n w = drop_bit n w
 | 
|  |    135 |     OR (of_bool (bit w (word_length TYPE('a) - Suc 0)) * NOT (mask (word_length TYPE('a) - Suc n)))\<close>
 | 
|  |    136 | 
 | 
|  |    137 | lemma msb_signed_drop_bit_iff [simp]:
 | 
|  |    138 |   \<open>msb (signed_drop_bit n w) \<longleftrightarrow> msb w\<close>
 | 
|  |    139 |   by (simp add: signed_drop_bit_def bit_simps not_le not_less)
 | 
|  |    140 |     (simp add: msb_def)
 | 
|  |    141 | 
 | 
| 75655 |    142 | end
 | 
|  |    143 | 
 | 
|  |    144 | end
 | 
|  |    145 | 
 | 
|  |    146 | end
 |