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