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