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