| author | bulwahn | 
| Wed, 31 Mar 2010 16:44:41 +0200 | |
| changeset 36052 | c240b2a5df90 | 
| parent 35440 | bdf8ad377877 | 
| permissions | -rw-r--r-- | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1 | (* Title: HOL/Library/Word.thy | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32456diff
changeset | 2 | Author: Sebastian Skalberg, TU Muenchen | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 3 | *) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 4 | |
| 14706 | 5 | header {* Binary Words *}
 | 
| 14589 | 6 | |
| 15131 | 7 | theory Word | 
| 33357 | 8 | imports Main | 
| 15131 | 9 | begin | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 10 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 11 | subsection {* Auxilary Lemmas *}
 | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 12 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 13 | lemma max_le [intro!]: "[| x \<le> z; y \<le> z |] ==> max x y \<le> z" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 14 | by (simp add: max_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 15 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 16 | lemma max_mono: | 
| 15067 | 17 | fixes x :: "'a::linorder" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 18 | assumes mf: "mono f" | 
| 15067 | 19 | shows "max (f x) (f y) \<le> f (max x y)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 20 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 21 | from mf and le_maxI1 [of x y] | 
| 23375 | 22 | have fx: "f x \<le> f (max x y)" by (rule monoD) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 23 | from mf and le_maxI2 [of y x] | 
| 23375 | 24 | have fy: "f y \<le> f (max x y)" by (rule monoD) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 25 | from fx and fy | 
| 23375 | 26 | show "max (f x) (f y) \<le> f (max x y)" by auto | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 27 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 28 | |
| 15067 | 29 | declare zero_le_power [intro] | 
| 23375 | 30 | and zero_less_power [intro] | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 31 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 32 | lemma int_nat_two_exp: "2 ^ k = int (2 ^ k)" | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 33 | by (simp add: zpower_int [symmetric]) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 34 | |
| 23375 | 35 | |
| 14589 | 36 | subsection {* Bits *}
 | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 37 | |
| 23375 | 38 | datatype bit = | 
| 39 |     Zero ("\<zero>")
 | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 40 |   | One ("\<one>")
 | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 41 | |
| 30960 | 42 | primrec bitval :: "bit => nat" where | 
| 43 | "bitval \<zero> = 0" | |
| 25961 | 44 | | "bitval \<one> = 1" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 45 | |
| 34941 | 46 | primrec bitnot :: "bit => bit" where | 
| 47 | bitnot_zero: "(bitnot \<zero>) = \<one>" | |
| 48 | | bitnot_one : "(bitnot \<one>) = \<zero>" | |
| 49 | ||
| 50 | primrec bitand :: "bit => bit => bit" (infixr "bitand" 35) where | |
| 51 | bitand_zero: "(\<zero> bitand y) = \<zero>" | |
| 52 | | bitand_one: "(\<one> bitand y) = y" | |
| 53 | ||
| 54 | primrec bitor :: "bit => bit => bit" (infixr "bitor" 30) where | |
| 55 | bitor_zero: "(\<zero> bitor y) = y" | |
| 56 | | bitor_one: "(\<one> bitor y) = \<one>" | |
| 57 | ||
| 58 | primrec bitxor :: "bit => bit => bit" (infixr "bitxor" 30) where | |
| 59 | bitxor_zero: "(\<zero> bitxor y) = y" | |
| 60 | | bitxor_one: "(\<one> bitxor y) = (bitnot y)" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 61 | |
| 21210 | 62 | notation (xsymbols) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 63 |   bitnot ("\<not>\<^sub>b _" [40] 40) and
 | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 64 | bitand (infixr "\<and>\<^sub>b" 35) and | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 65 | bitor (infixr "\<or>\<^sub>b" 30) and | 
| 19736 | 66 | bitxor (infixr "\<oplus>\<^sub>b" 30) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 67 | |
| 21210 | 68 | notation (HTML output) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 69 |   bitnot ("\<not>\<^sub>b _" [40] 40) and
 | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 70 | bitand (infixr "\<and>\<^sub>b" 35) and | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 71 | bitor (infixr "\<or>\<^sub>b" 30) and | 
| 19736 | 72 | bitxor (infixr "\<oplus>\<^sub>b" 30) | 
| 14565 | 73 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 74 | lemma bitnot_bitnot [simp]: "(bitnot (bitnot b)) = b" | 
| 23375 | 75 | by (cases b) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 76 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 77 | lemma bitand_cancel [simp]: "(b bitand b) = b" | 
| 23375 | 78 | by (cases b) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 79 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 80 | lemma bitor_cancel [simp]: "(b bitor b) = b" | 
| 23375 | 81 | by (cases b) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 82 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 83 | lemma bitxor_cancel [simp]: "(b bitxor b) = \<zero>" | 
| 23375 | 84 | by (cases b) simp_all | 
| 85 | ||
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 86 | |
| 14589 | 87 | subsection {* Bit Vectors *}
 | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 88 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 89 | text {* First, a couple of theorems expressing case analysis and
 | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 90 | induction principles for bit vectors. *} | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 91 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 92 | lemma bit_list_cases: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 93 | assumes empty: "w = [] ==> P w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 94 | and zero: "!!bs. w = \<zero> # bs ==> P w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 95 | and one: "!!bs. w = \<one> # bs ==> P w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 96 | shows "P w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 97 | proof (cases w) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 98 | assume "w = []" | 
| 23375 | 99 | thus ?thesis by (rule empty) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 100 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 101 | fix b bs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 102 | assume [simp]: "w = b # bs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 103 | show "P w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 104 | proof (cases b) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 105 | assume "b = \<zero>" | 
| 23375 | 106 | hence "w = \<zero> # bs" by simp | 
| 107 | thus ?thesis by (rule zero) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 108 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 109 | assume "b = \<one>" | 
| 23375 | 110 | hence "w = \<one> # bs" by simp | 
| 111 | thus ?thesis by (rule one) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 112 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 113 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 114 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 115 | lemma bit_list_induct: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 116 | assumes empty: "P []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 117 | and zero: "!!bs. P bs ==> P (\<zero>#bs)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 118 | and one: "!!bs. P bs ==> P (\<one>#bs)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 119 | shows "P w" | 
| 23375 | 120 | proof (induct w, simp_all add: empty) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 121 | fix b bs | 
| 23375 | 122 | assume "P bs" | 
| 123 | then show "P (b#bs)" | |
| 124 | by (cases b) (auto intro!: zero one) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 125 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 126 | |
| 19736 | 127 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 128 | bv_msb :: "bit list => bit" where | 
| 19736 | 129 | "bv_msb w = (if w = [] then \<zero> else hd w)" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 130 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 131 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 132 | bv_extend :: "[nat,bit,bit list]=>bit list" where | 
| 19736 | 133 | "bv_extend i b w = (replicate (i - length w) b) @ w" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 134 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 135 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 136 | bv_not :: "bit list => bit list" where | 
| 19736 | 137 | "bv_not w = map bitnot w" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 138 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 139 | lemma bv_length_extend [simp]: "length w \<le> i ==> length (bv_extend i b w) = i" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 140 | by (simp add: bv_extend_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 141 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 142 | lemma bv_not_Nil [simp]: "bv_not [] = []" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 143 | by (simp add: bv_not_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 144 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 145 | lemma bv_not_Cons [simp]: "bv_not (b#bs) = (bitnot b) # bv_not bs" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 146 | by (simp add: bv_not_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 147 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 148 | lemma bv_not_bv_not [simp]: "bv_not (bv_not w) = w" | 
| 23375 | 149 | by (rule bit_list_induct [of _ w]) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 150 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 151 | lemma bv_msb_Nil [simp]: "bv_msb [] = \<zero>" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 152 | by (simp add: bv_msb_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 153 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 154 | lemma bv_msb_Cons [simp]: "bv_msb (b#bs) = b" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 155 | by (simp add: bv_msb_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 156 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 157 | lemma bv_msb_bv_not [simp]: "0 < length w ==> bv_msb (bv_not w) = (bitnot (bv_msb w))" | 
| 23375 | 158 | by (cases w) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 159 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 160 | lemma bv_msb_one_length [simp,intro]: "bv_msb w = \<one> ==> 0 < length w" | 
| 23375 | 161 | by (cases w) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 162 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 163 | lemma length_bv_not [simp]: "length (bv_not w) = length w" | 
| 23375 | 164 | by (induct w) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 165 | |
| 19736 | 166 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 167 | bv_to_nat :: "bit list => nat" where | 
| 19736 | 168 | "bv_to_nat = foldl (%bn b. 2 * bn + bitval b) 0" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 169 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 170 | lemma bv_to_nat_Nil [simp]: "bv_to_nat [] = 0" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 171 | by (simp add: bv_to_nat_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 172 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 173 | lemma bv_to_nat_helper [simp]: "bv_to_nat (b # bs) = bitval b * 2 ^ length bs + bv_to_nat bs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 174 | proof - | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 175 | let ?bv_to_nat' = "foldl (\<lambda>bn b. 2 * bn + bitval b)" | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 176 | have helper: "\<And>base. ?bv_to_nat' base bs = base * 2 ^ length bs + ?bv_to_nat' 0 bs" | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 177 | proof (induct bs) | 
| 23375 | 178 | case Nil | 
| 179 | show ?case by simp | |
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 180 | next | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 181 | case (Cons x xs base) | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 182 | show ?case | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 183 | apply (simp only: foldl.simps) | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 184 | apply (subst Cons [of "2 * base + bitval x"]) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 185 | apply simp | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 186 | apply (subst Cons [of "bitval x"]) | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 187 | apply (simp add: add_mult_distrib) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 188 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 189 | qed | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 190 | show ?thesis by (simp add: bv_to_nat_def) (rule helper) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 191 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 192 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 193 | lemma bv_to_nat0 [simp]: "bv_to_nat (\<zero>#bs) = bv_to_nat bs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 194 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 195 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 196 | lemma bv_to_nat1 [simp]: "bv_to_nat (\<one>#bs) = 2 ^ length bs + bv_to_nat bs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 197 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 198 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 199 | lemma bv_to_nat_upper_range: "bv_to_nat w < 2 ^ length w" | 
| 23375 | 200 | proof (induct w, simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 201 | fix b bs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 202 | assume "bv_to_nat bs < 2 ^ length bs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 203 | show "bitval b * 2 ^ length bs + bv_to_nat bs < 2 * 2 ^ length bs" | 
| 23375 | 204 | proof (cases b, simp_all) | 
| 205 | have "bv_to_nat bs < 2 ^ length bs" by fact | |
| 206 | also have "... < 2 * 2 ^ length bs" by auto | |
| 207 | finally show "bv_to_nat bs < 2 * 2 ^ length bs" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 208 | next | 
| 23375 | 209 | have "bv_to_nat bs < 2 ^ length bs" by fact | 
| 210 | hence "2 ^ length bs + bv_to_nat bs < 2 ^ length bs + 2 ^ length bs" by arith | |
| 211 | also have "... = 2 * (2 ^ length bs)" by simp | |
| 212 | finally show "bv_to_nat bs < 2 ^ length bs" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 213 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 214 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 215 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 216 | lemma bv_extend_longer [simp]: | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 217 | assumes wn: "n \<le> length w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 218 | shows "bv_extend n b w = w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 219 | by (simp add: bv_extend_def wn) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 220 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 221 | lemma bv_extend_shorter [simp]: | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 222 | assumes wn: "length w < n" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 223 | shows "bv_extend n b w = bv_extend n b (b#w)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 224 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 225 | from wn | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 226 | have s: "n - Suc (length w) + 1 = n - length w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 227 | by arith | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 228 | have "bv_extend n b w = replicate (n - length w) b @ w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 229 | by (simp add: bv_extend_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 230 | also have "... = replicate (n - Suc (length w) + 1) b @ w" | 
| 23375 | 231 | by (subst s) rule | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 232 | also have "... = (replicate (n - Suc (length w)) b @ replicate 1 b) @ w" | 
| 23375 | 233 | by (subst replicate_add) rule | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 234 | also have "... = replicate (n - Suc (length w)) b @ b # w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 235 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 236 | also have "... = bv_extend n b (b#w)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 237 | by (simp add: bv_extend_def) | 
| 23375 | 238 | finally show "bv_extend n b w = bv_extend n b (b#w)" . | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 239 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 240 | |
| 34941 | 241 | primrec rem_initial :: "bit => bit list => bit list" where | 
| 242 | "rem_initial b [] = []" | |
| 243 | | "rem_initial b (x#xs) = (if b = x then rem_initial b xs else x#xs)" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 244 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 245 | lemma rem_initial_length: "length (rem_initial b w) \<le> length w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 246 | by (rule bit_list_induct [of _ w],simp_all (no_asm),safe,simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 247 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 248 | lemma rem_initial_equal: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 249 | assumes p: "length (rem_initial b w) = length w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 250 | shows "rem_initial b w = w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 251 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 252 | have "length (rem_initial b w) = length w --> rem_initial b w = w" | 
| 23375 | 253 | proof (induct w, simp_all, clarify) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 254 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 255 | assume "length (rem_initial b xs) = length xs --> rem_initial b xs = xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 256 | assume f: "length (rem_initial b xs) = Suc (length xs)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 257 | with rem_initial_length [of b xs] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 258 | show "rem_initial b xs = b#xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 259 | by auto | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 260 | qed | 
| 23375 | 261 | from this and p show ?thesis .. | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 262 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 263 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 264 | lemma bv_extend_rem_initial: "bv_extend (length w) b (rem_initial b w) = w" | 
| 23375 | 265 | proof (induct w, simp_all, safe) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 266 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 267 | assume ind: "bv_extend (length xs) b (rem_initial b xs) = xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 268 | from rem_initial_length [of b xs] | 
| 23375 | 269 | have [simp]: "Suc (length xs) - length (rem_initial b xs) = | 
| 270 | 1 + (length xs - length (rem_initial b xs))" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 271 | by arith | 
| 23375 | 272 | have "bv_extend (Suc (length xs)) b (rem_initial b xs) = | 
| 273 | replicate (Suc (length xs) - length (rem_initial b xs)) b @ (rem_initial b xs)" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 274 | by (simp add: bv_extend_def) | 
| 23375 | 275 | also have "... = | 
| 276 | replicate (1 + (length xs - length (rem_initial b xs))) b @ rem_initial b xs" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 277 | by simp | 
| 23375 | 278 | also have "... = | 
| 279 | (replicate 1 b @ replicate (length xs - length (rem_initial b xs)) b) @ rem_initial b xs" | |
| 280 | by (subst replicate_add) (rule refl) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 281 | also have "... = b # bv_extend (length xs) b (rem_initial b xs)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 282 | by (auto simp add: bv_extend_def [symmetric]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 283 | also have "... = b # xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 284 | by (simp add: ind) | 
| 23375 | 285 | finally show "bv_extend (Suc (length xs)) b (rem_initial b xs) = b # xs" . | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 286 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 287 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 288 | lemma rem_initial_append1: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 289 | assumes "rem_initial b xs ~= []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 290 | shows "rem_initial b (xs @ ys) = rem_initial b xs @ ys" | 
| 23375 | 291 | using assms by (induct xs) auto | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 292 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 293 | lemma rem_initial_append2: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 294 | assumes "rem_initial b xs = []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 295 | shows "rem_initial b (xs @ ys) = rem_initial b ys" | 
| 23375 | 296 | using assms by (induct xs) auto | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 297 | |
| 19736 | 298 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 299 | norm_unsigned :: "bit list => bit list" where | 
| 19736 | 300 | "norm_unsigned = rem_initial \<zero>" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 301 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 302 | lemma norm_unsigned_Nil [simp]: "norm_unsigned [] = []" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 303 | by (simp add: norm_unsigned_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 304 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 305 | lemma norm_unsigned_Cons0 [simp]: "norm_unsigned (\<zero>#bs) = norm_unsigned bs" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 306 | by (simp add: norm_unsigned_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 307 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 308 | lemma norm_unsigned_Cons1 [simp]: "norm_unsigned (\<one>#bs) = \<one>#bs" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 309 | by (simp add: norm_unsigned_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 310 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 311 | lemma norm_unsigned_idem [simp]: "norm_unsigned (norm_unsigned w) = norm_unsigned w" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 312 | by (rule bit_list_induct [of _ w],simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 313 | |
| 35440 | 314 | fun | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 315 | nat_to_bv_helper :: "nat => bit list => bit list" | 
| 35440 | 316 | where | 
| 317 | "nat_to_bv_helper n bs = (if n = 0 then bs | |
| 318 | else nat_to_bv_helper (n div 2) ((if n mod 2 = 0 then \<zero> else \<one>)#bs))" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 319 | |
| 19736 | 320 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 321 | nat_to_bv :: "nat => bit list" where | 
| 19736 | 322 | "nat_to_bv n = nat_to_bv_helper n []" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 323 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 324 | lemma nat_to_bv0 [simp]: "nat_to_bv 0 = []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 325 | by (simp add: nat_to_bv_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 326 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 327 | lemmas [simp del] = nat_to_bv_helper.simps | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 328 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 329 | lemma n_div_2_cases: | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 330 | assumes zero: "(n::nat) = 0 ==> R" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 331 | and div : "[| n div 2 < n ; 0 < n |] ==> R" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 332 | shows "R" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 333 | proof (cases "n = 0") | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 334 | assume "n = 0" | 
| 23375 | 335 | thus R by (rule zero) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 336 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 337 | assume "n ~= 0" | 
| 23375 | 338 | hence "0 < n" by simp | 
| 339 | hence "n div 2 < n" by arith | |
| 340 | from this and `0 < n` show R by (rule div) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 341 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 342 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 343 | lemma int_wf_ge_induct: | 
| 22059 | 344 | assumes ind : "!!i::int. (!!j. [| k \<le> j ; j < i |] ==> P j) ==> P i" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 345 | shows "P i" | 
| 22059 | 346 | proof (rule wf_induct_rule [OF wf_int_ge_less_than]) | 
| 347 | fix x | |
| 348 | assume ih: "(\<And>y\<Colon>int. (y, x) \<in> int_ge_less_than k \<Longrightarrow> P y)" | |
| 349 | thus "P x" | |
| 23375 | 350 | by (rule ind) (simp add: int_ge_less_than_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 351 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 352 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 353 | lemma unfold_nat_to_bv_helper: | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 354 | "nat_to_bv_helper b l = nat_to_bv_helper b [] @ l" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 355 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 356 | have "\<forall>l. nat_to_bv_helper b l = nat_to_bv_helper b [] @ l" | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 357 | proof (induct b rule: less_induct) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 358 | fix n | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 359 | assume ind: "!!j. j < n \<Longrightarrow> \<forall> l. nat_to_bv_helper j l = nat_to_bv_helper j [] @ l" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 360 | show "\<forall>l. nat_to_bv_helper n l = nat_to_bv_helper n [] @ l" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 361 | proof | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 362 | fix l | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 363 | show "nat_to_bv_helper n l = nat_to_bv_helper n [] @ l" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 364 | proof (cases "n < 0") | 
| 19736 | 365 | assume "n < 0" | 
| 366 | thus ?thesis | |
| 367 | by (simp add: nat_to_bv_helper.simps) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 368 | next | 
| 19736 | 369 | assume "~n < 0" | 
| 370 | show ?thesis | |
| 371 | proof (rule n_div_2_cases [of n]) | |
| 372 | assume [simp]: "n = 0" | |
| 373 | show ?thesis | |
| 374 | apply (simp only: nat_to_bv_helper.simps [of n]) | |
| 375 | apply simp | |
| 376 | done | |
| 377 | next | |
| 378 | assume n2n: "n div 2 < n" | |
| 379 | assume [simp]: "0 < n" | |
| 380 | hence n20: "0 \<le> n div 2" | |
| 381 | by arith | |
| 382 | from ind [of "n div 2"] and n2n n20 | |
| 383 | have ind': "\<forall>l. nat_to_bv_helper (n div 2) l = nat_to_bv_helper (n div 2) [] @ l" | |
| 384 | by blast | |
| 385 | show ?thesis | |
| 386 | apply (simp only: nat_to_bv_helper.simps [of n]) | |
| 387 | apply (cases "n=0") | |
| 388 | apply simp | |
| 389 | apply (simp only: if_False) | |
| 390 | apply simp | |
| 391 | apply (subst spec [OF ind',of "\<zero>#l"]) | |
| 392 | apply (subst spec [OF ind',of "\<one>#l"]) | |
| 393 | apply (subst spec [OF ind',of "[\<one>]"]) | |
| 394 | apply (subst spec [OF ind',of "[\<zero>]"]) | |
| 395 | apply simp | |
| 396 | done | |
| 397 | qed | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 398 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 399 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 400 | qed | 
| 23375 | 401 | thus ?thesis .. | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 402 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 403 | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 404 | lemma nat_to_bv_non0 [simp]: "n\<noteq>0 ==> nat_to_bv n = nat_to_bv (n div 2) @ [if n mod 2 = 0 then \<zero> else \<one>]" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 405 | proof - | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 406 | assume [simp]: "n\<noteq>0" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 407 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 408 | apply (subst nat_to_bv_def [of n]) | 
| 15481 | 409 | apply (simp only: nat_to_bv_helper.simps [of n]) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 410 | apply (subst unfold_nat_to_bv_helper) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 411 | using prems | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 412 | apply (simp) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 413 | apply (subst nat_to_bv_def [of "n div 2"]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 414 | apply auto | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 415 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 416 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 417 | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 418 | lemma bv_to_nat_dist_append: | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 419 | "bv_to_nat (l1 @ l2) = bv_to_nat l1 * 2 ^ length l2 + bv_to_nat l2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 420 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 421 | have "\<forall>l2. bv_to_nat (l1 @ l2) = bv_to_nat l1 * 2 ^ length l2 + bv_to_nat l2" | 
| 27368 | 422 | proof (induct l1, simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 423 | fix x xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 424 | assume ind: "\<forall>l2. bv_to_nat (xs @ l2) = bv_to_nat xs * 2 ^ length l2 + bv_to_nat l2" | 
| 27368 | 425 | show "\<forall>l2::bit list. bitval x * 2 ^ (length xs + length l2) + bv_to_nat xs * 2 ^ length l2 = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 426 | proof | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 427 | fix l2 | 
| 27368 | 428 | show "bitval x * 2 ^ (length xs + length l2) + bv_to_nat xs * 2 ^ length l2 = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 429 | proof - | 
| 19736 | 430 | have "(2::nat) ^ (length xs + length l2) = 2 ^ length xs * 2 ^ length l2" | 
| 34915 | 431 |           by (induct ("length xs")) simp_all
 | 
| 19736 | 432 | hence "bitval x * 2 ^ (length xs + length l2) + bv_to_nat xs * 2 ^ length l2 = | 
| 433 | bitval x * 2 ^ length xs * 2 ^ length l2 + bv_to_nat xs * 2 ^ length l2" | |
| 434 | by simp | |
| 435 | also have "... = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2" | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23431diff
changeset | 436 | by (simp add: ring_distribs) | 
| 25595 | 437 | finally show ?thesis by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 438 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 439 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 440 | qed | 
| 23375 | 441 | thus ?thesis .. | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 442 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 443 | |
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 444 | lemma bv_nat_bv [simp]: "bv_to_nat (nat_to_bv n) = n" | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 445 | proof (induct n rule: less_induct) | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 446 | fix n | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 447 | assume ind: "!!j. j < n \<Longrightarrow> bv_to_nat (nat_to_bv j) = j" | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 448 | show "bv_to_nat (nat_to_bv n) = n" | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 449 | proof (rule n_div_2_cases [of n]) | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 450 | assume "n = 0" then show ?thesis by simp | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 451 | next | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 452 | assume nn: "n div 2 < n" | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 453 | assume n0: "0 < n" | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 454 | from ind and nn | 
| 23375 | 455 | have ind': "bv_to_nat (nat_to_bv (n div 2)) = n div 2" by blast | 
| 456 | from n0 have n0': "n \<noteq> 0" by simp | |
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 457 | show ?thesis | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 458 | apply (subst nat_to_bv_def) | 
| 15481 | 459 | apply (simp only: nat_to_bv_helper.simps [of n]) | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 460 | apply (simp only: n0' if_False) | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 461 | apply (subst unfold_nat_to_bv_helper) | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 462 | apply (subst bv_to_nat_dist_append) | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 463 | apply (fold nat_to_bv_def) | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 464 | apply (simp add: ind' split del: split_if) | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 465 | apply (cases "n mod 2 = 0") | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 466 | proof (simp_all) | 
| 19736 | 467 | assume "n mod 2 = 0" | 
| 468 | with mod_div_equality [of n 2] | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 469 | show "n div 2 * 2 = n" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 470 | next | 
| 25162 | 471 | assume "n mod 2 = Suc 0" | 
| 19736 | 472 | with mod_div_equality [of n 2] | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 473 | show "Suc (n div 2 * 2) = n" by arith | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 474 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 475 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 476 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 477 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 478 | lemma bv_to_nat_type [simp]: "bv_to_nat (norm_unsigned w) = bv_to_nat w" | 
| 23375 | 479 | by (rule bit_list_induct) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 480 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 481 | lemma length_norm_unsigned_le [simp]: "length (norm_unsigned w) \<le> length w" | 
| 23375 | 482 | by (rule bit_list_induct) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 483 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 484 | lemma bv_to_nat_rew_msb: "bv_msb w = \<one> ==> bv_to_nat w = 2 ^ (length w - 1) + bv_to_nat (tl w)" | 
| 23375 | 485 | by (rule bit_list_cases [of w]) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 486 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 487 | lemma norm_unsigned_result: "norm_unsigned xs = [] \<or> bv_msb (norm_unsigned xs) = \<one>" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 488 | proof (rule length_induct [of _ xs]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 489 | fix xs :: "bit list" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 490 | assume ind: "\<forall>ys. length ys < length xs --> norm_unsigned ys = [] \<or> bv_msb (norm_unsigned ys) = \<one>" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 491 | show "norm_unsigned xs = [] \<or> bv_msb (norm_unsigned xs) = \<one>" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 492 | proof (rule bit_list_cases [of xs],simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 493 | fix bs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 494 | assume [simp]: "xs = \<zero>#bs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 495 | from ind | 
| 23375 | 496 | have "length bs < length xs --> norm_unsigned bs = [] \<or> bv_msb (norm_unsigned bs) = \<one>" .. | 
| 497 | thus "norm_unsigned bs = [] \<or> bv_msb (norm_unsigned bs) = \<one>" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 498 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 499 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 500 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 501 | lemma norm_empty_bv_to_nat_zero: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 502 | assumes nw: "norm_unsigned w = []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 503 | shows "bv_to_nat w = 0" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 504 | proof - | 
| 23375 | 505 | have "bv_to_nat w = bv_to_nat (norm_unsigned w)" by simp | 
| 506 | also have "... = bv_to_nat []" by (subst nw) (rule refl) | |
| 507 | also have "... = 0" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 508 | finally show ?thesis . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 509 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 510 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 511 | lemma bv_to_nat_lower_limit: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 512 | assumes w0: "0 < bv_to_nat w" | 
| 23375 | 513 | shows "2 ^ (length (norm_unsigned w) - 1) \<le> bv_to_nat w" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 514 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 515 | from w0 and norm_unsigned_result [of w] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 516 | have msbw: "bv_msb (norm_unsigned w) = \<one>" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 517 | by (auto simp add: norm_empty_bv_to_nat_zero) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 518 | have "2 ^ (length (norm_unsigned w) - 1) \<le> bv_to_nat (norm_unsigned w)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 519 | by (subst bv_to_nat_rew_msb [OF msbw],simp) | 
| 23375 | 520 | thus ?thesis by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 521 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 522 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 523 | lemmas [simp del] = nat_to_bv_non0 | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 524 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 525 | lemma norm_unsigned_length [intro!]: "length (norm_unsigned w) \<le> length w" | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 526 | by (subst norm_unsigned_def,rule rem_initial_length) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 527 | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 528 | lemma norm_unsigned_equal: | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 529 | "length (norm_unsigned w) = length w ==> norm_unsigned w = w" | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 530 | by (simp add: norm_unsigned_def,rule rem_initial_equal) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 531 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 532 | lemma bv_extend_norm_unsigned: "bv_extend (length w) \<zero> (norm_unsigned w) = w" | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 533 | by (simp add: norm_unsigned_def,rule bv_extend_rem_initial) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 534 | |
| 23375 | 535 | lemma norm_unsigned_append1 [simp]: | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 536 | "norm_unsigned xs \<noteq> [] ==> norm_unsigned (xs @ ys) = norm_unsigned xs @ ys" | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 537 | by (simp add: norm_unsigned_def,rule rem_initial_append1) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 538 | |
| 23375 | 539 | lemma norm_unsigned_append2 [simp]: | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 540 | "norm_unsigned xs = [] ==> norm_unsigned (xs @ ys) = norm_unsigned ys" | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 541 | by (simp add: norm_unsigned_def,rule rem_initial_append2) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 542 | |
| 23375 | 543 | lemma bv_to_nat_zero_imp_empty: | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 544 | "bv_to_nat w = 0 \<Longrightarrow> norm_unsigned w = []" | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 545 | by (atomize (full), induct w rule: bit_list_induct) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 546 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 547 | lemma bv_to_nat_nzero_imp_nempty: | 
| 23375 | 548 | "bv_to_nat w \<noteq> 0 \<Longrightarrow> norm_unsigned w \<noteq> []" | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 549 | by (induct w rule: bit_list_induct) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 550 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 551 | lemma nat_helper1: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 552 | assumes ass: "nat_to_bv (bv_to_nat w) = norm_unsigned w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 553 | shows "nat_to_bv (2 * bv_to_nat w + bitval x) = norm_unsigned (w @ [x])" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 554 | proof (cases x) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 555 | assume [simp]: "x = \<one>" | 
| 33357 | 556 | have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\<one>] = | 
| 557 | nat_to_bv ((1 + 2 * bv_to_nat w) div 2) @ [\<one>]" | |
| 558 | by (simp add: add_commute) | |
| 559 | also have "... = nat_to_bv (bv_to_nat w) @ [\<one>]" | |
| 560 | by (subst div_add1_eq) simp | |
| 561 | also have "... = norm_unsigned w @ [\<one>]" | |
| 562 | by (subst ass) (rule refl) | |
| 563 | also have "... = norm_unsigned (w @ [\<one>])" | |
| 564 | by (cases "norm_unsigned w") simp_all | |
| 565 | finally have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\<one>] = norm_unsigned (w @ [\<one>])" . | |
| 566 | then show ?thesis by (simp add: nat_to_bv_non0) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 567 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 568 | assume [simp]: "x = \<zero>" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 569 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 570 | proof (cases "bv_to_nat w = 0") | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 571 | assume "bv_to_nat w = 0" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 572 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 573 | by (simp add: bv_to_nat_zero_imp_empty) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 574 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 575 | assume "bv_to_nat w \<noteq> 0" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 576 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 577 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 578 | apply (subst nat_to_bv_non0) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 579 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 580 | apply auto | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 581 | apply (subst ass) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 582 | apply (cases "norm_unsigned w") | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 583 | apply (simp_all add: norm_empty_bv_to_nat_zero) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 584 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 585 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 586 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 587 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 588 | lemma nat_helper2: "nat_to_bv (2 ^ length xs + bv_to_nat xs) = \<one> # xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 589 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 590 | have "\<forall>xs. nat_to_bv (2 ^ length (rev xs) + bv_to_nat (rev xs)) = \<one> # (rev xs)" (is "\<forall>xs. ?P xs") | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 591 | proof | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 592 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 593 | show "?P xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 594 | proof (rule length_induct [of _ xs]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 595 | fix xs :: "bit list" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 596 | assume ind: "\<forall>ys. length ys < length xs --> ?P ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 597 | show "?P xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 598 | proof (cases xs) | 
| 23375 | 599 | assume "xs = []" | 
| 600 | then show ?thesis by (simp add: nat_to_bv_non0) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 601 | next | 
| 19736 | 602 | fix y ys | 
| 603 | assume [simp]: "xs = y # ys" | |
| 604 | show ?thesis | |
| 605 | apply simp | |
| 606 | apply (subst bv_to_nat_dist_append) | |
| 607 | apply simp | |
| 608 | proof - | |
| 609 | have "nat_to_bv (2 * 2 ^ length ys + (bv_to_nat (rev ys) * 2 + bitval y)) = | |
| 610 | nat_to_bv (2 * (2 ^ length ys + bv_to_nat (rev ys)) + bitval y)" | |
| 611 | by (simp add: add_ac mult_ac) | |
| 612 | also have "... = nat_to_bv (2 * (bv_to_nat (\<one>#rev ys)) + bitval y)" | |
| 613 | by simp | |
| 614 | also have "... = norm_unsigned (\<one>#rev ys) @ [y]" | |
| 615 | proof - | |
| 616 | from ind | |
| 617 | have "nat_to_bv (2 ^ length (rev ys) + bv_to_nat (rev ys)) = \<one> # rev ys" | |
| 618 | by auto | |
| 619 | hence [simp]: "nat_to_bv (2 ^ length ys + bv_to_nat (rev ys)) = \<one> # rev ys" | |
| 620 | by simp | |
| 621 | show ?thesis | |
| 622 | apply (subst nat_helper1) | |
| 623 | apply simp_all | |
| 624 | done | |
| 625 | qed | |
| 626 | also have "... = (\<one>#rev ys) @ [y]" | |
| 627 | by simp | |
| 628 | also have "... = \<one> # rev ys @ [y]" | |
| 629 | by simp | |
| 23375 | 630 | finally show "nat_to_bv (2 * 2 ^ length ys + (bv_to_nat (rev ys) * 2 + bitval y)) = | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32456diff
changeset | 631 | \<one> # rev ys @ [y]" . | 
| 19736 | 632 | qed | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 633 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 634 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 635 | qed | 
| 23375 | 636 | hence "nat_to_bv (2 ^ length (rev (rev xs)) + bv_to_nat (rev (rev xs))) = | 
| 637 | \<one> # rev (rev xs)" .. | |
| 638 | thus ?thesis by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 639 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 640 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 641 | lemma nat_bv_nat [simp]: "nat_to_bv (bv_to_nat w) = norm_unsigned w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 642 | proof (rule bit_list_induct [of _ w],simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 643 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 644 | assume "nat_to_bv (bv_to_nat xs) = norm_unsigned xs" | 
| 23375 | 645 | have "bv_to_nat xs = bv_to_nat (norm_unsigned xs)" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 646 | have "bv_to_nat xs < 2 ^ length xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 647 | by (rule bv_to_nat_upper_range) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 648 | show "nat_to_bv (2 ^ length xs + bv_to_nat xs) = \<one> # xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 649 | by (rule nat_helper2) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 650 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 651 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 652 | lemma bv_to_nat_qinj: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 653 | assumes one: "bv_to_nat xs = bv_to_nat ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 654 | and len: "length xs = length ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 655 | shows "xs = ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 656 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 657 | from one | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 658 | have "nat_to_bv (bv_to_nat xs) = nat_to_bv (bv_to_nat ys)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 659 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 660 | hence xsys: "norm_unsigned xs = norm_unsigned ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 661 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 662 | have "xs = bv_extend (length xs) \<zero> (norm_unsigned xs)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 663 | by (simp add: bv_extend_norm_unsigned) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 664 | also have "... = bv_extend (length ys) \<zero> (norm_unsigned ys)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 665 | by (simp add: xsys len) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 666 | also have "... = ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 667 | by (simp add: bv_extend_norm_unsigned) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 668 | finally show ?thesis . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 669 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 670 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 671 | lemma norm_unsigned_nat_to_bv [simp]: | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 672 | "norm_unsigned (nat_to_bv n) = nat_to_bv n" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 673 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 674 | have "norm_unsigned (nat_to_bv n) = nat_to_bv (bv_to_nat (norm_unsigned (nat_to_bv n)))" | 
| 23375 | 675 | by (subst nat_bv_nat) simp | 
| 676 | also have "... = nat_to_bv n" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 677 | finally show ?thesis . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 678 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 679 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 680 | lemma length_nat_to_bv_upper_limit: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 681 | assumes nk: "n \<le> 2 ^ k - 1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 682 | shows "length (nat_to_bv n) \<le> k" | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 683 | proof (cases "n = 0") | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 684 | case True | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 685 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 686 | by (simp add: nat_to_bv_def nat_to_bv_helper.simps) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 687 | next | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 688 | case False | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 689 | hence n0: "0 < n" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 690 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 691 | proof (rule ccontr) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 692 | assume "~ length (nat_to_bv n) \<le> k" | 
| 23375 | 693 | hence "k < length (nat_to_bv n)" by simp | 
| 694 | hence "k \<le> length (nat_to_bv n) - 1" by arith | |
| 695 | hence "(2::nat) ^ k \<le> 2 ^ (length (nat_to_bv n) - 1)" by simp | |
| 696 | also have "... = 2 ^ (length (norm_unsigned (nat_to_bv n)) - 1)" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 697 | also have "... \<le> bv_to_nat (nat_to_bv n)" | 
| 23375 | 698 | by (rule bv_to_nat_lower_limit) (simp add: n0) | 
| 699 | also have "... = n" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 700 | finally have "2 ^ k \<le> n" . | 
| 23375 | 701 | with n0 have "2 ^ k - 1 < n" by arith | 
| 702 | with nk show False by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 703 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 704 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 705 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 706 | lemma length_nat_to_bv_lower_limit: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 707 | assumes nk: "2 ^ k \<le> n" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 708 | shows "k < length (nat_to_bv n)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 709 | proof (rule ccontr) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 710 | assume "~ k < length (nat_to_bv n)" | 
| 23375 | 711 | hence lnk: "length (nat_to_bv n) \<le> k" by simp | 
| 712 | have "n = bv_to_nat (nat_to_bv n)" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 713 | also have "... < 2 ^ length (nat_to_bv n)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 714 | by (rule bv_to_nat_upper_range) | 
| 23375 | 715 | also from lnk have "... \<le> 2 ^ k" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 716 | finally have "n < 2 ^ k" . | 
| 23375 | 717 | with nk show False by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 718 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 719 | |
| 23375 | 720 | |
| 14589 | 721 | subsection {* Unsigned Arithmetic Operations *}
 | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 722 | |
| 19736 | 723 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 724 | bv_add :: "[bit list, bit list ] => bit list" where | 
| 19736 | 725 | "bv_add w1 w2 = nat_to_bv (bv_to_nat w1 + bv_to_nat w2)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 726 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 727 | lemma bv_add_type1 [simp]: "bv_add (norm_unsigned w1) w2 = bv_add w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 728 | by (simp add: bv_add_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 729 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 730 | lemma bv_add_type2 [simp]: "bv_add w1 (norm_unsigned w2) = bv_add w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 731 | by (simp add: bv_add_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 732 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 733 | lemma bv_add_returntype [simp]: "norm_unsigned (bv_add w1 w2) = bv_add w1 w2" | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 734 | by (simp add: bv_add_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 735 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 736 | lemma bv_add_length: "length (bv_add w1 w2) \<le> Suc (max (length w1) (length w2))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 737 | proof (unfold bv_add_def,rule length_nat_to_bv_upper_limit) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 738 | from bv_to_nat_upper_range [of w1] and bv_to_nat_upper_range [of w2] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 739 | have "bv_to_nat w1 + bv_to_nat w2 \<le> (2 ^ length w1 - 1) + (2 ^ length w2 - 1)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 740 | by arith | 
| 23375 | 741 | also have "... \<le> | 
| 742 | max (2 ^ length w1 - 1) (2 ^ length w2 - 1) + max (2 ^ length w1 - 1) (2 ^ length w2 - 1)" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 743 | by (rule add_mono,safe intro!: le_maxI1 le_maxI2) | 
| 23375 | 744 | also have "... = 2 * max (2 ^ length w1 - 1) (2 ^ length w2 - 1)" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 745 | also have "... \<le> 2 ^ Suc (max (length w1) (length w2)) - 2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 746 | proof (cases "length w1 \<le> length w2") | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 747 | assume w1w2: "length w1 \<le> length w2" | 
| 23375 | 748 | hence "(2::nat) ^ length w1 \<le> 2 ^ length w2" by simp | 
| 749 | hence "(2::nat) ^ length w1 - 1 \<le> 2 ^ length w2 - 1" by arith | |
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 750 | with w1w2 show ?thesis | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 751 | by (simp add: diff_mult_distrib2 split: split_max) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 752 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 753 | assume [simp]: "~ (length w1 \<le> length w2)" | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 754 | have "~ ((2::nat) ^ length w1 - 1 \<le> 2 ^ length w2 - 1)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 755 | proof | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 756 | assume "(2::nat) ^ length w1 - 1 \<le> 2 ^ length w2 - 1" | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 757 | hence "((2::nat) ^ length w1 - 1) + 1 \<le> (2 ^ length w2 - 1) + 1" | 
| 19736 | 758 | by (rule add_right_mono) | 
| 23375 | 759 | hence "(2::nat) ^ length w1 \<le> 2 ^ length w2" by simp | 
| 760 | hence "length w1 \<le> length w2" by simp | |
| 761 | thus False by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 762 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 763 | thus ?thesis | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 764 | by (simp add: diff_mult_distrib2 split: split_max) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 765 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 766 | finally show "bv_to_nat w1 + bv_to_nat w2 \<le> 2 ^ Suc (max (length w1) (length w2)) - 1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 767 | by arith | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 768 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 769 | |
| 19736 | 770 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 771 | bv_mult :: "[bit list, bit list ] => bit list" where | 
| 19736 | 772 | "bv_mult w1 w2 = nat_to_bv (bv_to_nat w1 * bv_to_nat w2)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 773 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 774 | lemma bv_mult_type1 [simp]: "bv_mult (norm_unsigned w1) w2 = bv_mult w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 775 | by (simp add: bv_mult_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 776 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 777 | lemma bv_mult_type2 [simp]: "bv_mult w1 (norm_unsigned w2) = bv_mult w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 778 | by (simp add: bv_mult_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 779 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 780 | lemma bv_mult_returntype [simp]: "norm_unsigned (bv_mult w1 w2) = bv_mult w1 w2" | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 781 | by (simp add: bv_mult_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 782 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 783 | lemma bv_mult_length: "length (bv_mult w1 w2) \<le> length w1 + length w2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 784 | proof (unfold bv_mult_def,rule length_nat_to_bv_upper_limit) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 785 | from bv_to_nat_upper_range [of w1] and bv_to_nat_upper_range [of w2] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 786 | have h: "bv_to_nat w1 \<le> 2 ^ length w1 - 1 \<and> bv_to_nat w2 \<le> 2 ^ length w2 - 1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 787 | by arith | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 788 | have "bv_to_nat w1 * bv_to_nat w2 \<le> (2 ^ length w1 - 1) * (2 ^ length w2 - 1)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 789 | apply (cut_tac h) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 790 | apply (rule mult_mono) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 791 | apply auto | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 792 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 793 | also have "... < 2 ^ length w1 * 2 ^ length w2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 794 | by (rule mult_strict_mono,auto) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 795 | also have "... = 2 ^ (length w1 + length w2)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 796 | by (simp add: power_add) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 797 | finally show "bv_to_nat w1 * bv_to_nat w2 \<le> 2 ^ (length w1 + length w2) - 1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 798 | by arith | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 799 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 800 | |
| 14589 | 801 | subsection {* Signed Vectors *}
 | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 802 | |
| 34941 | 803 | primrec norm_signed :: "bit list => bit list" where | 
| 804 | norm_signed_Nil: "norm_signed [] = []" | |
| 805 | | norm_signed_Cons: "norm_signed (b#bs) = | |
| 806 | (case b of | |
| 807 | \<zero> => if norm_unsigned bs = [] then [] else b#norm_unsigned bs | |
| 808 | | \<one> => b#rem_initial b bs)" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 809 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 810 | lemma norm_signed0 [simp]: "norm_signed [\<zero>] = []" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 811 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 812 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 813 | lemma norm_signed1 [simp]: "norm_signed [\<one>] = [\<one>]" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 814 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 815 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 816 | lemma norm_signed01 [simp]: "norm_signed (\<zero>#\<one>#xs) = \<zero>#\<one>#xs" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 817 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 818 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 819 | lemma norm_signed00 [simp]: "norm_signed (\<zero>#\<zero>#xs) = norm_signed (\<zero>#xs)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 820 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 821 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 822 | lemma norm_signed10 [simp]: "norm_signed (\<one>#\<zero>#xs) = \<one>#\<zero>#xs" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 823 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 824 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 825 | lemma norm_signed11 [simp]: "norm_signed (\<one>#\<one>#xs) = norm_signed (\<one>#xs)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 826 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 827 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 828 | lemmas [simp del] = norm_signed_Cons | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 829 | |
| 19736 | 830 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 831 | int_to_bv :: "int => bit list" where | 
| 19736 | 832 | "int_to_bv n = (if 0 \<le> n | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 833 | then norm_signed (\<zero>#nat_to_bv (nat n)) | 
| 19736 | 834 | else norm_signed (bv_not (\<zero>#nat_to_bv (nat (-n- 1)))))" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 835 | |
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 836 | lemma int_to_bv_ge0 [simp]: "0 \<le> n ==> int_to_bv n = norm_signed (\<zero> # nat_to_bv (nat n))" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 837 | by (simp add: int_to_bv_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 838 | |
| 23375 | 839 | lemma int_to_bv_lt0 [simp]: | 
| 840 | "n < 0 ==> int_to_bv n = norm_signed (bv_not (\<zero>#nat_to_bv (nat (-n- 1))))" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 841 | by (simp add: int_to_bv_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 842 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 843 | lemma norm_signed_idem [simp]: "norm_signed (norm_signed w) = norm_signed w" | 
| 23375 | 844 | proof (rule bit_list_induct [of _ w], simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 845 | fix xs | 
| 23375 | 846 | assume eq: "norm_signed (norm_signed xs) = norm_signed xs" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 847 | show "norm_signed (norm_signed (\<zero>#xs)) = norm_signed (\<zero>#xs)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 848 | proof (rule bit_list_cases [of xs],simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 849 | fix ys | 
| 23375 | 850 | assume "xs = \<zero>#ys" | 
| 851 | from this [symmetric] and eq | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 852 | show "norm_signed (norm_signed (\<zero>#ys)) = norm_signed (\<zero>#ys)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 853 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 854 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 855 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 856 | fix xs | 
| 23375 | 857 | assume eq: "norm_signed (norm_signed xs) = norm_signed xs" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 858 | show "norm_signed (norm_signed (\<one>#xs)) = norm_signed (\<one>#xs)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 859 | proof (rule bit_list_cases [of xs],simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 860 | fix ys | 
| 23375 | 861 | assume "xs = \<one>#ys" | 
| 862 | from this [symmetric] and eq | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 863 | show "norm_signed (norm_signed (\<one>#ys)) = norm_signed (\<one>#ys)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 864 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 865 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 866 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 867 | |
| 19736 | 868 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 869 | bv_to_int :: "bit list => int" where | 
| 19736 | 870 | "bv_to_int w = | 
| 871 | (case bv_msb w of \<zero> => int (bv_to_nat w) | |
| 872 | | \<one> => - int (bv_to_nat (bv_not w) + 1))" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 873 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 874 | lemma bv_to_int_Nil [simp]: "bv_to_int [] = 0" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 875 | by (simp add: bv_to_int_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 876 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 877 | lemma bv_to_int_Cons0 [simp]: "bv_to_int (\<zero>#bs) = int (bv_to_nat bs)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 878 | by (simp add: bv_to_int_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 879 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 880 | lemma bv_to_int_Cons1 [simp]: "bv_to_int (\<one>#bs) = - int (bv_to_nat (bv_not bs) + 1)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 881 | by (simp add: bv_to_int_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 882 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 883 | lemma bv_to_int_type [simp]: "bv_to_int (norm_signed w) = bv_to_int w" | 
| 23375 | 884 | proof (rule bit_list_induct [of _ w], simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 885 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 886 | assume ind: "bv_to_int (norm_signed xs) = bv_to_int xs" | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 887 | show "bv_to_int (norm_signed (\<zero>#xs)) = int (bv_to_nat xs)" | 
| 23375 | 888 | proof (rule bit_list_cases [of xs], simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 889 | fix ys | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 890 | assume [simp]: "xs = \<zero>#ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 891 | from ind | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 892 | show "bv_to_int (norm_signed (\<zero>#ys)) = int (bv_to_nat ys)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 893 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 894 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 895 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 896 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 897 | assume ind: "bv_to_int (norm_signed xs) = bv_to_int xs" | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23375diff
changeset | 898 | show "bv_to_int (norm_signed (\<one>#xs)) = -1 - int (bv_to_nat (bv_not xs))" | 
| 23375 | 899 | proof (rule bit_list_cases [of xs], simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 900 | fix ys | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 901 | assume [simp]: "xs = \<one>#ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 902 | from ind | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23375diff
changeset | 903 | show "bv_to_int (norm_signed (\<one>#ys)) = -1 - int (bv_to_nat (bv_not ys))" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 904 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 905 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 906 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 907 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 908 | lemma bv_to_int_upper_range: "bv_to_int w < 2 ^ (length w - 1)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 909 | proof (rule bit_list_cases [of w],simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 910 | fix bs | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 911 | from bv_to_nat_upper_range | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 912 | show "int (bv_to_nat bs) < 2 ^ length bs" | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 913 | by (simp add: int_nat_two_exp) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 914 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 915 | fix bs | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23375diff
changeset | 916 | have "-1 - int (bv_to_nat (bv_not bs)) \<le> 0" by simp | 
| 23375 | 917 | also have "... < 2 ^ length bs" by (induct bs) simp_all | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23375diff
changeset | 918 | finally show "-1 - int (bv_to_nat (bv_not bs)) < 2 ^ length bs" . | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 919 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 920 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 921 | lemma bv_to_int_lower_range: "- (2 ^ (length w - 1)) \<le> bv_to_int w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 922 | proof (rule bit_list_cases [of w],simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 923 | fix bs :: "bit list" | 
| 23375 | 924 | have "- (2 ^ length bs) \<le> (0::int)" by (induct bs) simp_all | 
| 925 | also have "... \<le> int (bv_to_nat bs)" by simp | |
| 926 | finally show "- (2 ^ length bs) \<le> int (bv_to_nat bs)" . | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 927 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 928 | fix bs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 929 | from bv_to_nat_upper_range [of "bv_not bs"] | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23375diff
changeset | 930 | show "- (2 ^ length bs) \<le> -1 - int (bv_to_nat (bv_not bs))" | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 931 | by (simp add: int_nat_two_exp) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 932 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 933 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 934 | lemma int_bv_int [simp]: "int_to_bv (bv_to_int w) = norm_signed w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 935 | proof (rule bit_list_cases [of w],simp) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 936 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 937 | assume [simp]: "w = \<zero>#xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 938 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 939 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 940 | apply (subst norm_signed_Cons [of "\<zero>" "xs"]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 941 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 942 | using norm_unsigned_result [of xs] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 943 | apply safe | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 944 | apply (rule bit_list_cases [of "norm_unsigned xs"]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 945 | apply simp_all | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 946 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 947 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 948 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 949 | assume [simp]: "w = \<one>#xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 950 | show ?thesis | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 951 | apply (simp del: int_to_bv_lt0) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 952 | apply (rule bit_list_induct [of _ xs]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 953 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 954 | apply (subst int_to_bv_lt0) | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 955 | apply (subgoal_tac "- int (bv_to_nat (bv_not (\<zero> # bs))) + -1 < 0 + 0") | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 956 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 957 | apply (rule add_le_less_mono) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 958 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 959 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 960 | apply (simp del: bv_to_nat1 bv_to_nat_helper) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 961 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 962 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 963 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 964 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 965 | lemma bv_int_bv [simp]: "bv_to_int (int_to_bv i) = i" | 
| 23375 | 966 | by (cases "0 \<le> i") simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 967 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 968 | lemma bv_msb_norm [simp]: "bv_msb (norm_signed w) = bv_msb w" | 
| 23375 | 969 | by (rule bit_list_cases [of w]) (simp_all add: norm_signed_Cons) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 970 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 971 | lemma norm_signed_length: "length (norm_signed w) \<le> length w" | 
| 23375 | 972 | apply (cases w, simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 973 | apply (subst norm_signed_Cons) | 
| 23375 | 974 | apply (case_tac a, simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 975 | apply (rule rem_initial_length) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 976 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 977 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 978 | lemma norm_signed_equal: "length (norm_signed w) = length w ==> norm_signed w = w" | 
| 23375 | 979 | proof (rule bit_list_cases [of w], simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 980 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 981 | assume "length (norm_signed (\<zero>#xs)) = Suc (length xs)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 982 | thus "norm_signed (\<zero>#xs) = \<zero>#xs" | 
| 35175 | 983 | by (simp add: norm_signed_Cons norm_unsigned_equal [THEN eqTrueI] | 
| 984 | split: split_if_asm) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 985 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 986 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 987 | assume "length (norm_signed (\<one>#xs)) = Suc (length xs)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 988 | thus "norm_signed (\<one>#xs) = \<one>#xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 989 | apply (simp add: norm_signed_Cons) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 990 | apply (rule rem_initial_equal) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 991 | apply assumption | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 992 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 993 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 994 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 995 | lemma bv_extend_norm_signed: "bv_msb w = b ==> bv_extend (length w) b (norm_signed w) = w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 996 | proof (rule bit_list_cases [of w],simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 997 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 998 | show "bv_extend (Suc (length xs)) \<zero> (norm_signed (\<zero>#xs)) = \<zero>#xs" | 
| 34941 | 999 | proof (simp add: norm_signed_def,auto) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1000 | assume "norm_unsigned xs = []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1001 | hence xx: "rem_initial \<zero> xs = []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1002 | by (simp add: norm_unsigned_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1003 | have "bv_extend (Suc (length xs)) \<zero> (\<zero>#rem_initial \<zero> xs) = \<zero>#xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1004 | apply (simp add: bv_extend_def replicate_app_Cons_same) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1005 | apply (fold bv_extend_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1006 | apply (rule bv_extend_rem_initial) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1007 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1008 | thus "bv_extend (Suc (length xs)) \<zero> [\<zero>] = \<zero>#xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1009 | by (simp add: xx) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1010 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1011 | show "bv_extend (Suc (length xs)) \<zero> (\<zero>#norm_unsigned xs) = \<zero>#xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1012 | apply (simp add: norm_unsigned_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1013 | apply (simp add: bv_extend_def replicate_app_Cons_same) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1014 | apply (fold bv_extend_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1015 | apply (rule bv_extend_rem_initial) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1016 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1017 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1018 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1019 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1020 | show "bv_extend (Suc (length xs)) \<one> (norm_signed (\<one>#xs)) = \<one>#xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1021 | apply (simp add: norm_signed_Cons) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1022 | apply (simp add: bv_extend_def replicate_app_Cons_same) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1023 | apply (fold bv_extend_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1024 | apply (rule bv_extend_rem_initial) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1025 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1026 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1027 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1028 | lemma bv_to_int_qinj: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1029 | assumes one: "bv_to_int xs = bv_to_int ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1030 | and len: "length xs = length ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1031 | shows "xs = ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1032 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1033 | from one | 
| 23375 | 1034 | have "int_to_bv (bv_to_int xs) = int_to_bv (bv_to_int ys)" by simp | 
| 1035 | hence xsys: "norm_signed xs = norm_signed ys" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1036 | hence xsys': "bv_msb xs = bv_msb ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1037 | proof - | 
| 23375 | 1038 | have "bv_msb xs = bv_msb (norm_signed xs)" by simp | 
| 1039 | also have "... = bv_msb (norm_signed ys)" by (simp add: xsys) | |
| 1040 | also have "... = bv_msb ys" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1041 | finally show ?thesis . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1042 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1043 | have "xs = bv_extend (length xs) (bv_msb xs) (norm_signed xs)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1044 | by (simp add: bv_extend_norm_signed) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1045 | also have "... = bv_extend (length ys) (bv_msb ys) (norm_signed ys)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1046 | by (simp add: xsys xsys' len) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1047 | also have "... = ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1048 | by (simp add: bv_extend_norm_signed) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1049 | finally show ?thesis . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1050 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1051 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1052 | lemma int_to_bv_returntype [simp]: "norm_signed (int_to_bv w) = int_to_bv w" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1053 | by (simp add: int_to_bv_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1054 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1055 | lemma bv_to_int_msb0: "0 \<le> bv_to_int w1 ==> bv_msb w1 = \<zero>" | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 1056 | by (rule bit_list_cases,simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1057 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1058 | lemma bv_to_int_msb1: "bv_to_int w1 < 0 ==> bv_msb w1 = \<one>" | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 1059 | by (rule bit_list_cases,simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1060 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1061 | lemma bv_to_int_lower_limit_gt0: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1062 | assumes w0: "0 < bv_to_int w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1063 | shows "2 ^ (length (norm_signed w) - 2) \<le> bv_to_int w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1064 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1065 | from w0 | 
| 23375 | 1066 | have "0 \<le> bv_to_int w" by simp | 
| 1067 | hence [simp]: "bv_msb w = \<zero>" by (rule bv_to_int_msb0) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1068 | have "2 ^ (length (norm_signed w) - 2) \<le> bv_to_int (norm_signed w)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1069 | proof (rule bit_list_cases [of w]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1070 | assume "w = []" | 
| 23375 | 1071 | with w0 show ?thesis by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1072 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1073 | fix w' | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1074 | assume weq: "w = \<zero> # w'" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1075 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1076 | proof (simp add: norm_signed_Cons,safe) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1077 | assume "norm_unsigned w' = []" | 
| 23375 | 1078 | with weq and w0 show False | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32456diff
changeset | 1079 | by (simp add: norm_empty_bv_to_nat_zero) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1080 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1081 | assume w'0: "norm_unsigned w' \<noteq> []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1082 | have "0 < bv_to_nat w'" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1083 | proof (rule ccontr) | 
| 19736 | 1084 | assume "~ (0 < bv_to_nat w')" | 
| 1085 | hence "bv_to_nat w' = 0" | |
| 1086 | by arith | |
| 1087 | hence "norm_unsigned w' = []" | |
| 1088 | by (simp add: bv_to_nat_zero_imp_empty) | |
| 1089 | with w'0 | |
| 23375 | 1090 | show False by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1091 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1092 | with bv_to_nat_lower_limit [of w'] | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 1093 | show "2 ^ (length (norm_unsigned w') - Suc 0) \<le> int (bv_to_nat w')" | 
| 19736 | 1094 | by (simp add: int_nat_two_exp) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1095 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1096 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1097 | fix w' | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1098 | assume "w = \<one> # w'" | 
| 23375 | 1099 | from w0 have "bv_msb w = \<zero>" by simp | 
| 1100 | with prems show ?thesis by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1101 | qed | 
| 23375 | 1102 | also have "... = bv_to_int w" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1103 | finally show ?thesis . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1104 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1105 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1106 | lemma norm_signed_result: "norm_signed w = [] \<or> norm_signed w = [\<one>] \<or> bv_msb (norm_signed w) \<noteq> bv_msb (tl (norm_signed w))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1107 | apply (rule bit_list_cases [of w],simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1108 | apply (case_tac "bs",simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1109 | apply (case_tac "a",simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1110 | apply (simp add: norm_signed_Cons) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1111 | apply safe | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1112 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1113 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1114 | fix l | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1115 | assume msb: "\<zero> = bv_msb (norm_unsigned l)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1116 | assume "norm_unsigned l \<noteq> []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1117 | with norm_unsigned_result [of l] | 
| 23375 | 1118 | have "bv_msb (norm_unsigned l) = \<one>" by simp | 
| 1119 | with msb show False by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1120 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1121 | fix xs | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1122 | assume p: "\<one> = bv_msb (tl (norm_signed (\<one> # xs)))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1123 | have "\<one> \<noteq> bv_msb (tl (norm_signed (\<one> # xs)))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1124 | by (rule bit_list_induct [of _ xs],simp_all) | 
| 23375 | 1125 | with p show False by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1126 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1127 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1128 | lemma bv_to_int_upper_limit_lem1: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1129 | assumes w0: "bv_to_int w < -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1130 | shows "bv_to_int w < - (2 ^ (length (norm_signed w) - 2))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1131 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1132 | from w0 | 
| 23375 | 1133 | have "bv_to_int w < 0" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1134 | hence msbw [simp]: "bv_msb w = \<one>" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1135 | by (rule bv_to_int_msb1) | 
| 23375 | 1136 | have "bv_to_int w = bv_to_int (norm_signed w)" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1137 | also from norm_signed_result [of w] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1138 | have "... < - (2 ^ (length (norm_signed w) - 2))" | 
| 23375 | 1139 | proof safe | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1140 | assume "norm_signed w = []" | 
| 23375 | 1141 | hence "bv_to_int (norm_signed w) = 0" by simp | 
| 1142 | with w0 show ?thesis by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1143 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1144 | assume "norm_signed w = [\<one>]" | 
| 23375 | 1145 | hence "bv_to_int (norm_signed w) = -1" by simp | 
| 1146 | with w0 show ?thesis by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1147 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1148 | assume "bv_msb (norm_signed w) \<noteq> bv_msb (tl (norm_signed w))" | 
| 23375 | 1149 | hence msb_tl: "\<one> \<noteq> bv_msb (tl (norm_signed w))" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1150 | show "bv_to_int (norm_signed w) < - (2 ^ (length (norm_signed w) - 2))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1151 | proof (rule bit_list_cases [of "norm_signed w"]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1152 | assume "norm_signed w = []" | 
| 23375 | 1153 | hence "bv_to_int (norm_signed w) = 0" by simp | 
| 1154 | with w0 show ?thesis by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1155 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1156 | fix w' | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1157 | assume nw: "norm_signed w = \<zero> # w'" | 
| 23375 | 1158 | from msbw have "bv_msb (norm_signed w) = \<one>" by simp | 
| 1159 | with nw show ?thesis by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1160 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1161 | fix w' | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1162 | assume weq: "norm_signed w = \<one> # w'" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1163 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1164 | proof (rule bit_list_cases [of w']) | 
| 19736 | 1165 | assume w'eq: "w' = []" | 
| 23375 | 1166 | from w0 have "bv_to_int (norm_signed w) < -1" by simp | 
| 1167 | with w'eq and weq show ?thesis by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1168 | next | 
| 19736 | 1169 | fix w'' | 
| 1170 | assume w'eq: "w' = \<zero> # w''" | |
| 1171 | show ?thesis | |
| 1172 | apply (simp add: weq w'eq) | |
| 1173 | apply (subgoal_tac "- int (bv_to_nat (bv_not w'')) + -1 < 0 + 0") | |
| 1174 | apply (simp add: int_nat_two_exp) | |
| 1175 | apply (rule add_le_less_mono) | |
| 1176 | apply simp_all | |
| 1177 | done | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1178 | next | 
| 19736 | 1179 | fix w'' | 
| 1180 | assume w'eq: "w' = \<one> # w''" | |
| 23375 | 1181 | with weq and msb_tl show ?thesis by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1182 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1183 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1184 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1185 | finally show ?thesis . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1186 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1187 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1188 | lemma length_int_to_bv_upper_limit_gt0: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1189 | assumes w0: "0 < i" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1190 | and wk: "i \<le> 2 ^ (k - 1) - 1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1191 | shows "length (int_to_bv i) \<le> k" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1192 | proof (rule ccontr) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1193 | from w0 wk | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1194 | have k1: "1 < k" | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19736diff
changeset | 1195 | by (cases "k - 1",simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1196 | assume "~ length (int_to_bv i) \<le> k" | 
| 23375 | 1197 | hence "k < length (int_to_bv i)" by simp | 
| 1198 | hence "k \<le> length (int_to_bv i) - 1" by arith | |
| 1199 | hence a: "k - 1 \<le> length (int_to_bv i) - 2" by arith | |
| 15067 | 1200 | hence "(2::int) ^ (k - 1) \<le> 2 ^ (length (int_to_bv i) - 2)" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1201 | also have "... \<le> i" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1202 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1203 | have "2 ^ (length (norm_signed (int_to_bv i)) - 2) \<le> bv_to_int (int_to_bv i)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1204 | proof (rule bv_to_int_lower_limit_gt0) | 
| 23375 | 1205 | from w0 show "0 < bv_to_int (int_to_bv i)" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1206 | qed | 
| 23375 | 1207 | thus ?thesis by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1208 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1209 | finally have "2 ^ (k - 1) \<le> i" . | 
| 23375 | 1210 | with wk show False by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1211 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1212 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1213 | lemma pos_length_pos: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1214 | assumes i0: "0 < bv_to_int w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1215 | shows "0 < length w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1216 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1217 | from norm_signed_result [of w] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1218 | have "0 < length (norm_signed w)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1219 | proof (auto) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1220 | assume ii: "norm_signed w = []" | 
| 23375 | 1221 | have "bv_to_int (norm_signed w) = 0" by (subst ii) simp | 
| 1222 | hence "bv_to_int w = 0" by simp | |
| 1223 | with i0 show False by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1224 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1225 | assume ii: "norm_signed w = []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1226 | assume jj: "bv_msb w \<noteq> \<zero>" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1227 | have "\<zero> = bv_msb (norm_signed w)" | 
| 23375 | 1228 | by (subst ii) simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1229 | also have "... \<noteq> \<zero>" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1230 | by (simp add: jj) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1231 | finally show False by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1232 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1233 | also have "... \<le> length w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1234 | by (rule norm_signed_length) | 
| 23375 | 1235 | finally show ?thesis . | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1236 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1237 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1238 | lemma neg_length_pos: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1239 | assumes i0: "bv_to_int w < -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1240 | shows "0 < length w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1241 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1242 | from norm_signed_result [of w] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1243 | have "0 < length (norm_signed w)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1244 | proof (auto) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1245 | assume ii: "norm_signed w = []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1246 | have "bv_to_int (norm_signed w) = 0" | 
| 23375 | 1247 | by (subst ii) simp | 
| 1248 | hence "bv_to_int w = 0" by simp | |
| 1249 | with i0 show False by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1250 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1251 | assume ii: "norm_signed w = []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1252 | assume jj: "bv_msb w \<noteq> \<zero>" | 
| 23375 | 1253 | have "\<zero> = bv_msb (norm_signed w)" by (subst ii) simp | 
| 1254 | also have "... \<noteq> \<zero>" by (simp add: jj) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1255 | finally show False by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1256 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1257 | also have "... \<le> length w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1258 | by (rule norm_signed_length) | 
| 23375 | 1259 | finally show ?thesis . | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1260 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1261 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1262 | lemma length_int_to_bv_lower_limit_gt0: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1263 | assumes wk: "2 ^ (k - 1) \<le> i" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1264 | shows "k < length (int_to_bv i)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1265 | proof (rule ccontr) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1266 | have "0 < (2::int) ^ (k - 1)" | 
| 23375 | 1267 | by (rule zero_less_power) simp | 
| 1268 | also have "... \<le> i" by (rule wk) | |
| 1269 | finally have i0: "0 < i" . | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1270 | have lii0: "0 < length (int_to_bv i)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1271 | apply (rule pos_length_pos) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1272 | apply (simp,rule i0) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1273 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1274 | assume "~ k < length (int_to_bv i)" | 
| 23375 | 1275 | hence "length (int_to_bv i) \<le> k" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1276 | with lii0 | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1277 | have a: "length (int_to_bv i) - 1 \<le> k - 1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1278 | by arith | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1279 | have "i < 2 ^ (length (int_to_bv i) - 1)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1280 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1281 | have "i = bv_to_int (int_to_bv i)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1282 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1283 | also have "... < 2 ^ (length (int_to_bv i) - 1)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1284 | by (rule bv_to_int_upper_range) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1285 | finally show ?thesis . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1286 | qed | 
| 15067 | 1287 | also have "(2::int) ^ (length (int_to_bv i) - 1) \<le> 2 ^ (k - 1)" using a | 
| 23375 | 1288 | by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1289 | finally have "i < 2 ^ (k - 1)" . | 
| 23375 | 1290 | with wk show False by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1291 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1292 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1293 | lemma length_int_to_bv_upper_limit_lem1: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1294 | assumes w1: "i < -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1295 | and wk: "- (2 ^ (k - 1)) \<le> i" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1296 | shows "length (int_to_bv i) \<le> k" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1297 | proof (rule ccontr) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1298 | from w1 wk | 
| 23375 | 1299 | have k1: "1 < k" by (cases "k - 1") simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1300 | assume "~ length (int_to_bv i) \<le> k" | 
| 23375 | 1301 | hence "k < length (int_to_bv i)" by simp | 
| 1302 | hence "k \<le> length (int_to_bv i) - 1" by arith | |
| 1303 | hence a: "k - 1 \<le> length (int_to_bv i) - 2" by arith | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1304 | have "i < - (2 ^ (length (int_to_bv i) - 2))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1305 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1306 | have "i = bv_to_int (int_to_bv i)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1307 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1308 | also have "... < - (2 ^ (length (norm_signed (int_to_bv i)) - 2))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1309 | by (rule bv_to_int_upper_limit_lem1,simp,rule w1) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1310 | finally show ?thesis by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1311 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1312 | also have "... \<le> -(2 ^ (k - 1))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1313 | proof - | 
| 23375 | 1314 | have "(2::int) ^ (k - 1) \<le> 2 ^ (length (int_to_bv i) - 2)" using a by simp | 
| 1315 | thus ?thesis by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1316 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1317 | finally have "i < -(2 ^ (k - 1))" . | 
| 23375 | 1318 | with wk show False by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1319 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1320 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1321 | lemma length_int_to_bv_lower_limit_lem1: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1322 | assumes wk: "i < -(2 ^ (k - 1))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1323 | shows "k < length (int_to_bv i)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1324 | proof (rule ccontr) | 
| 23375 | 1325 | from wk have "i \<le> -(2 ^ (k - 1)) - 1" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1326 | also have "... < -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1327 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1328 | have "0 < (2::int) ^ (k - 1)" | 
| 23375 | 1329 | by (rule zero_less_power) simp | 
| 1330 | hence "-((2::int) ^ (k - 1)) < 0" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1331 | thus ?thesis by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1332 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1333 | finally have i1: "i < -1" . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1334 | have lii0: "0 < length (int_to_bv i)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1335 | apply (rule neg_length_pos) | 
| 23375 | 1336 | apply (simp, rule i1) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1337 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1338 | assume "~ k < length (int_to_bv i)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1339 | hence "length (int_to_bv i) \<le> k" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1340 | by simp | 
| 23375 | 1341 | with lii0 have a: "length (int_to_bv i) - 1 \<le> k - 1" by arith | 
| 15067 | 1342 | hence "(2::int) ^ (length (int_to_bv i) - 1) \<le> 2 ^ (k - 1)" by simp | 
| 23375 | 1343 | hence "-((2::int) ^ (k - 1)) \<le> - (2 ^ (length (int_to_bv i) - 1))" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1344 | also have "... \<le> i" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1345 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1346 | have "- (2 ^ (length (int_to_bv i) - 1)) \<le> bv_to_int (int_to_bv i)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1347 | by (rule bv_to_int_lower_range) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1348 | also have "... = i" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1349 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1350 | finally show ?thesis . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1351 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1352 | finally have "-(2 ^ (k - 1)) \<le> i" . | 
| 23375 | 1353 | with wk show False by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1354 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1355 | |
| 23375 | 1356 | |
| 14589 | 1357 | subsection {* Signed Arithmetic Operations *}
 | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1358 | |
| 14589 | 1359 | subsubsection {* Conversion from unsigned to signed *}
 | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1360 | |
| 19736 | 1361 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1362 | utos :: "bit list => bit list" where | 
| 19736 | 1363 | "utos w = norm_signed (\<zero> # w)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1364 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1365 | lemma utos_type [simp]: "utos (norm_unsigned w) = utos w" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1366 | by (simp add: utos_def norm_signed_Cons) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1367 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1368 | lemma utos_returntype [simp]: "norm_signed (utos w) = utos w" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1369 | by (simp add: utos_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1370 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1371 | lemma utos_length: "length (utos w) \<le> Suc (length w)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1372 | by (simp add: utos_def norm_signed_Cons) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1373 | |
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 1374 | lemma bv_to_int_utos: "bv_to_int (utos w) = int (bv_to_nat w)" | 
| 23375 | 1375 | proof (simp add: utos_def norm_signed_Cons, safe) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1376 | assume "norm_unsigned w = []" | 
| 23375 | 1377 | hence "bv_to_nat (norm_unsigned w) = 0" by simp | 
| 1378 | thus "bv_to_nat w = 0" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1379 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1380 | |
| 23375 | 1381 | |
| 14589 | 1382 | subsubsection {* Unary minus *}
 | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1383 | |
| 19736 | 1384 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1385 | bv_uminus :: "bit list => bit list" where | 
| 19736 | 1386 | "bv_uminus w = int_to_bv (- bv_to_int w)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1387 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1388 | lemma bv_uminus_type [simp]: "bv_uminus (norm_signed w) = bv_uminus w" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1389 | by (simp add: bv_uminus_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1390 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1391 | lemma bv_uminus_returntype [simp]: "norm_signed (bv_uminus w) = bv_uminus w" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1392 | by (simp add: bv_uminus_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1393 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1394 | lemma bv_uminus_length: "length (bv_uminus w) \<le> Suc (length w)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1395 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1396 | have "1 < -bv_to_int w \<or> -bv_to_int w = 1 \<or> -bv_to_int w = 0 \<or> -bv_to_int w = -1 \<or> -bv_to_int w < -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1397 | by arith | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1398 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1399 | proof safe | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1400 | assume p: "1 < - bv_to_int w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1401 | have lw: "0 < length w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1402 | apply (rule neg_length_pos) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1403 | using p | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1404 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1405 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1406 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1407 | proof (simp add: bv_uminus_def,rule length_int_to_bv_upper_limit_gt0,simp_all) | 
| 23375 | 1408 | from prems show "bv_to_int w < 0" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1409 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1410 | have "-(2^(length w - 1)) \<le> bv_to_int w" | 
| 19736 | 1411 | by (rule bv_to_int_lower_range) | 
| 23375 | 1412 | hence "- bv_to_int w \<le> 2^(length w - 1)" by simp | 
| 1413 | also from lw have "... < 2 ^ length w" by simp | |
| 1414 | finally show "- bv_to_int w < 2 ^ length w" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1415 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1416 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1417 | assume p: "- bv_to_int w = 1" | 
| 23375 | 1418 | hence lw: "0 < length w" by (cases w) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1419 | from p | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1420 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1421 | apply (simp add: bv_uminus_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1422 | using lw | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1423 | apply (simp (no_asm) add: nat_to_bv_non0) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1424 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1425 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1426 | assume "- bv_to_int w = 0" | 
| 23375 | 1427 | thus ?thesis by (simp add: bv_uminus_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1428 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1429 | assume p: "- bv_to_int w = -1" | 
| 23375 | 1430 | thus ?thesis by (simp add: bv_uminus_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1431 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1432 | assume p: "- bv_to_int w < -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1433 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1434 | apply (simp add: bv_uminus_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1435 | apply (rule length_int_to_bv_upper_limit_lem1) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1436 | apply (rule p) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1437 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1438 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1439 | have "bv_to_int w < 2 ^ (length w - 1)" | 
| 19736 | 1440 | by (rule bv_to_int_upper_range) | 
| 15067 | 1441 | also have "... \<le> 2 ^ length w" by simp | 
| 23375 | 1442 | finally show "bv_to_int w \<le> 2 ^ length w" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1443 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1444 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1445 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1446 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1447 | lemma bv_uminus_length_utos: "length (bv_uminus (utos w)) \<le> Suc (length w)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1448 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1449 | have "-bv_to_int (utos w) = 0 \<or> -bv_to_int (utos w) = -1 \<or> -bv_to_int (utos w) < -1" | 
| 23375 | 1450 | by (simp add: bv_to_int_utos, arith) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1451 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1452 | proof safe | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1453 | assume "-bv_to_int (utos w) = 0" | 
| 23375 | 1454 | thus ?thesis by (simp add: bv_uminus_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1455 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1456 | assume "-bv_to_int (utos w) = -1" | 
| 23375 | 1457 | thus ?thesis by (simp add: bv_uminus_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1458 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1459 | assume p: "-bv_to_int (utos w) < -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1460 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1461 | apply (simp add: bv_uminus_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1462 | apply (rule length_int_to_bv_upper_limit_lem1) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1463 | apply (rule p) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1464 | apply (simp add: bv_to_int_utos) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1465 | using bv_to_nat_upper_range [of w] | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 1466 | apply (simp add: int_nat_two_exp) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1467 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1468 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1469 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1470 | |
| 19736 | 1471 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1472 | bv_sadd :: "[bit list, bit list ] => bit list" where | 
| 19736 | 1473 | "bv_sadd w1 w2 = int_to_bv (bv_to_int w1 + bv_to_int w2)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1474 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1475 | lemma bv_sadd_type1 [simp]: "bv_sadd (norm_signed w1) w2 = bv_sadd w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1476 | by (simp add: bv_sadd_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1477 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1478 | lemma bv_sadd_type2 [simp]: "bv_sadd w1 (norm_signed w2) = bv_sadd w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1479 | by (simp add: bv_sadd_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1480 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1481 | lemma bv_sadd_returntype [simp]: "norm_signed (bv_sadd w1 w2) = bv_sadd w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1482 | by (simp add: bv_sadd_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1483 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1484 | lemma adder_helper: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1485 | assumes lw: "0 < max (length w1) (length w2)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1486 | shows "((2::int) ^ (length w1 - 1)) + (2 ^ (length w2 - 1)) \<le> 2 ^ max (length w1) (length w2)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1487 | proof - | 
| 23375 | 1488 | have "((2::int) ^ (length w1 - 1)) + (2 ^ (length w2 - 1)) \<le> | 
| 1489 | 2 ^ (max (length w1) (length w2) - 1) + 2 ^ (max (length w1) (length w2) - 1)" | |
| 32438 | 1490 | by (auto simp:max_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1491 | also have "... = 2 ^ max (length w1) (length w2)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1492 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1493 | from lw | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1494 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1495 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1496 | apply (subst power_Suc [symmetric]) | 
| 30960 | 1497 | apply simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1498 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1499 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1500 | finally show ?thesis . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1501 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1502 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1503 | lemma bv_sadd_length: "length (bv_sadd w1 w2) \<le> Suc (max (length w1) (length w2))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1504 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1505 | let ?Q = "bv_to_int w1 + bv_to_int w2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1506 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1507 | have helper: "?Q \<noteq> 0 ==> 0 < max (length w1) (length w2)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1508 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1509 | assume p: "?Q \<noteq> 0" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1510 | show "0 < max (length w1) (length w2)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1511 | proof (simp add: less_max_iff_disj,rule) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1512 | assume [simp]: "w1 = []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1513 | show "w2 \<noteq> []" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1514 | proof (rule ccontr,simp) | 
| 19736 | 1515 | assume [simp]: "w2 = []" | 
| 23375 | 1516 | from p show False by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1517 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1518 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1519 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1520 | |
| 23375 | 1521 | have "0 < ?Q \<or> ?Q = 0 \<or> ?Q = -1 \<or> ?Q < -1" by arith | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1522 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1523 | proof safe | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1524 | assume "?Q = 0" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1525 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1526 | by (simp add: bv_sadd_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1527 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1528 | assume "?Q = -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1529 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1530 | by (simp add: bv_sadd_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1531 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1532 | assume p: "0 < ?Q" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1533 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1534 | apply (simp add: bv_sadd_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1535 | apply (rule length_int_to_bv_upper_limit_gt0) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1536 | apply (rule p) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1537 | proof simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1538 | from bv_to_int_upper_range [of w2] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1539 | have "bv_to_int w2 \<le> 2 ^ (length w2 - 1)" | 
| 19736 | 1540 | by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1541 | with bv_to_int_upper_range [of w1] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1542 | have "bv_to_int w1 + bv_to_int w2 < (2 ^ (length w1 - 1)) + (2 ^ (length w2 - 1))" | 
| 19736 | 1543 | by (rule zadd_zless_mono) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1544 | also have "... \<le> 2 ^ max (length w1) (length w2)" | 
| 19736 | 1545 | apply (rule adder_helper) | 
| 1546 | apply (rule helper) | |
| 1547 | using p | |
| 1548 | apply simp | |
| 1549 | done | |
| 23375 | 1550 | finally show "?Q < 2 ^ max (length w1) (length w2)" . | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1551 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1552 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1553 | assume p: "?Q < -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1554 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1555 | apply (simp add: bv_sadd_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1556 | apply (rule length_int_to_bv_upper_limit_lem1,simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1557 | apply (rule p) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1558 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1559 | have "(2 ^ (length w1 - 1)) + 2 ^ (length w2 - 1) \<le> (2::int) ^ max (length w1) (length w2)" | 
| 19736 | 1560 | apply (rule adder_helper) | 
| 1561 | apply (rule helper) | |
| 1562 | using p | |
| 1563 | apply simp | |
| 1564 | done | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1565 | hence "-((2::int) ^ max (length w1) (length w2)) \<le> - (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1))" | 
| 19736 | 1566 | by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1567 | also have "- (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1)) \<le> ?Q" | 
| 19736 | 1568 | apply (rule add_mono) | 
| 1569 | apply (rule bv_to_int_lower_range [of w1]) | |
| 1570 | apply (rule bv_to_int_lower_range [of w2]) | |
| 1571 | done | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1572 | finally show "- (2^max (length w1) (length w2)) \<le> ?Q" . | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1573 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1574 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1575 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1576 | |
| 19736 | 1577 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1578 | bv_sub :: "[bit list, bit list] => bit list" where | 
| 19736 | 1579 | "bv_sub w1 w2 = bv_sadd w1 (bv_uminus w2)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1580 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1581 | lemma bv_sub_type1 [simp]: "bv_sub (norm_signed w1) w2 = bv_sub w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1582 | by (simp add: bv_sub_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1583 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1584 | lemma bv_sub_type2 [simp]: "bv_sub w1 (norm_signed w2) = bv_sub w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1585 | by (simp add: bv_sub_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1586 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1587 | lemma bv_sub_returntype [simp]: "norm_signed (bv_sub w1 w2) = bv_sub w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1588 | by (simp add: bv_sub_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1589 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1590 | lemma bv_sub_length: "length (bv_sub w1 w2) \<le> Suc (max (length w1) (length w2))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1591 | proof (cases "bv_to_int w2 = 0") | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1592 | assume p: "bv_to_int w2 = 0" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1593 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1594 | proof (simp add: bv_sub_def bv_sadd_def bv_uminus_def p) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1595 | have "length (norm_signed w1) \<le> length w1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1596 | by (rule norm_signed_length) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1597 | also have "... \<le> max (length w1) (length w2)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1598 | by (rule le_maxI1) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1599 | also have "... \<le> Suc (max (length w1) (length w2))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1600 | by arith | 
| 23375 | 1601 | finally show "length (norm_signed w1) \<le> Suc (max (length w1) (length w2))" . | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1602 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1603 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1604 | assume "bv_to_int w2 \<noteq> 0" | 
| 23375 | 1605 | hence "0 < length w2" by (cases w2,simp_all) | 
| 1606 | hence lmw: "0 < max (length w1) (length w2)" by arith | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1607 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1608 | let ?Q = "bv_to_int w1 - bv_to_int w2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1609 | |
| 23375 | 1610 | have "0 < ?Q \<or> ?Q = 0 \<or> ?Q = -1 \<or> ?Q < -1" by arith | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1611 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1612 | proof safe | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1613 | assume "?Q = 0" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1614 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1615 | by (simp add: bv_sub_def bv_sadd_def bv_uminus_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1616 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1617 | assume "?Q = -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1618 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1619 | by (simp add: bv_sub_def bv_sadd_def bv_uminus_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1620 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1621 | assume p: "0 < ?Q" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1622 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1623 | apply (simp add: bv_sub_def bv_sadd_def bv_uminus_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1624 | apply (rule length_int_to_bv_upper_limit_gt0) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1625 | apply (rule p) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1626 | proof simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1627 | from bv_to_int_lower_range [of w2] | 
| 23375 | 1628 | have v2: "- bv_to_int w2 \<le> 2 ^ (length w2 - 1)" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1629 | have "bv_to_int w1 + - bv_to_int w2 < (2 ^ (length w1 - 1)) + (2 ^ (length w2 - 1))" | 
| 19736 | 1630 | apply (rule zadd_zless_mono) | 
| 1631 | apply (rule bv_to_int_upper_range [of w1]) | |
| 1632 | apply (rule v2) | |
| 1633 | done | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1634 | also have "... \<le> 2 ^ max (length w1) (length w2)" | 
| 19736 | 1635 | apply (rule adder_helper) | 
| 1636 | apply (rule lmw) | |
| 1637 | done | |
| 23375 | 1638 | finally show "?Q < 2 ^ max (length w1) (length w2)" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1639 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1640 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1641 | assume p: "?Q < -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1642 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1643 | apply (simp add: bv_sub_def bv_sadd_def bv_uminus_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1644 | apply (rule length_int_to_bv_upper_limit_lem1) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1645 | apply (rule p) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1646 | proof simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1647 | have "(2 ^ (length w1 - 1)) + 2 ^ (length w2 - 1) \<le> (2::int) ^ max (length w1) (length w2)" | 
| 19736 | 1648 | apply (rule adder_helper) | 
| 1649 | apply (rule lmw) | |
| 1650 | done | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1651 | hence "-((2::int) ^ max (length w1) (length w2)) \<le> - (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1))" | 
| 19736 | 1652 | by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1653 | also have "- (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1)) \<le> bv_to_int w1 + -bv_to_int w2" | 
| 19736 | 1654 | apply (rule add_mono) | 
| 1655 | apply (rule bv_to_int_lower_range [of w1]) | |
| 1656 | using bv_to_int_upper_range [of w2] | |
| 1657 | apply simp | |
| 1658 | done | |
| 23375 | 1659 | finally show "- (2^max (length w1) (length w2)) \<le> ?Q" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1660 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1661 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1662 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1663 | |
| 19736 | 1664 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1665 | bv_smult :: "[bit list, bit list] => bit list" where | 
| 19736 | 1666 | "bv_smult w1 w2 = int_to_bv (bv_to_int w1 * bv_to_int w2)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1667 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1668 | lemma bv_smult_type1 [simp]: "bv_smult (norm_signed w1) w2 = bv_smult w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1669 | by (simp add: bv_smult_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1670 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1671 | lemma bv_smult_type2 [simp]: "bv_smult w1 (norm_signed w2) = bv_smult w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1672 | by (simp add: bv_smult_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1673 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 1674 | lemma bv_smult_returntype [simp]: "norm_signed (bv_smult w1 w2) = bv_smult w1 w2" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1675 | by (simp add: bv_smult_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1676 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1677 | lemma bv_smult_length: "length (bv_smult w1 w2) \<le> length w1 + length w2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1678 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1679 | let ?Q = "bv_to_int w1 * bv_to_int w2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1680 | |
| 23375 | 1681 | have lmw: "?Q \<noteq> 0 ==> 0 < length w1 \<and> 0 < length w2" by auto | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1682 | |
| 23375 | 1683 | have "0 < ?Q \<or> ?Q = 0 \<or> ?Q = -1 \<or> ?Q < -1" by arith | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1684 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1685 | proof (safe dest!: iffD1 [OF mult_eq_0_iff]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1686 | assume "bv_to_int w1 = 0" | 
| 23375 | 1687 | thus ?thesis by (simp add: bv_smult_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1688 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1689 | assume "bv_to_int w2 = 0" | 
| 23375 | 1690 | thus ?thesis by (simp add: bv_smult_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1691 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1692 | assume p: "?Q = -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1693 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1694 | apply (simp add: bv_smult_def p) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1695 | apply (cut_tac lmw) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1696 | apply arith | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1697 | using p | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1698 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1699 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1700 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1701 | assume p: "0 < ?Q" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1702 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1703 | proof (simp add: zero_less_mult_iff,safe) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1704 | assume bi1: "0 < bv_to_int w1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1705 | assume bi2: "0 < bv_to_int w2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1706 | show ?thesis | 
| 19736 | 1707 | apply (simp add: bv_smult_def) | 
| 1708 | apply (rule length_int_to_bv_upper_limit_gt0) | |
| 1709 | apply (rule p) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1710 | proof simp | 
| 19736 | 1711 | have "?Q < 2 ^ (length w1 - 1) * 2 ^ (length w2 - 1)" | 
| 1712 | apply (rule mult_strict_mono) | |
| 1713 | apply (rule bv_to_int_upper_range) | |
| 1714 | apply (rule bv_to_int_upper_range) | |
| 1715 | apply (rule zero_less_power) | |
| 1716 | apply simp | |
| 1717 | using bi2 | |
| 1718 | apply simp | |
| 1719 | done | |
| 1720 | also have "... \<le> 2 ^ (length w1 + length w2 - Suc 0)" | |
| 1721 | apply simp | |
| 1722 | apply (subst zpower_zadd_distrib [symmetric]) | |
| 1723 | apply simp | |
| 1724 | done | |
| 23375 | 1725 | finally show "?Q < 2 ^ (length w1 + length w2 - Suc 0)" . | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1726 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1727 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1728 | assume bi1: "bv_to_int w1 < 0" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1729 | assume bi2: "bv_to_int w2 < 0" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1730 | show ?thesis | 
| 19736 | 1731 | apply (simp add: bv_smult_def) | 
| 1732 | apply (rule length_int_to_bv_upper_limit_gt0) | |
| 1733 | apply (rule p) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1734 | proof simp | 
| 19736 | 1735 | have "-bv_to_int w1 * -bv_to_int w2 \<le> 2 ^ (length w1 - 1) * 2 ^(length w2 - 1)" | 
| 1736 | apply (rule mult_mono) | |
| 1737 | using bv_to_int_lower_range [of w1] | |
| 1738 | apply simp | |
| 1739 | using bv_to_int_lower_range [of w2] | |
| 1740 | apply simp | |
| 1741 | apply (rule zero_le_power,simp) | |
| 1742 | using bi2 | |
| 1743 | apply simp | |
| 1744 | done | |
| 1745 | hence "?Q \<le> 2 ^ (length w1 - 1) * 2 ^(length w2 - 1)" | |
| 1746 | by simp | |
| 1747 | also have "... < 2 ^ (length w1 + length w2 - Suc 0)" | |
| 1748 | apply simp | |
| 1749 | apply (subst zpower_zadd_distrib [symmetric]) | |
| 1750 | apply simp | |
| 1751 | apply (cut_tac lmw) | |
| 1752 | apply arith | |
| 1753 | apply (cut_tac p) | |
| 1754 | apply arith | |
| 1755 | done | |
| 1756 | finally show "?Q < 2 ^ (length w1 + length w2 - Suc 0)" . | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1757 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1758 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1759 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1760 | assume p: "?Q < -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1761 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1762 | apply (subst bv_smult_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1763 | apply (rule length_int_to_bv_upper_limit_lem1) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1764 | apply (rule p) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1765 | proof simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1766 | have "(2::int) ^ (length w1 - 1) * 2 ^(length w2 - 1) \<le> 2 ^ (length w1 + length w2 - Suc 0)" | 
| 19736 | 1767 | apply simp | 
| 1768 | apply (subst zpower_zadd_distrib [symmetric]) | |
| 1769 | apply simp | |
| 1770 | done | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1771 | hence "-((2::int) ^ (length w1 + length w2 - Suc 0)) \<le> -(2^(length w1 - 1) * 2 ^ (length w2 - 1))" | 
| 19736 | 1772 | by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1773 | also have "... \<le> ?Q" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1774 | proof - | 
| 19736 | 1775 | from p | 
| 1776 | have q: "bv_to_int w1 * bv_to_int w2 < 0" | |
| 1777 | by simp | |
| 1778 | thus ?thesis | |
| 1779 | proof (simp add: mult_less_0_iff,safe) | |
| 1780 | assume bi1: "0 < bv_to_int w1" | |
| 1781 | assume bi2: "bv_to_int w2 < 0" | |
| 1782 | have "-bv_to_int w2 * bv_to_int w1 \<le> ((2::int)^(length w2 - 1)) * (2 ^ (length w1 - 1))" | |
| 1783 | apply (rule mult_mono) | |
| 1784 | using bv_to_int_lower_range [of w2] | |
| 1785 | apply simp | |
| 1786 | using bv_to_int_upper_range [of w1] | |
| 1787 | apply simp | |
| 1788 | apply (rule zero_le_power,simp) | |
| 1789 | using bi1 | |
| 1790 | apply simp | |
| 1791 | done | |
| 1792 | hence "-?Q \<le> ((2::int)^(length w1 - 1)) * (2 ^ (length w2 - 1))" | |
| 1793 | by (simp add: zmult_ac) | |
| 1794 | thus "-(((2::int)^(length w1 - Suc 0)) * (2 ^ (length w2 - Suc 0))) \<le> ?Q" | |
| 1795 | by simp | |
| 1796 | next | |
| 1797 | assume bi1: "bv_to_int w1 < 0" | |
| 1798 | assume bi2: "0 < bv_to_int w2" | |
| 1799 | have "-bv_to_int w1 * bv_to_int w2 \<le> ((2::int)^(length w1 - 1)) * (2 ^ (length w2 - 1))" | |
| 1800 | apply (rule mult_mono) | |
| 1801 | using bv_to_int_lower_range [of w1] | |
| 1802 | apply simp | |
| 1803 | using bv_to_int_upper_range [of w2] | |
| 1804 | apply simp | |
| 1805 | apply (rule zero_le_power,simp) | |
| 1806 | using bi2 | |
| 1807 | apply simp | |
| 1808 | done | |
| 1809 | hence "-?Q \<le> ((2::int)^(length w1 - 1)) * (2 ^ (length w2 - 1))" | |
| 1810 | by (simp add: zmult_ac) | |
| 1811 | thus "-(((2::int)^(length w1 - Suc 0)) * (2 ^ (length w2 - Suc 0))) \<le> ?Q" | |
| 1812 | by simp | |
| 1813 | qed | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1814 | qed | 
| 23375 | 1815 | finally show "-(2 ^ (length w1 + length w2 - Suc 0)) \<le> ?Q" . | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1816 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1817 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1818 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1819 | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1820 | lemma bv_msb_one: "bv_msb w = \<one> ==> bv_to_nat w \<noteq> 0" | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1821 | by (cases w) simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1822 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1823 | lemma bv_smult_length_utos: "length (bv_smult (utos w1) w2) \<le> length w1 + length w2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1824 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1825 | let ?Q = "bv_to_int (utos w1) * bv_to_int w2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1826 | |
| 23375 | 1827 | have lmw: "?Q \<noteq> 0 ==> 0 < length (utos w1) \<and> 0 < length w2" by auto | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1828 | |
| 23375 | 1829 | have "0 < ?Q \<or> ?Q = 0 \<or> ?Q = -1 \<or> ?Q < -1" by arith | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1830 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1831 | proof (safe dest!: iffD1 [OF mult_eq_0_iff]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1832 | assume "bv_to_int (utos w1) = 0" | 
| 23375 | 1833 | thus ?thesis by (simp add: bv_smult_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1834 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1835 | assume "bv_to_int w2 = 0" | 
| 23375 | 1836 | thus ?thesis by (simp add: bv_smult_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1837 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1838 | assume p: "0 < ?Q" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1839 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1840 | proof (simp add: zero_less_mult_iff,safe) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1841 | assume biw2: "0 < bv_to_int w2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1842 | show ?thesis | 
| 19736 | 1843 | apply (simp add: bv_smult_def) | 
| 1844 | apply (rule length_int_to_bv_upper_limit_gt0) | |
| 1845 | apply (rule p) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1846 | proof simp | 
| 19736 | 1847 | have "?Q < 2 ^ length w1 * 2 ^ (length w2 - 1)" | 
| 1848 | apply (rule mult_strict_mono) | |
| 1849 | apply (simp add: bv_to_int_utos int_nat_two_exp) | |
| 1850 | apply (rule bv_to_nat_upper_range) | |
| 1851 | apply (rule bv_to_int_upper_range) | |
| 1852 | apply (rule zero_less_power,simp) | |
| 1853 | using biw2 | |
| 1854 | apply simp | |
| 1855 | done | |
| 1856 | also have "... \<le> 2 ^ (length w1 + length w2 - Suc 0)" | |
| 1857 | apply simp | |
| 1858 | apply (subst zpower_zadd_distrib [symmetric]) | |
| 1859 | apply simp | |
| 1860 | apply (cut_tac lmw) | |
| 1861 | apply arith | |
| 1862 | using p | |
| 1863 | apply auto | |
| 1864 | done | |
| 23375 | 1865 | finally show "?Q < 2 ^ (length w1 + length w2 - Suc 0)" . | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1866 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1867 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1868 | assume "bv_to_int (utos w1) < 0" | 
| 23375 | 1869 | thus ?thesis by (simp add: bv_to_int_utos) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1870 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1871 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1872 | assume p: "?Q = -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1873 | thus ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1874 | apply (simp add: bv_smult_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1875 | apply (cut_tac lmw) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1876 | apply arith | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1877 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1878 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1879 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1880 | assume p: "?Q < -1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1881 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1882 | apply (subst bv_smult_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1883 | apply (rule length_int_to_bv_upper_limit_lem1) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1884 | apply (rule p) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1885 | proof simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1886 | have "(2::int) ^ length w1 * 2 ^(length w2 - 1) \<le> 2 ^ (length w1 + length w2 - Suc 0)" | 
| 19736 | 1887 | apply simp | 
| 1888 | apply (subst zpower_zadd_distrib [symmetric]) | |
| 1889 | apply simp | |
| 1890 | apply (cut_tac lmw) | |
| 1891 | apply arith | |
| 1892 | apply (cut_tac p) | |
| 1893 | apply arith | |
| 1894 | done | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1895 | hence "-((2::int) ^ (length w1 + length w2 - Suc 0)) \<le> -(2^ length w1 * 2 ^ (length w2 - 1))" | 
| 19736 | 1896 | by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1897 | also have "... \<le> ?Q" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1898 | proof - | 
| 19736 | 1899 | from p | 
| 1900 | have q: "bv_to_int (utos w1) * bv_to_int w2 < 0" | |
| 1901 | by simp | |
| 1902 | thus ?thesis | |
| 1903 | proof (simp add: mult_less_0_iff,safe) | |
| 1904 | assume bi1: "0 < bv_to_int (utos w1)" | |
| 1905 | assume bi2: "bv_to_int w2 < 0" | |
| 1906 | have "-bv_to_int w2 * bv_to_int (utos w1) \<le> ((2::int)^(length w2 - 1)) * (2 ^ length w1)" | |
| 1907 | apply (rule mult_mono) | |
| 1908 | using bv_to_int_lower_range [of w2] | |
| 1909 | apply simp | |
| 1910 | apply (simp add: bv_to_int_utos) | |
| 1911 | using bv_to_nat_upper_range [of w1] | |
| 1912 | apply (simp add: int_nat_two_exp) | |
| 1913 | apply (rule zero_le_power,simp) | |
| 1914 | using bi1 | |
| 1915 | apply simp | |
| 1916 | done | |
| 1917 | hence "-?Q \<le> ((2::int)^length w1) * (2 ^ (length w2 - 1))" | |
| 1918 | by (simp add: zmult_ac) | |
| 1919 | thus "-(((2::int)^length w1) * (2 ^ (length w2 - Suc 0))) \<le> ?Q" | |
| 1920 | by simp | |
| 1921 | next | |
| 1922 | assume bi1: "bv_to_int (utos w1) < 0" | |
| 1923 | thus "-(((2::int)^length w1) * (2 ^ (length w2 - Suc 0))) \<le> ?Q" | |
| 1924 | by (simp add: bv_to_int_utos) | |
| 1925 | qed | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1926 | qed | 
| 23375 | 1927 | finally show "-(2 ^ (length w1 + length w2 - Suc 0)) \<le> ?Q" . | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1928 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1929 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1930 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1931 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1932 | lemma bv_smult_sym: "bv_smult w1 w2 = bv_smult w2 w1" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1933 | by (simp add: bv_smult_def zmult_ac) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1934 | |
| 14589 | 1935 | subsection {* Structural operations *}
 | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1936 | |
| 19736 | 1937 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1938 | bv_select :: "[bit list,nat] => bit" where | 
| 19736 | 1939 | "bv_select w i = w ! (length w - 1 - i)" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1940 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1941 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1942 | bv_chop :: "[bit list,nat] => bit list * bit list" where | 
| 19736 | 1943 | "bv_chop w i = (let len = length w in (take (len - i) w,drop (len - i) w))" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1944 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1945 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 1946 | bv_slice :: "[bit list,nat*nat] => bit list" where | 
| 19736 | 1947 | "bv_slice w = (\<lambda>(b,e). fst (bv_chop (snd (bv_chop w (b+1))) e))" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1948 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1949 | lemma bv_select_rev: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1950 | assumes notnull: "n < length w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1951 | shows "bv_select w n = rev w ! n" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1952 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1953 | have "\<forall>n. n < length w --> bv_select w n = rev w ! n" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1954 | proof (rule length_induct [of _ w],auto simp add: bv_select_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1955 | fix xs :: "bit list" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1956 | fix n | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1957 | assume ind: "\<forall>ys::bit list. length ys < length xs --> (\<forall>n. n < length ys --> ys ! (length ys - Suc n) = rev ys ! n)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1958 | assume notx: "n < length xs" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1959 | show "xs ! (length xs - Suc n) = rev xs ! n" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1960 | proof (cases xs) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1961 | assume "xs = []" | 
| 23375 | 1962 | with notx show ?thesis by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1963 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1964 | fix y ys | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1965 | assume [simp]: "xs = y # ys" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1966 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1967 | proof (auto simp add: nth_append) | 
| 19736 | 1968 | assume noty: "n < length ys" | 
| 1969 | from spec [OF ind,of ys] | |
| 1970 | have "\<forall>n. n < length ys --> ys ! (length ys - Suc n) = rev ys ! n" | |
| 1971 | by simp | |
| 23375 | 1972 | hence "n < length ys --> ys ! (length ys - Suc n) = rev ys ! n" .. | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32456diff
changeset | 1973 | from this and noty | 
| 23375 | 1974 | have "ys ! (length ys - Suc n) = rev ys ! n" .. | 
| 19736 | 1975 | thus "(y # ys) ! (length ys - n) = rev ys ! n" | 
| 1976 | by (simp add: nth_Cons' noty linorder_not_less [symmetric]) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1977 | next | 
| 19736 | 1978 | assume "~ n < length ys" | 
| 23375 | 1979 | hence x: "length ys \<le> n" by simp | 
| 1980 | from notx have "n < Suc (length ys)" by simp | |
| 1981 | hence "n \<le> length ys" by simp | |
| 1982 | with x have "length ys = n" by simp | |
| 1983 | thus "y = [y] ! (n - length ys)" by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1984 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1985 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1986 | qed | 
| 23375 | 1987 | then have "n < length w --> bv_select w n = rev w ! n" .. | 
| 1988 | from this and notnull show ?thesis .. | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1989 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1990 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1991 | lemma bv_chop_append: "bv_chop (w1 @ w2) (length w2) = (w1,w2)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1992 | by (simp add: bv_chop_def Let_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1993 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1994 | lemma append_bv_chop_id: "fst (bv_chop w l) @ snd (bv_chop w l) = w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1995 | by (simp add: bv_chop_def Let_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1996 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1997 | lemma bv_chop_length_fst [simp]: "length (fst (bv_chop w i)) = length w - i" | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19736diff
changeset | 1998 | by (simp add: bv_chop_def Let_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 1999 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2000 | lemma bv_chop_length_snd [simp]: "length (snd (bv_chop w i)) = min i (length w)" | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19736diff
changeset | 2001 | by (simp add: bv_chop_def Let_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2002 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2003 | lemma bv_slice_length [simp]: "[| j \<le> i; i < length w |] ==> length (bv_slice w (i,j)) = i - j + 1" | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19736diff
changeset | 2004 | by (auto simp add: bv_slice_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2005 | |
| 19736 | 2006 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 2007 | length_nat :: "nat => nat" where | 
| 28562 | 2008 | [code del]: "length_nat x = (LEAST n. x < 2 ^ n)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2009 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2010 | lemma length_nat: "length (nat_to_bv n) = length_nat n" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2011 | apply (simp add: length_nat_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2012 | apply (rule Least_equality [symmetric]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2013 | prefer 2 | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2014 | apply (rule length_nat_to_bv_upper_limit) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2015 | apply arith | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2016 | apply (rule ccontr) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2017 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2018 | assume "~ n < 2 ^ length (nat_to_bv n)" | 
| 23375 | 2019 | hence "2 ^ length (nat_to_bv n) \<le> n" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2020 | hence "length (nat_to_bv n) < length (nat_to_bv n)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2021 | by (rule length_nat_to_bv_lower_limit) | 
| 23375 | 2022 | thus False by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2023 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2024 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2025 | lemma length_nat_0 [simp]: "length_nat 0 = 0" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2026 | by (simp add: length_nat_def Least_equality) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2027 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2028 | lemma length_nat_non0: | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 2029 | assumes n0: "n \<noteq> 0" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2030 | shows "length_nat n = Suc (length_nat (n div 2))" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2031 | apply (simp add: length_nat [symmetric]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2032 | apply (subst nat_to_bv_non0 [of n]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2033 | apply (simp_all add: n0) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2034 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2035 | |
| 19736 | 2036 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 2037 | length_int :: "int => nat" where | 
| 19736 | 2038 | "length_int x = | 
| 2039 | (if 0 < x then Suc (length_nat (nat x)) | |
| 2040 | else if x = 0 then 0 | |
| 2041 | else Suc (length_nat (nat (-x - 1))))" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2042 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2043 | lemma length_int: "length (int_to_bv i) = length_int i" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2044 | proof (cases "0 < i") | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2045 | assume i0: "0 < i" | 
| 23375 | 2046 | hence "length (int_to_bv i) = | 
| 2047 | length (norm_signed (\<zero> # norm_unsigned (nat_to_bv (nat i))))" by simp | |
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2048 | also from norm_unsigned_result [of "nat_to_bv (nat i)"] | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2049 | have "... = Suc (length_nat (nat i))" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2050 | apply safe | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2051 | apply (simp del: norm_unsigned_nat_to_bv) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2052 | apply (drule norm_empty_bv_to_nat_zero) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2053 | using prems | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2054 | apply simp | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2055 | apply (cases "norm_unsigned (nat_to_bv (nat i))") | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2056 | apply (drule norm_empty_bv_to_nat_zero [of "nat_to_bv (nat i)"]) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2057 | using prems | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2058 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2059 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2060 | using prems | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2061 | apply (simp add: length_nat [symmetric]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2062 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2063 | finally show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2064 | using i0 | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2065 | by (simp add: length_int_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2066 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2067 | assume "~ 0 < i" | 
| 23375 | 2068 | hence i0: "i \<le> 0" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2069 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2070 | proof (cases "i = 0") | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2071 | assume "i = 0" | 
| 23375 | 2072 | thus ?thesis by (simp add: length_int_def) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2073 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2074 | assume "i \<noteq> 0" | 
| 23375 | 2075 | with i0 have i0: "i < 0" by simp | 
| 2076 | hence "length (int_to_bv i) = | |
| 2077 | length (norm_signed (\<one> # bv_not (norm_unsigned (nat_to_bv (nat (- i) - 1)))))" | |
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2078 | by (simp add: int_to_bv_def nat_diff_distrib) | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2079 | also from norm_unsigned_result [of "nat_to_bv (nat (- i) - 1)"] | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2080 | have "... = Suc (length_nat (nat (- i) - 1))" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2081 | apply safe | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2082 | apply (simp del: norm_unsigned_nat_to_bv) | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2083 | apply (drule norm_empty_bv_to_nat_zero [of "nat_to_bv (nat (-i) - Suc 0)"]) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2084 | using prems | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2085 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2086 | apply (cases "- i - 1 = 0") | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2087 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2088 | apply (simp add: length_nat [symmetric]) | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2089 | apply (cases "norm_unsigned (nat_to_bv (nat (- i) - 1))") | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2090 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2091 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2092 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2093 | finally | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2094 | show ?thesis | 
| 23375 | 2095 | using i0 by (simp add: length_int_def nat_diff_distrib del: int_to_bv_lt0) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2096 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2097 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2098 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2099 | lemma length_int_0 [simp]: "length_int 0 = 0" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2100 | by (simp add: length_int_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2101 | |
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2102 | lemma length_int_gt0: "0 < i ==> length_int i = Suc (length_nat (nat i))" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2103 | by (simp add: length_int_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2104 | |
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2105 | lemma length_int_lt0: "i < 0 ==> length_int i = Suc (length_nat (nat (- i) - 1))" | 
| 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2106 | by (simp add: length_int_def nat_diff_distrib) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2107 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2108 | lemma bv_chopI: "[| w = w1 @ w2 ; i = length w2 |] ==> bv_chop w i = (w1,w2)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2109 | by (simp add: bv_chop_def Let_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2110 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2111 | lemma bv_sliceI: "[| j \<le> i ; i < length w ; w = w1 @ w2 @ w3 ; Suc i = length w2 + j ; j = length w3 |] ==> bv_slice w (i,j) = w2" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2112 | apply (simp add: bv_slice_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2113 | apply (subst bv_chopI [of "w1 @ w2 @ w3" w1 "w2 @ w3"]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2114 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2115 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2116 | apply simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2117 | apply (subst bv_chopI [of "w2 @ w3" w2 w3],simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2118 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2119 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2120 | lemma bv_slice_bv_slice: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2121 | assumes ki: "k \<le> i" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2122 | and ij: "i \<le> j" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2123 | and jl: "j \<le> l" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2124 | and lw: "l < length w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2125 | shows "bv_slice w (j,i) = bv_slice (bv_slice w (l,k)) (j-k,i-k)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2126 | proof - | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2127 | def w1 == "fst (bv_chop w (Suc l))" | 
| 19736 | 2128 | and w2 == "fst (bv_chop (snd (bv_chop w (Suc l))) (Suc j))" | 
| 2129 | and w3 == "fst (bv_chop (snd (bv_chop (snd (bv_chop w (Suc l))) (Suc j))) i)" | |
| 2130 | and w4 == "fst (bv_chop (snd (bv_chop (snd (bv_chop (snd (bv_chop w (Suc l))) (Suc j))) i)) k)" | |
| 2131 | and w5 == "snd (bv_chop (snd (bv_chop (snd (bv_chop (snd (bv_chop w (Suc l))) (Suc j))) i)) k)" | |
| 2132 | note w_defs = this | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2133 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2134 | have w_def: "w = w1 @ w2 @ w3 @ w4 @ w5" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2135 | by (simp add: w_defs append_bv_chop_id) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2136 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2137 | from ki ij jl lw | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2138 | show ?thesis | 
| 15488 | 2139 | apply (subst bv_sliceI [where ?j = i and ?i = j and ?w = w and ?w1.0 = "w1 @ w2" and ?w2.0 = w3 and ?w3.0 = "w4 @ w5"]) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2140 | apply simp_all | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2141 | apply (rule w_def) | 
| 32438 | 2142 | apply (simp add: w_defs) | 
| 2143 | apply (simp add: w_defs) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2144 | apply (subst bv_sliceI [where ?j = k and ?i = l and ?w = w and ?w1.0 = w1 and ?w2.0 = "w2 @ w3 @ w4" and ?w3.0 = w5]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2145 | apply simp_all | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2146 | apply (rule w_def) | 
| 32438 | 2147 | apply (simp add: w_defs) | 
| 2148 | apply (simp add: w_defs) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2149 | apply (subst bv_sliceI [where ?j = "i-k" and ?i = "j-k" and ?w = "w2 @ w3 @ w4" and ?w1.0 = w2 and ?w2.0 = w3 and ?w3.0 = w4]) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2150 | apply simp_all | 
| 32438 | 2151 | apply (simp_all add: w_defs) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2152 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2153 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2154 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2155 | lemma bv_to_nat_extend [simp]: "bv_to_nat (bv_extend n \<zero> w) = bv_to_nat w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2156 | apply (simp add: bv_extend_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2157 | apply (subst bv_to_nat_dist_append) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2158 | apply simp | 
| 34915 | 2159 |   apply (induct ("n - length w"))
 | 
| 19736 | 2160 | apply simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2161 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2162 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2163 | lemma bv_msb_extend_same [simp]: "bv_msb w = b ==> bv_msb (bv_extend n b w) = b" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2164 | apply (simp add: bv_extend_def) | 
| 34915 | 2165 | apply (cases "n - length w") | 
| 19736 | 2166 | apply simp_all | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2167 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2168 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2169 | lemma bv_to_int_extend [simp]: | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2170 | assumes a: "bv_msb w = b" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2171 | shows "bv_to_int (bv_extend n b w) = bv_to_int w" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2172 | proof (cases "bv_msb w") | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2173 | assume [simp]: "bv_msb w = \<zero>" | 
| 23375 | 2174 | with a have [simp]: "b = \<zero>" by simp | 
| 2175 | show ?thesis by (simp add: bv_to_int_def) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2176 | next | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2177 | assume [simp]: "bv_msb w = \<one>" | 
| 23375 | 2178 | with a have [simp]: "b = \<one>" by simp | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2179 | show ?thesis | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2180 | apply (simp add: bv_to_int_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2181 | apply (simp add: bv_extend_def) | 
| 34915 | 2182 |     apply (induct ("n - length w"), simp_all)
 | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2183 | done | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2184 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2185 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2186 | lemma length_nat_mono [simp]: "x \<le> y ==> length_nat x \<le> length_nat y" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2187 | proof (rule ccontr) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2188 | assume xy: "x \<le> y" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2189 | assume "~ length_nat x \<le> length_nat y" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2190 | hence lxly: "length_nat y < length_nat x" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2191 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2192 | hence "length_nat y < (LEAST n. x < 2 ^ n)" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2193 | by (simp add: length_nat_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2194 | hence "~ x < 2 ^ length_nat y" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2195 | by (rule not_less_Least) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2196 | hence xx: "2 ^ length_nat y \<le> x" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2197 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2198 | have yy: "y < 2 ^ length_nat y" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2199 | apply (simp add: length_nat_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2200 | apply (rule LeastI) | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2201 | apply (subgoal_tac "y < 2 ^ y",assumption) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2202 | apply (cases "0 \<le> y") | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2203 | apply (induct y,simp_all) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2204 | done | 
| 23375 | 2205 | with xx have "y < x" by simp | 
| 2206 | with xy show False by simp | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2207 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2208 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2209 | lemma length_nat_mono_int [simp]: "x \<le> y ==> length_nat x \<le> length_nat y" | 
| 23375 | 2210 | by (rule length_nat_mono) arith | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2211 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2212 | lemma length_nat_pos [simp,intro!]: "0 < x ==> 0 < length_nat x" | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2213 | by (simp add: length_nat_non0) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2214 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2215 | lemma length_int_mono_gt0: "[| 0 \<le> x ; x \<le> y |] ==> length_int x \<le> length_int y" | 
| 23375 | 2216 | by (cases "x = 0") (simp_all add: length_int_gt0 nat_le_eq_zle) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2217 | |
| 23375 | 2218 | lemma length_int_mono_lt0: "[| x \<le> y ; y \<le> 0 |] ==> length_int y \<le> length_int x" | 
| 2219 | by (cases "y = 0") (simp_all add: length_int_lt0) | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2220 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2221 | lemmas [simp] = length_nat_non0 | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2222 | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25595diff
changeset | 2223 | lemma "nat_to_bv (number_of Int.Pls) = []" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2224 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2225 | |
| 34941 | 2226 | primrec fast_bv_to_nat_helper :: "[bit list, int] => int" where | 
| 2227 | fast_bv_to_nat_Nil: "fast_bv_to_nat_helper [] k = k" | |
| 2228 | | fast_bv_to_nat_Cons: "fast_bv_to_nat_helper (b#bs) k = | |
| 2229 | fast_bv_to_nat_helper bs ((bit_case Int.Bit0 Int.Bit1 b) k)" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2230 | |
| 28562 | 2231 | declare fast_bv_to_nat_helper.simps [code del] | 
| 28229 | 2232 | |
| 23375 | 2233 | lemma fast_bv_to_nat_Cons0: "fast_bv_to_nat_helper (\<zero>#bs) bin = | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 2234 | fast_bv_to_nat_helper bs (Int.Bit0 bin)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2235 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2236 | |
| 23375 | 2237 | lemma fast_bv_to_nat_Cons1: "fast_bv_to_nat_helper (\<one>#bs) bin = | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 2238 | fast_bv_to_nat_helper bs (Int.Bit1 bin)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2239 | by simp | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2240 | |
| 23375 | 2241 | lemma fast_bv_to_nat_def: | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25595diff
changeset | 2242 | "bv_to_nat bs == number_of (fast_bv_to_nat_helper bs Int.Pls)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2243 | proof (simp add: bv_to_nat_def) | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2244 | have "\<forall> bin. \<not> (neg (number_of bin :: int)) \<longrightarrow> (foldl (%bn b. 2 * bn + bitval b) (number_of bin) bs) = number_of (fast_bv_to_nat_helper bs bin)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2245 | apply (induct bs,simp) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2246 | apply (case_tac a,simp_all) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2247 | done | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25595diff
changeset | 2248 | thus "foldl (\<lambda>bn b. 2 * bn + bitval b) 0 bs \<equiv> number_of (fast_bv_to_nat_helper bs Int.Pls)" | 
| 15325 
50ac7d2c34c9
Functions nat_to_bv and bv_to_nat now really operate on natural numbers.
 berghofe parents: 
15140diff
changeset | 2249 | by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric]) | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2250 | qed | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2251 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2252 | declare fast_bv_to_nat_Cons [simp del] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2253 | declare fast_bv_to_nat_Cons0 [simp] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2254 | declare fast_bv_to_nat_Cons1 [simp] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2255 | |
| 22993 | 2256 | setup {*
 | 
| 2257 | (*comments containing lcp are the removal of fast_bv_of_nat*) | |
| 2258 | let | |
| 2259 |   fun is_const_bool (Const("True",_)) = true
 | |
| 2260 |     | is_const_bool (Const("False",_)) = true
 | |
| 2261 | | is_const_bool _ = false | |
| 2262 |   fun is_const_bit (Const("Word.bit.Zero",_)) = true
 | |
| 2263 |     | is_const_bit (Const("Word.bit.One",_)) = true
 | |
| 2264 | | is_const_bit _ = false | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25595diff
changeset | 2265 |   fun num_is_usable (Const(@{const_name Int.Pls},_)) = true
 | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25595diff
changeset | 2266 |     | num_is_usable (Const(@{const_name Int.Min},_)) = false
 | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 2267 |     | num_is_usable (Const(@{const_name Int.Bit0},_) $ w) =
 | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 2268 | num_is_usable w | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 2269 |     | num_is_usable (Const(@{const_name Int.Bit1},_) $ w) =
 | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 2270 | num_is_usable w | 
| 22993 | 2271 | | num_is_usable _ = false | 
| 2272 |   fun vec_is_usable (Const("List.list.Nil",_)) = true
 | |
| 2273 |     | vec_is_usable (Const("List.list.Cons",_) $ b $ bs) =
 | |
| 2274 | vec_is_usable bs andalso is_const_bit b | |
| 2275 | | vec_is_usable _ = false | |
| 2276 | (*lcp** val fast1_th = PureThy.get_thm thy "Word.fast_nat_to_bv_def"*) | |
| 2277 |   val fast2_th = @{thm "Word.fast_bv_to_nat_def"};
 | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25595diff
changeset | 2278 |   (*lcp** fun f sg thms (Const("Word.nat_to_bv",_) $ (Const(@{const_name Int.number_of},_) $ t)) =
 | 
| 22993 | 2279 | if num_is_usable t | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25595diff
changeset | 2280 |       then SOME (Drule.cterm_instantiate [(cterm_of sg (Var (("w", 0), @{typ int})), cterm_of sg t)] fast1_th)
 | 
| 22993 | 2281 | else NONE | 
| 2282 | | f _ _ _ = NONE *) | |
| 2283 |   fun g sg thms (Const("Word.bv_to_nat",_) $ (t as (Const("List.list.Cons",_) $ _ $ _))) =
 | |
| 2284 | if vec_is_usable t then | |
| 2285 |           SOME (Drule.cterm_instantiate [(cterm_of sg (Var(("bs",0),Type("List.list",[Type("Word.bit",[])]))),cterm_of sg t)] fast2_th)
 | |
| 2286 | else NONE | |
| 2287 | | g _ _ _ = NONE | |
| 2288 | (*lcp** val simproc1 = Simplifier.simproc thy "nat_to_bv" ["Word.nat_to_bv (number_of w)"] f *) | |
| 2289 |   val simproc2 = Simplifier.simproc @{theory} "bv_to_nat" ["Word.bv_to_nat (x # xs)"] g
 | |
| 2290 | in | |
| 26496 
49ae9456eba9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26086diff
changeset | 2291 | Simplifier.map_simpset (fn ss => ss addsimprocs [(*lcp*simproc1,*)simproc2]) | 
| 22993 | 2292 | end*} | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2293 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2294 | declare bv_to_nat1 [simp del] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2295 | declare bv_to_nat_helper [simp del] | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2296 | |
| 19736 | 2297 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 2298 | bv_mapzip :: "[bit => bit => bit,bit list, bit list] => bit list" where | 
| 19736 | 2299 | "bv_mapzip f w1 w2 = | 
| 2300 | (let g = bv_extend (max (length w1) (length w2)) \<zero> | |
| 2301 | in map (split f) (zip (g w1) (g w2)))" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2302 | |
| 19736 | 2303 | lemma bv_length_bv_mapzip [simp]: | 
| 23375 | 2304 | "length (bv_mapzip f w1 w2) = max (length w1) (length w2)" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2305 | by (simp add: bv_mapzip_def Let_def split: split_max) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2306 | |
| 17650 
44b135d04cc4
Made sure all lemmas now have names (especially so that certain of them
 skalberg parents: 
16796diff
changeset | 2307 | lemma bv_mapzip_Nil [simp]: "bv_mapzip f [] [] = []" | 
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2308 | by (simp add: bv_mapzip_def Let_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2309 | |
| 19736 | 2310 | lemma bv_mapzip_Cons [simp]: "length w1 = length w2 ==> | 
| 2311 | bv_mapzip f (x#w1) (y#w2) = f x y # bv_mapzip f w1 w2" | |
| 14494 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2312 | by (simp add: bv_mapzip_def Let_def) | 
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2313 | |
| 
48ae8d678d88
Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
 skalberg parents: diff
changeset | 2314 | end |