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