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