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