src/HOL/Word/WordShift.thy
author wenzelm
Mon, 16 Mar 2009 18:24:30 +0100
changeset 30549 d2d7874648bd
parent 30042 31039ee583fa
child 30952 7ab2716dd93b
permissions -rw-r--r--
simplified method setup;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     1
(* 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     2
    Author:     Jeremy Dawson and Gerwin Klein, NICTA
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     3
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     4
  contains theorems to do with shifting, rotating, splitting words
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     5
*)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     6
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
     7
header {* Shifting, Rotating, and Splitting Words *}
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
     8
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
     9
theory WordShift
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
    10
imports WordBitwise
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
    11
begin
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    12
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
    13
subsection "Bit shifting"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    14
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    15
lemma shiftl1_number [simp] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    16
  "shiftl1 (number_of w) = number_of (w BIT bit.B0)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    17
  apply (unfold shiftl1_def word_number_of_def)
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26072
diff changeset
    18
  apply (simp add: word_ubin.norm_eq_iff [symmetric] word_ubin.eq_norm
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26072
diff changeset
    19
              del: BIT_simps)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    20
  apply (subst refl [THEN bintrunc_BIT_I, symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    21
  apply (subst bintrunc_bintrunc_min)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    22
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    23
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    24
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    25
lemma shiftl1_0 [simp] : "shiftl1 0 = 0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    26
  unfolding word_0_no shiftl1_number by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    27
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    28
lemmas shiftl1_def_u = shiftl1_def [folded word_number_of_def]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    29
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    30
lemma shiftl1_def_s: "shiftl1 w = number_of (sint w BIT bit.B0)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    31
  by (rule trans [OF _ shiftl1_number]) simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    32
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    33
lemma shiftr1_0 [simp] : "shiftr1 0 = 0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    34
  unfolding shiftr1_def 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    35
  by simp (simp add: word_0_wi)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    36
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    37
lemma sshiftr1_0 [simp] : "sshiftr1 0 = 0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    38
  apply (unfold sshiftr1_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    39
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    40
  apply (simp add : word_0_wi)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    41
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    42
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    43
lemma sshiftr1_n1 [simp] : "sshiftr1 -1 = -1"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    44
  unfolding sshiftr1_def by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    45
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
    46
lemma shiftl_0 [simp] : "(0::'a::len0 word) << n = 0"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    47
  unfolding shiftl_def by (induct n) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    48
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
    49
lemma shiftr_0 [simp] : "(0::'a::len0 word) >> n = 0"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    50
  unfolding shiftr_def by (induct n) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    51
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    52
lemma sshiftr_0 [simp] : "0 >>> n = 0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    53
  unfolding sshiftr_def by (induct n) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    54
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    55
lemma sshiftr_n1 [simp] : "-1 >>> n = -1"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    56
  unfolding sshiftr_def by (induct n) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    57
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    58
lemma nth_shiftl1: "shiftl1 w !! n = (n < size w & n > 0 & w !! (n - 1))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    59
  apply (unfold shiftl1_def word_test_bit_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    60
  apply (simp add: nth_bintr word_ubin.eq_norm word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    61
  apply (cases n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    62
   apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    63
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    64
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    65
lemma nth_shiftl' [rule_format]:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
    66
  "ALL n. ((w::'a::len0 word) << m) !! n = (n < size w & n >= m & w !! (n - m))"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    67
  apply (unfold shiftl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    68
  apply (induct "m")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    69
   apply (force elim!: test_bit_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    70
  apply (clarsimp simp add : nth_shiftl1 word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    71
  apply arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    72
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    73
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    74
lemmas nth_shiftl = nth_shiftl' [unfolded word_size] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    75
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    76
lemma nth_shiftr1: "shiftr1 w !! n = w !! Suc n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    77
  apply (unfold shiftr1_def word_test_bit_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    78
  apply (simp add: nth_bintr word_ubin.eq_norm)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    79
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    80
  apply (drule bin_nth.Suc [THEN iffD2, THEN bin_nth_uint_imp])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    81
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    82
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    83
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    84
lemma nth_shiftr: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
    85
  "\<And>n. ((w::'a::len0 word) >> m) !! n = w !! (n + m)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    86
  apply (unfold shiftr_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    87
  apply (induct "m")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    88
   apply (auto simp add : nth_shiftr1)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    89
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    90
   
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    91
(* see paper page 10, (1), (2), shiftr1_def is of the form of (1),
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    92
  where f (ie bin_rest) takes normal arguments to normal results,
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    93
  thus we get (2) from (1) *)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    94
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    95
lemma uint_shiftr1: "uint (shiftr1 w) = bin_rest (uint w)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    96
  apply (unfold shiftr1_def word_ubin.eq_norm bin_rest_trunc_i)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    97
  apply (subst bintr_uint [symmetric, OF order_refl])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    98
  apply (simp only : bintrunc_bintrunc_l)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    99
  apply simp 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   100
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   101
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   102
lemma nth_sshiftr1: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   103
  "sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   104
  apply (unfold sshiftr1_def word_test_bit_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   105
  apply (simp add: nth_bintr word_ubin.eq_norm
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   106
                   bin_nth.Suc [symmetric] word_size 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   107
             del: bin_nth.simps)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   108
  apply (simp add: nth_bintr uint_sint del : bin_nth.simps)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   109
  apply (auto simp add: bin_nth_sint)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   110
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   111
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   112
lemma nth_sshiftr [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   113
  "ALL n. sshiftr w m !! n = (n < size w & 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   114
    (if n + m >= size w then w !! (size w - 1) else w !! (n + m)))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   115
  apply (unfold sshiftr_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   116
  apply (induct_tac "m")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   117
   apply (simp add: test_bit_bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   118
  apply (clarsimp simp add: nth_sshiftr1 word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   119
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   120
       apply arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   121
      apply arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   122
     apply (erule thin_rl)
27136
06a8f65e32f6 recovered nat_induct as default for induct_tac;
wenzelm
parents: 27132
diff changeset
   123
     apply (case_tac n)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   124
      apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   125
      apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   126
     apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   127
    apply (erule thin_rl)
27136
06a8f65e32f6 recovered nat_induct as default for induct_tac;
wenzelm
parents: 27132
diff changeset
   128
    apply (case_tac n)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   129
     apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   130
     apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   131
    apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   132
   apply arith+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   133
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   134
    
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   135
lemma shiftr1_div_2: "uint (shiftr1 w) = uint w div 2"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   136
  apply (unfold shiftr1_def bin_rest_div)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   137
  apply (rule word_uint.Abs_inverse)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   138
  apply (simp add: uints_num pos_imp_zdiv_nonneg_iff)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   139
  apply (rule xtr7)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   140
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   141
   apply (rule zdiv_le_dividend)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   142
    apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   143
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   144
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   145
lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   146
  apply (unfold sshiftr1_def bin_rest_div [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   147
  apply (simp add: word_sbin.eq_norm)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   148
  apply (rule trans)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   149
   defer
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   150
   apply (subst word_sbin.norm_Rep [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   151
   apply (rule refl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   152
  apply (subst word_sbin.norm_Rep [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   153
  apply (unfold One_nat_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   154
  apply (rule sbintrunc_rest)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   155
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   156
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   157
lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   158
  apply (unfold shiftr_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   159
  apply (induct "n")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   160
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   161
  apply (simp add: shiftr1_div_2 mult_commute
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   162
                   zdiv_zmult2_eq [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   163
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   164
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   165
lemma sshiftr_div_2n: "sint (sshiftr w n) = sint w div 2 ^ n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   166
  apply (unfold sshiftr_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   167
  apply (induct "n")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   168
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   169
  apply (simp add: sshiftr1_div_2 mult_commute
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   170
                   zdiv_zmult2_eq [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   171
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   172
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
   173
subsubsection "shift functions in terms of lists of bools"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   174
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   175
lemmas bshiftr1_no_bin [simp] = 
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   176
  bshiftr1_def [where w="number_of w", unfolded to_bl_no_bin, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   177
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   178
lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   179
  unfolding bshiftr1_def by (rule word_bl.Abs_inverse) simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   180
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   181
lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   182
  unfolding uint_bl of_bl_no 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   183
  by (simp add: bl_to_bin_aux_append bl_to_bin_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   184
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   185
lemma shiftl1_bl: "shiftl1 (w::'a::len0 word) = of_bl (to_bl w @ [False])"
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   186
proof -
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   187
  have "shiftl1 w = shiftl1 (of_bl (to_bl w))" by simp
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   188
  also have "\<dots> = of_bl (to_bl w @ [False])" by (rule shiftl1_of_bl)
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   189
  finally show ?thesis .
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   190
qed
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   191
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   192
lemma bl_shiftl1:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   193
  "to_bl (shiftl1 (w :: 'a :: len word)) = tl (to_bl w) @ [False]"
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   194
  apply (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons')
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   195
  apply (fast intro!: Suc_leI)
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   196
  done
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   197
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   198
lemma shiftr1_bl: "shiftr1 w = of_bl (butlast (to_bl w))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   199
  apply (unfold shiftr1_def uint_bl of_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   200
  apply (simp add: butlast_rest_bin word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   201
  apply (simp add: bin_rest_trunc [symmetric, unfolded One_nat_def])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   202
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   203
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   204
lemma bl_shiftr1: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   205
  "to_bl (shiftr1 (w :: 'a :: len word)) = False # butlast (to_bl w)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   206
  unfolding shiftr1_bl
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   207
  by (simp add : word_rep_drop len_gt_0 [THEN Suc_leI])
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   208
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   209
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   210
(* relate the two above : TODO - remove the :: len restriction on
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   211
  this theorem and others depending on it *)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   212
lemma shiftl1_rev: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   213
  "shiftl1 (w :: 'a :: len word) = word_reverse (shiftr1 (word_reverse w))"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   214
  apply (unfold word_reverse_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   215
  apply (rule word_bl.Rep_inverse' [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   216
  apply (simp add: bl_shiftl1 bl_shiftr1 word_bl.Abs_inverse)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   217
  apply (cases "to_bl w")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   218
   apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   219
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   220
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   221
lemma shiftl_rev: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   222
  "shiftl (w :: 'a :: len word) n = word_reverse (shiftr (word_reverse w) n)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   223
  apply (unfold shiftl_def shiftr_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   224
  apply (induct "n")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   225
   apply (auto simp add : shiftl1_rev)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   226
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   227
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   228
lemmas rev_shiftl =
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   229
  shiftl_rev [where w = "word_reverse w", simplified, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   230
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   231
lemmas shiftr_rev = rev_shiftl [THEN word_rev_gal', standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   232
lemmas rev_shiftr = shiftl_rev [THEN word_rev_gal', standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   233
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   234
lemma bl_sshiftr1:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   235
  "to_bl (sshiftr1 (w :: 'a :: len word)) = hd (to_bl w) # butlast (to_bl w)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   236
  apply (unfold sshiftr1_def uint_bl word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   237
  apply (simp add: butlast_rest_bin word_ubin.eq_norm)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   238
  apply (simp add: sint_uint)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   239
  apply (rule nth_equalityI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   240
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   241
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   242
  apply (case_tac i)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   243
   apply (simp_all add: hd_conv_nth length_0_conv [symmetric]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   244
                        nth_bin_to_bl bin_nth.Suc [symmetric] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   245
                        nth_sbintr 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   246
                   del: bin_nth.Suc)
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   247
   apply force
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   248
  apply (rule impI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   249
  apply (rule_tac f = "bin_nth (uint w)" in arg_cong)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   250
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   251
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   252
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   253
lemma drop_shiftr: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   254
  "drop n (to_bl ((w :: 'a :: len word) >> n)) = take (size w - n) (to_bl w)" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   255
  apply (unfold shiftr_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   256
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   257
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   258
   apply (simp add: drop_Suc bl_shiftr1 butlast_drop [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   259
   apply (rule butlast_take [THEN trans])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   260
  apply (auto simp: word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   261
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   262
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   263
lemma drop_sshiftr: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   264
  "drop n (to_bl ((w :: 'a :: len word) >>> n)) = take (size w - n) (to_bl w)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   265
  apply (unfold sshiftr_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   266
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   267
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   268
   apply (simp add: drop_Suc bl_sshiftr1 butlast_drop [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   269
   apply (rule butlast_take [THEN trans])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   270
  apply (auto simp: word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   271
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   272
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   273
lemma take_shiftr [rule_format] :
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   274
  "n <= size (w :: 'a :: len word) --> take n (to_bl (w >> n)) = 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   275
    replicate n False" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   276
  apply (unfold shiftr_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   277
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   278
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   279
   apply (simp add: bl_shiftr1)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   280
   apply (rule impI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   281
   apply (rule take_butlast [THEN trans])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   282
  apply (auto simp: word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   283
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   284
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   285
lemma take_sshiftr' [rule_format] :
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   286
  "n <= size (w :: 'a :: len word) --> hd (to_bl (w >>> n)) = hd (to_bl w) & 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   287
    take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   288
  apply (unfold sshiftr_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   289
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   290
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   291
   apply (simp add: bl_sshiftr1)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   292
   apply (rule impI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   293
   apply (rule take_butlast [THEN trans])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   294
  apply (auto simp: word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   295
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   296
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   297
lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   298
lemmas take_sshiftr = take_sshiftr' [THEN conjunct2, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   299
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   300
lemma atd_lem: "take n xs = t ==> drop n xs = d ==> xs = t @ d"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   301
  by (auto intro: append_take_drop_id [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   302
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   303
lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   304
lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   305
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   306
lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   307
  unfolding shiftl_def
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   308
  by (induct n) (auto simp: shiftl1_of_bl replicate_app_Cons_same)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   309
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   310
lemma shiftl_bl:
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   311
  "(w::'a::len0 word) << (n::nat) = of_bl (to_bl w @ replicate n False)"
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   312
proof -
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   313
  have "w << n = of_bl (to_bl w) << n" by simp
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   314
  also have "\<dots> = of_bl (to_bl w @ replicate n False)" by (rule shiftl_of_bl)
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   315
  finally show ?thesis .
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   316
qed
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   317
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   318
lemmas shiftl_number [simp] = shiftl_def [where w="number_of w", standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   319
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   320
lemma bl_shiftl:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   321
  "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   322
  by (simp add: shiftl_bl word_rep_drop word_size min_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   323
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   324
lemma shiftl_zero_size: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   325
  fixes x :: "'a::len0 word"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   326
  shows "size x <= n ==> x << n = 0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   327
  apply (unfold word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   328
  apply (rule word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   329
  apply (clarsimp simp add: shiftl_bl word_size test_bit_of_bl nth_append)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   330
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   331
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   332
(* note - the following results use 'a :: len word < number_ring *)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   333
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   334
lemma shiftl1_2t: "shiftl1 (w :: 'a :: len word) = 2 * w"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   335
  apply (simp add: shiftl1_def_u)
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26072
diff changeset
   336
  apply (simp only:  double_number_of_Bit0 [symmetric])
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   337
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   338
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   339
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   340
lemma shiftl1_p: "shiftl1 (w :: 'a :: len word) = w + w"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   341
  apply (simp add: shiftl1_def_u)
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26072
diff changeset
   342
  apply (simp only: double_number_of_Bit0 [symmetric])
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   343
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   344
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   345
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   346
lemma shiftl_t2n: "shiftl (w :: 'a :: len word) n = 2 ^ n * w"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   347
  unfolding shiftl_def 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   348
  by (induct n) (auto simp: shiftl1_2t power_Suc)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   349
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   350
lemma shiftr1_bintr [simp]:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   351
  "(shiftr1 (number_of w) :: 'a :: len0 word) = 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   352
    number_of (bin_rest (bintrunc (len_of TYPE ('a)) w))" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   353
  unfolding shiftr1_def word_number_of_def
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   354
  by (simp add : word_ubin.eq_norm)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   355
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   356
lemma sshiftr1_sbintr [simp] :
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   357
  "(sshiftr1 (number_of w) :: 'a :: len word) = 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   358
    number_of (bin_rest (sbintrunc (len_of TYPE ('a) - 1) w))" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   359
  unfolding sshiftr1_def word_number_of_def
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   360
  by (simp add : word_sbin.eq_norm)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   361
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   362
lemma shiftr_no': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   363
  "w = number_of bin ==> 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   364
  (w::'a::len0 word) >> n = number_of ((bin_rest ^ n) (bintrunc (size w) bin))"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   365
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   366
  apply (rule word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   367
  apply (auto simp: nth_shiftr nth_rest_power_bin nth_bintr word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   368
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   369
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   370
lemma sshiftr_no': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   371
  "w = number_of bin ==> w >>> n = number_of ((bin_rest ^ n) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   372
    (sbintrunc (size w - 1) bin))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   373
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   374
  apply (rule word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   375
  apply (auto simp: nth_sshiftr nth_rest_power_bin nth_sbintr word_size)
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   376
   apply (subgoal_tac "na + n = len_of TYPE('a) - Suc 0", simp, simp)+
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   377
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   378
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   379
lemmas sshiftr_no [simp] = 
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   380
  sshiftr_no' [where w = "number_of w", OF refl, unfolded word_size, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   381
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   382
lemmas shiftr_no [simp] = 
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   383
  shiftr_no' [where w = "number_of w", OF refl, unfolded word_size, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   384
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   385
lemma shiftr1_bl_of': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   386
  "us = shiftr1 (of_bl bl) ==> length bl <= size us ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   387
    us = of_bl (butlast bl)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   388
  by (clarsimp simp: shiftr1_def of_bl_def word_size butlast_rest_bl2bin 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   389
                     word_ubin.eq_norm trunc_bl2bin)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   390
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   391
lemmas shiftr1_bl_of = refl [THEN shiftr1_bl_of', unfolded word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   392
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   393
lemma shiftr_bl_of' [rule_format]: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   394
  "us = of_bl bl >> n ==> length bl <= size us --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   395
   us = of_bl (take (length bl - n) bl)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   396
  apply (unfold shiftr_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   397
  apply hypsubst
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   398
  apply (unfold word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   399
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   400
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   401
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   402
  apply (subst shiftr1_bl_of)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   403
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   404
  apply (simp add: butlast_take)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   405
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   406
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   407
lemmas shiftr_bl_of = refl [THEN shiftr_bl_of', unfolded word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   408
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   409
lemmas shiftr_bl = word_bl.Rep' [THEN eq_imp_le, THEN shiftr_bl_of,
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   410
  simplified word_size, simplified, THEN eq_reflection, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   411
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   412
lemma msb_shift': "msb (w::'a::len word) <-> (w >> (size w - 1)) ~= 0"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   413
  apply (unfold shiftr_bl word_msb_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   414
  apply (simp add: word_size Suc_le_eq take_Suc)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   415
  apply (cases "hd (to_bl w)")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   416
   apply (auto simp: word_1_bl word_0_bl 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   417
                     of_bl_rep_False [where n=1 and bs="[]", simplified])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   418
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   419
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   420
lemmas msb_shift = msb_shift' [unfolded word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   421
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   422
lemma align_lem_or [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   423
  "ALL x m. length x = n + m --> length y = n + m --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   424
    drop m x = replicate n False --> take m y = replicate m False --> 
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
   425
    map2 op | x y = take m x @ drop m y"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   426
  apply (induct_tac y)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   427
   apply force
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   428
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   429
  apply (case_tac x, force)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   430
  apply (case_tac m, auto)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   431
  apply (drule sym)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   432
  apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   433
  apply (induct_tac list, auto)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   434
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   435
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   436
lemma align_lem_and [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   437
  "ALL x m. length x = n + m --> length y = n + m --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   438
    drop m x = replicate n False --> take m y = replicate m False --> 
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
   439
    map2 op & x y = replicate (n + m) False"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   440
  apply (induct_tac y)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   441
   apply force
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   442
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   443
  apply (case_tac x, force)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   444
  apply (case_tac m, auto)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   445
  apply (drule sym)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   446
  apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   447
  apply (induct_tac list, auto)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   448
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   449
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   450
lemma aligned_bl_add_size':
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   451
  "size x - n = m ==> n <= size x ==> drop m (to_bl x) = replicate n False ==>
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   452
    take m (to_bl y) = replicate m False ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   453
    to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   454
  apply (subgoal_tac "x AND y = 0")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   455
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   456
   apply (rule word_bl.Rep_eqD)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   457
   apply (simp add: bl_word_and to_bl_0)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   458
   apply (rule align_lem_and [THEN trans])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   459
       apply (simp_all add: word_size)[5]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   460
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   461
  apply (subst word_plus_and_or [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   462
  apply (simp add : bl_word_or)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   463
  apply (rule align_lem_or)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   464
     apply (simp_all add: word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   465
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   466
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   467
lemmas aligned_bl_add_size = refl [THEN aligned_bl_add_size']
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   468
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
   469
subsubsection "Mask"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   470
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   471
lemma nth_mask': "m = mask n ==> test_bit m i = (i < n & i < size m)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   472
  apply (unfold mask_def test_bit_bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   473
  apply (simp only: word_1_bl [symmetric] shiftl_of_bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   474
  apply (clarsimp simp add: word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   475
  apply (simp only: of_bl_no mask_lem number_of_succ add_diff_cancel2)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   476
  apply (fold of_bl_no)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   477
  apply (simp add: word_1_bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   478
  apply (rule test_bit_of_bl [THEN trans, unfolded test_bit_bl word_size])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   479
  apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   480
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   481
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   482
lemmas nth_mask [simp] = refl [THEN nth_mask']
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   483
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   484
lemma mask_bl: "mask n = of_bl (replicate n True)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   485
  by (auto simp add : test_bit_of_bl word_size intro: word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   486
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25350
diff changeset
   487
lemma mask_bin: "mask n = number_of (bintrunc n Int.Min)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   488
  by (auto simp add: nth_bintr word_size intro: word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   489
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   490
lemma and_mask_bintr: "w AND mask n = number_of (bintrunc n (uint w))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   491
  apply (rule word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   492
  apply (simp add: nth_bintr word_size word_ops_nth_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   493
  apply (auto simp add: test_bit_bin)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   494
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   495
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   496
lemma and_mask_no: "number_of i AND mask n = number_of (bintrunc n i)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   497
  by (auto simp add : nth_bintr word_size word_ops_nth_size intro: word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   498
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   499
lemmas and_mask_wi = and_mask_no [unfolded word_number_of_def] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   500
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   501
lemma bl_and_mask:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   502
  "to_bl (w AND mask n :: 'a :: len word) = 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   503
    replicate (len_of TYPE('a) - n) False @ 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   504
    drop (len_of TYPE('a) - n) (to_bl w)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   505
  apply (rule nth_equalityI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   506
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   507
  apply (clarsimp simp add: to_bl_nth word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   508
  apply (simp add: word_size word_ops_nth_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   509
  apply (auto simp add: word_size test_bit_bl nth_append nth_rev)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   510
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   511
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   512
lemmas and_mask_mod_2p = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   513
  and_mask_bintr [unfolded word_number_of_alt no_bintr_alt]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   514
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   515
lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   516
  apply (simp add : and_mask_bintr no_bintr_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   517
  apply (rule xtr8)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   518
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   519
   apply (rule pos_mod_bound)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   520
  apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   521
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   522
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   523
lemmas eq_mod_iff = trans [symmetric, OF int_mod_lem eq_sym_conv]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   524
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   525
lemma mask_eq_iff: "(w AND mask n) = w <-> uint w < 2 ^ n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   526
  apply (simp add: and_mask_bintr word_number_of_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   527
  apply (simp add: word_ubin.inverse_norm)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   528
  apply (simp add: eq_mod_iff bintrunc_mod2p min_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   529
  apply (fast intro!: lt2p_lem)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   530
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   531
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   532
lemma and_mask_dvd: "2 ^ n dvd uint w = (w AND mask n = 0)"
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29631
diff changeset
   533
  apply (simp add: dvd_eq_mod_eq_0 and_mask_mod_2p)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   534
  apply (simp add: word_uint.norm_eq_iff [symmetric] word_of_int_homs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   535
  apply (subst word_uint.norm_Rep [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   536
  apply (simp only: bintrunc_bintrunc_min bintrunc_mod2p [symmetric] min_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   537
  apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   538
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   539
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   540
lemma and_mask_dvd_nat: "2 ^ n dvd unat w = (w AND mask n = 0)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   541
  apply (unfold unat_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   542
  apply (rule trans [OF _ and_mask_dvd])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   543
  apply (unfold dvd_def) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   544
  apply auto 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   545
  apply (drule uint_ge_0 [THEN nat_int.Abs_inverse' [simplified], symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   546
  apply (simp add : int_mult int_power)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   547
  apply (simp add : nat_mult_distrib nat_power_eq)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   548
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   549
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   550
lemma word_2p_lem: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   551
  "n < size w ==> w < 2 ^ n = (uint (w :: 'a :: len word) < 2 ^ n)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   552
  apply (unfold word_size word_less_alt word_number_of_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   553
  apply (clarsimp simp add: word_of_int_power_hom word_uint.eq_norm 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   554
                            int_mod_eq'
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   555
                  simp del: word_of_int_bin)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   556
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   557
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   558
lemma less_mask_eq: "x < 2 ^ n ==> x AND mask n = (x :: 'a :: len word)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   559
  apply (unfold word_less_alt word_number_of_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   560
  apply (clarsimp simp add: and_mask_mod_2p word_of_int_power_hom 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   561
                            word_uint.eq_norm
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   562
                  simp del: word_of_int_bin)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   563
  apply (drule xtr8 [rotated])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   564
  apply (rule int_mod_le)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   565
  apply (auto simp add : mod_pos_pos_trivial)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   566
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   567
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   568
lemmas mask_eq_iff_w2p =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   569
  trans [OF mask_eq_iff word_2p_lem [symmetric], standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   570
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   571
lemmas and_mask_less' = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   572
  iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   573
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   574
lemma and_mask_less_size: "n < size x ==> x AND mask n < 2^n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   575
  unfolding word_size by (erule and_mask_less')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   576
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   577
lemma word_mod_2p_is_mask':
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   578
  "c = 2 ^ n ==> c > 0 ==> x mod c = (x :: 'a :: len word) AND mask n" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   579
  by (clarsimp simp add: word_mod_def uint_2p and_mask_mod_2p) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   580
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   581
lemmas word_mod_2p_is_mask = refl [THEN word_mod_2p_is_mask'] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   582
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   583
lemma mask_eqs:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   584
  "(a AND mask n) + b AND mask n = a + b AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   585
  "a + (b AND mask n) AND mask n = a + b AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   586
  "(a AND mask n) - b AND mask n = a - b AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   587
  "a - (b AND mask n) AND mask n = a - b AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   588
  "a * (b AND mask n) AND mask n = a * b AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   589
  "(b AND mask n) * a AND mask n = b * a AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   590
  "(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   591
  "(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   592
  "(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   593
  "- (a AND mask n) AND mask n = - a AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   594
  "word_succ (a AND mask n) AND mask n = word_succ a AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   595
  "word_pred (a AND mask n) AND mask n = word_pred a AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   596
  using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   597
  by (auto simp: and_mask_wi bintr_ariths bintr_arith1s new_word_of_int_homs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   598
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   599
lemma mask_power_eq:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   600
  "(x AND mask n) ^ k AND mask n = x ^ k AND mask n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   601
  using word_of_int_Ex [where x=x]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   602
  by (clarsimp simp: and_mask_wi word_of_int_power_hom bintr_ariths)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   603
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   604
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
   605
subsubsection "Revcast"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   606
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   607
lemmas revcast_def' = revcast_def [simplified]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   608
lemmas revcast_def'' = revcast_def' [simplified word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   609
lemmas revcast_no_def [simp] =
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   610
  revcast_def' [where w="number_of w", unfolded word_size, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   611
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   612
lemma to_bl_revcast: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   613
  "to_bl (revcast w :: 'a :: len0 word) = 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   614
   takefill False (len_of TYPE ('a)) (to_bl w)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   615
  apply (unfold revcast_def' word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   616
  apply (rule word_bl.Abs_inverse)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   617
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   618
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   619
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   620
lemma revcast_rev_ucast': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   621
  "cs = [rc, uc] ==> rc = revcast (word_reverse w) ==> uc = ucast w ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   622
    rc = word_reverse uc"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   623
  apply (unfold ucast_def revcast_def' Let_def word_reverse_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   624
  apply (clarsimp simp add : to_bl_of_bin takefill_bintrunc)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   625
  apply (simp add : word_bl.Abs_inverse word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   626
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   627
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   628
lemmas revcast_rev_ucast = revcast_rev_ucast' [OF refl refl refl]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   629
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   630
lemmas revcast_ucast = revcast_rev_ucast
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   631
  [where w = "word_reverse w", simplified word_rev_rev, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   632
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   633
lemmas ucast_revcast = revcast_rev_ucast [THEN word_rev_gal', standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   634
lemmas ucast_rev_revcast = revcast_ucast [THEN word_rev_gal', standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   635
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   636
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   637
-- "linking revcast and cast via shift"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   638
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   639
lemmas wsst_TYs = source_size target_size word_size
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   640
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   641
lemma revcast_down_uu': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   642
  "rc = revcast ==> source_size rc = target_size rc + n ==> 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   643
    rc (w :: 'a :: len word) = ucast (w >> n)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   644
  apply (simp add: revcast_def')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   645
  apply (rule word_bl.Rep_inverse')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   646
  apply (rule trans, rule ucast_down_drop)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   647
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   648
   apply (rule trans, rule drop_shiftr)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   649
   apply (auto simp: takefill_alt wsst_TYs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   650
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   651
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   652
lemma revcast_down_us': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   653
  "rc = revcast ==> source_size rc = target_size rc + n ==> 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   654
    rc (w :: 'a :: len word) = ucast (w >>> n)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   655
  apply (simp add: revcast_def')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   656
  apply (rule word_bl.Rep_inverse')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   657
  apply (rule trans, rule ucast_down_drop)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   658
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   659
   apply (rule trans, rule drop_sshiftr)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   660
   apply (auto simp: takefill_alt wsst_TYs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   661
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   662
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   663
lemma revcast_down_su': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   664
  "rc = revcast ==> source_size rc = target_size rc + n ==> 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   665
    rc (w :: 'a :: len word) = scast (w >> n)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   666
  apply (simp add: revcast_def')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   667
  apply (rule word_bl.Rep_inverse')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   668
  apply (rule trans, rule scast_down_drop)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   669
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   670
   apply (rule trans, rule drop_shiftr)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   671
   apply (auto simp: takefill_alt wsst_TYs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   672
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   673
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   674
lemma revcast_down_ss': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   675
  "rc = revcast ==> source_size rc = target_size rc + n ==> 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   676
    rc (w :: 'a :: len word) = scast (w >>> n)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   677
  apply (simp add: revcast_def')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   678
  apply (rule word_bl.Rep_inverse')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   679
  apply (rule trans, rule scast_down_drop)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   680
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   681
   apply (rule trans, rule drop_sshiftr)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   682
   apply (auto simp: takefill_alt wsst_TYs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   683
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   684
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   685
lemmas revcast_down_uu = refl [THEN revcast_down_uu']
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   686
lemmas revcast_down_us = refl [THEN revcast_down_us']
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   687
lemmas revcast_down_su = refl [THEN revcast_down_su']
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   688
lemmas revcast_down_ss = refl [THEN revcast_down_ss']
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   689
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   690
lemma cast_down_rev: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   691
  "uc = ucast ==> source_size uc = target_size uc + n ==> 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   692
    uc w = revcast ((w :: 'a :: len word) << n)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   693
  apply (unfold shiftl_rev)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   694
  apply clarify
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   695
  apply (simp add: revcast_rev_ucast)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   696
  apply (rule word_rev_gal')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   697
  apply (rule trans [OF _ revcast_rev_ucast])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   698
  apply (rule revcast_down_uu [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   699
  apply (auto simp add: wsst_TYs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   700
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   701
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   702
lemma revcast_up': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   703
  "rc = revcast ==> source_size rc + n = target_size rc ==> 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   704
    rc w = (ucast w :: 'a :: len word) << n" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   705
  apply (simp add: revcast_def')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   706
  apply (rule word_bl.Rep_inverse')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   707
  apply (simp add: takefill_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   708
  apply (rule bl_shiftl [THEN trans])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   709
  apply (subst ucast_up_app)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   710
  apply (auto simp add: wsst_TYs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   711
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   712
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   713
lemmas revcast_up = refl [THEN revcast_up']
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   714
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   715
lemmas rc1 = revcast_up [THEN 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   716
  revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   717
lemmas rc2 = revcast_down_uu [THEN 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   718
  revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   719
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   720
lemmas ucast_up =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   721
 rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   722
lemmas ucast_down = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   723
  rc2 [simplified rev_shiftr revcast_ucast [symmetric]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   724
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   725
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
   726
subsubsection "Slices"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   727
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   728
lemmas slice1_no_bin [simp] =
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
   729
  slice1_def [where w="number_of w", unfolded to_bl_no_bin, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   730
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   731
lemmas slice_no_bin [simp] = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   732
   trans [OF slice_def [THEN meta_eq_to_obj_eq] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   733
             slice1_no_bin [THEN meta_eq_to_obj_eq], 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   734
          unfolded word_size, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   735
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   736
lemma slice1_0 [simp] : "slice1 n 0 = 0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   737
  unfolding slice1_def by (simp add : to_bl_0)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   738
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   739
lemma slice_0 [simp] : "slice n 0 = 0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   740
  unfolding slice_def by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   741
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   742
lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   743
  unfolding slice_def' slice1_def
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   744
  by (simp add : takefill_alt word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   745
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   746
lemmas slice_take = slice_take' [unfolded word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   747
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   748
-- "shiftr to a word of the same size is just slice, 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   749
    slice is just shiftr then ucast"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   750
lemmas shiftr_slice = trans
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   751
  [OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric], standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   752
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   753
lemma slice_shiftr: "slice n w = ucast (w >> n)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   754
  apply (unfold slice_take shiftr_bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   755
  apply (rule ucast_of_bl_up [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   756
  apply (simp add: word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   757
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   758
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   759
lemma nth_slice: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   760
  "(slice n w :: 'a :: len0 word) !! m = 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   761
   (w !! (m + n) & m < len_of TYPE ('a))"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   762
  unfolding slice_shiftr 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   763
  by (simp add : nth_ucast nth_shiftr)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   764
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   765
lemma slice1_down_alt': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   766
  "sl = slice1 n w ==> fs = size sl ==> fs + k = n ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   767
    to_bl sl = takefill False fs (drop k (to_bl w))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   768
  unfolding slice1_def word_size of_bl_def uint_bl
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   769
  by (clarsimp simp: word_ubin.eq_norm bl_bin_bl_rep_drop drop_takefill)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   770
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   771
lemma slice1_up_alt': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   772
  "sl = slice1 n w ==> fs = size sl ==> fs = n + k ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   773
    to_bl sl = takefill False fs (replicate k False @ (to_bl w))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   774
  apply (unfold slice1_def word_size of_bl_def uint_bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   775
  apply (clarsimp simp: word_ubin.eq_norm bl_bin_bl_rep_drop 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   776
                        takefill_append [symmetric])
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   777
  apply (rule_tac f = "%k. takefill False (len_of TYPE('a))
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   778
    (replicate k False @ bin_to_bl (len_of TYPE('b)) (uint w))" in arg_cong)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   779
  apply arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   780
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   781
    
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   782
lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   783
lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   784
lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   785
lemmas slice1_up_alts = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   786
  le_add_diff_inverse [symmetric, THEN su1] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   787
  le_add_diff_inverse2 [symmetric, THEN su1]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   788
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   789
lemma ucast_slice1: "ucast w = slice1 (size w) w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   790
  unfolding slice1_def ucast_bl
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   791
  by (simp add : takefill_same' word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   792
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   793
lemma ucast_slice: "ucast w = slice 0 w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   794
  unfolding slice_def by (simp add : ucast_slice1)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   795
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   796
lemmas slice_id = trans [OF ucast_slice [symmetric] ucast_id]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   797
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   798
lemma revcast_slice1': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   799
  "rc = revcast w ==> slice1 (size rc) w = rc"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   800
  unfolding slice1_def revcast_def' by (simp add : word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   801
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   802
lemmas revcast_slice1 = refl [THEN revcast_slice1']
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   803
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   804
lemma slice1_tf_tf': 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   805
  "to_bl (slice1 n w :: 'a :: len0 word) = 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   806
   rev (takefill False (len_of TYPE('a)) (rev (takefill False n (to_bl w))))"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   807
  unfolding slice1_def by (rule word_rev_tf)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   808
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   809
lemmas slice1_tf_tf = slice1_tf_tf'
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   810
  [THEN word_bl.Rep_inverse', symmetric, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   811
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   812
lemma rev_slice1:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   813
  "n + k = len_of TYPE('a) + len_of TYPE('b) \<Longrightarrow> 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   814
  slice1 n (word_reverse w :: 'b :: len0 word) = 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   815
  word_reverse (slice1 k w :: 'a :: len0 word)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   816
  apply (unfold word_reverse_def slice1_tf_tf)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   817
  apply (rule word_bl.Rep_inverse')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   818
  apply (rule rev_swap [THEN iffD1])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   819
  apply (rule trans [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   820
  apply (rule tf_rev)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   821
   apply (simp add: word_bl.Abs_inverse)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   822
  apply (simp add: word_bl.Abs_inverse)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   823
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   824
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   825
lemma rev_slice': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   826
  "res = slice n (word_reverse w) ==> n + k + size res = size w ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   827
    res = word_reverse (slice k w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   828
  apply (unfold slice_def word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   829
  apply clarify
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   830
  apply (rule rev_slice1)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   831
  apply arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   832
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   833
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   834
lemmas rev_slice = refl [THEN rev_slice', unfolded word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   835
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   836
lemmas sym_notr = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   837
  not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   838
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   839
-- {* problem posed by TPHOLs referee:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   840
      criterion for overflow of addition of signed integers *}
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   841
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   842
lemma sofl_test:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   843
  "(sint (x :: 'a :: len word) + sint y = sint (x + y)) = 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   844
     ((((x+y) XOR x) AND ((x+y) XOR y)) >> (size x - 1) = 0)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   845
  apply (unfold word_size)
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   846
  apply (cases "len_of TYPE('a)", simp) 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   847
  apply (subst msb_shift [THEN sym_notr])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   848
  apply (simp add: word_ops_msb)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   849
  apply (simp add: word_msb_sint)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   850
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   851
       apply simp_all
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   852
     apply (unfold sint_word_ariths)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   853
     apply (unfold word_sbin.set_iff_norm [symmetric] sints_num)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   854
     apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   855
        apply (insert sint_range' [where x=x])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   856
        apply (insert sint_range' [where x=y])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   857
        defer 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   858
        apply (simp (no_asm), arith)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   859
       apply (simp (no_asm), arith)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   860
      defer
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   861
      defer
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   862
      apply (simp (no_asm), arith)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   863
     apply (simp (no_asm), arith)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   864
    apply (rule notI [THEN notnotD],
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   865
           drule leI not_leE,
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   866
           drule sbintrunc_inc sbintrunc_dec,      
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   867
           simp)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   868
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   869
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   870
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
   871
subsection "Split and cat"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   872
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   873
lemmas word_split_bin' = word_split_def [THEN meta_eq_to_obj_eq, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   874
lemmas word_cat_bin' = word_cat_def [THEN meta_eq_to_obj_eq, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   875
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   876
lemma word_rsplit_no:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   877
  "(word_rsplit (number_of bin :: 'b :: len0 word) :: 'a word list) = 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   878
    map number_of (bin_rsplit (len_of TYPE('a :: len)) 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   879
      (len_of TYPE('b), bintrunc (len_of TYPE('b)) bin))"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   880
  apply (unfold word_rsplit_def word_no_wi)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   881
  apply (simp add: word_ubin.eq_norm)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   882
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   883
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   884
lemmas word_rsplit_no_cl [simp] = word_rsplit_no
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   885
  [unfolded bin_rsplitl_def bin_rsplit_l [symmetric]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   886
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   887
lemma test_bit_cat:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   888
  "wc = word_cat a b ==> wc !! n = (n < size wc & 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   889
    (if n < size b then b !! n else a !! (n - size b)))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   890
  apply (unfold word_cat_bin' test_bit_bin)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   891
  apply (auto simp add : word_ubin.eq_norm nth_bintr bin_nth_cat word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   892
  apply (erule bin_nth_uint_imp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   893
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   894
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   895
lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   896
  apply (unfold of_bl_def to_bl_def word_cat_bin')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   897
  apply (simp add: bl_to_bin_app_cat)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   898
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   899
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   900
lemma of_bl_append:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   901
  "(of_bl (xs @ ys) :: 'a :: len word) = of_bl xs * 2^(length ys) + of_bl ys"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   902
  apply (unfold of_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   903
  apply (simp add: bl_to_bin_app_cat bin_cat_num)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   904
  apply (simp add: word_of_int_power_hom [symmetric] new_word_of_int_hom_syms)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   905
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   906
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   907
lemma of_bl_False [simp]:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   908
  "of_bl (False#xs) = of_bl xs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   909
  by (rule word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   910
     (auto simp add: test_bit_of_bl nth_append)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   911
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   912
lemma of_bl_True: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   913
  "(of_bl (True#xs)::'a::len word) = 2^length xs + of_bl xs"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   914
  by (subst of_bl_append [where xs="[True]", simplified])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   915
     (simp add: word_1_bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   916
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   917
lemma of_bl_Cons:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   918
  "of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   919
  by (cases x) (simp_all add: of_bl_True)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   920
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   921
lemma split_uint_lem: "bin_split n (uint (w :: 'a :: len0 word)) = (a, b) ==> 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   922
  a = bintrunc (len_of TYPE('a) - n) a & b = bintrunc (len_of TYPE('a)) b"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   923
  apply (frule word_ubin.norm_Rep [THEN ssubst])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   924
  apply (drule bin_split_trunc1)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   925
  apply (drule sym [THEN trans])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   926
  apply assumption
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   927
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   928
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   929
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   930
lemma word_split_bl': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   931
  "std = size c - size b ==> (word_split c = (a, b)) ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   932
    (a = of_bl (take std (to_bl c)) & b = of_bl (drop std (to_bl c)))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   933
  apply (unfold word_split_bin')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   934
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   935
   defer
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   936
   apply (clarsimp split: prod.splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   937
   apply (drule word_ubin.norm_Rep [THEN ssubst])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   938
   apply (drule split_bintrunc)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   939
   apply (simp add : of_bl_def bl2bin_drop word_size
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   940
       word_ubin.norm_eq_iff [symmetric] min_def del : word_ubin.norm_Rep)    
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   941
  apply (clarsimp split: prod.splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   942
  apply (frule split_uint_lem [THEN conjunct1])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   943
  apply (unfold word_size)
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   944
  apply (cases "len_of TYPE('a) >= len_of TYPE('b)")
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   945
   defer
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   946
   apply (simp add: word_0_bl word_0_wi_Pls)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   947
  apply (simp add : of_bl_def to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   948
  apply (subst bin_split_take1 [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   949
    prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   950
    apply assumption
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   951
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   952
  apply (erule thin_rl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   953
  apply (erule arg_cong [THEN trans])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   954
  apply (simp add : word_ubin.norm_eq_iff [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   955
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   956
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   957
lemma word_split_bl: "std = size c - size b ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   958
    (a = of_bl (take std (to_bl c)) & b = of_bl (drop std (to_bl c))) <-> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   959
    word_split c = (a, b)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   960
  apply (rule iffI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   961
   defer
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   962
   apply (erule (1) word_split_bl')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   963
  apply (case_tac "word_split c")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   964
  apply (auto simp add : word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   965
  apply (frule word_split_bl' [rotated])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   966
  apply (auto simp add : word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   967
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   968
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   969
lemma word_split_bl_eq:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   970
   "(word_split (c::'a::len word) :: ('c :: len0 word * 'd :: len0 word)) =
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   971
      (of_bl (take (len_of TYPE('a::len) - len_of TYPE('d::len0)) (to_bl c)),
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
   972
       of_bl (drop (len_of TYPE('a) - len_of TYPE('d)) (to_bl c)))"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   973
  apply (rule word_split_bl [THEN iffD1])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   974
  apply (unfold word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   975
  apply (rule refl conjI)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   976
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   977
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   978
-- "keep quantifiers for use in simplification"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   979
lemma test_bit_split':
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   980
  "word_split c = (a, b) --> (ALL n m. b !! n = (n < size b & c !! n) & 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   981
    a !! m = (m < size a & c !! (m + size b)))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   982
  apply (unfold word_split_bin' test_bit_bin)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   983
  apply (clarify)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   984
  apply (clarsimp simp: word_ubin.eq_norm nth_bintr word_size split: prod.splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   985
  apply (drule bin_nth_split)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   986
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   987
       apply (simp_all add: add_commute)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   988
   apply (erule bin_nth_uint_imp)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   989
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   990
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   991
lemmas test_bit_split = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   992
  test_bit_split' [THEN mp, simplified all_simps, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   993
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   994
lemma test_bit_split_eq: "word_split c = (a, b) <-> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   995
  ((ALL n::nat. b !! n = (n < size b & c !! n)) &
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   996
    (ALL m::nat. a !! m = (m < size a & c !! (m + size b))))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   997
  apply (rule_tac iffI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   998
   apply (rule_tac conjI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   999
    apply (erule test_bit_split [THEN conjunct1])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1000
   apply (erule test_bit_split [THEN conjunct2])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1001
  apply (case_tac "word_split c")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1002
  apply (frule test_bit_split)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1003
  apply (erule trans)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1004
  apply (fastsimp intro ! : word_eqI simp add : word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1005
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1006
26008
24c82bef5696 eliminated escaped white space;
wenzelm
parents: 25919
diff changeset
  1007
-- {* this odd result is analogous to @{text "ucast_id"}, 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1008
      result to the length given by the result type *}
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1009
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1010
lemma word_cat_id: "word_cat a b = b"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1011
  unfolding word_cat_bin' by (simp add: word_ubin.inverse_norm)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1012
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1013
-- "limited hom result"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1014
lemma word_cat_hom:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1015
  "len_of TYPE('a::len0) <= len_of TYPE('b::len0) + len_of TYPE ('c::len0)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1016
  ==>
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1017
  (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1018
  word_of_int (bin_cat w (size b) (uint b))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1019
  apply (unfold word_cat_def word_size) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1020
  apply (clarsimp simp add : word_ubin.norm_eq_iff [symmetric]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1021
      word_ubin.eq_norm bintr_cat min_def)
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1022
  apply arith
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1023
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1024
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1025
lemma word_cat_split_alt:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1026
  "size w <= size u + size v ==> word_split w = (u, v) ==> word_cat u v = w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1027
  apply (rule word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1028
  apply (drule test_bit_split)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1029
  apply (clarsimp simp add : test_bit_cat word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1030
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1031
  apply arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1032
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1033
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1034
lemmas word_cat_split_size = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1035
  sym [THEN [2] word_cat_split_alt [symmetric], standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1036
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1037
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
  1038
subsubsection "Split and slice"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1039
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1040
lemma split_slices: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1041
  "word_split w = (u, v) ==> u = slice (size v) w & v = slice 0 w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1042
  apply (drule test_bit_split)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1043
  apply (rule conjI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1044
   apply (rule word_eqI, clarsimp simp: nth_slice word_size)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1045
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1046
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1047
lemma slice_cat1':
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1048
  "wc = word_cat a b ==> size wc >= size a + size b ==> slice (size b) wc = a"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1049
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1050
  apply (rule word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1051
  apply (simp add: nth_slice test_bit_cat word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1052
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1053
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1054
lemmas slice_cat1 = refl [THEN slice_cat1']
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1055
lemmas slice_cat2 = trans [OF slice_id word_cat_id]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1056
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1057
lemma cat_slices:
26008
24c82bef5696 eliminated escaped white space;
wenzelm
parents: 25919
diff changeset
  1058
  "a = slice n c ==> b = slice 0 c ==> n = size b ==>
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1059
    size a + size b >= size c ==> word_cat a b = c"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1060
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1061
  apply (rule word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1062
  apply (simp add: nth_slice test_bit_cat word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1063
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1064
  apply arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1065
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1066
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1067
lemma word_split_cat_alt:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1068
  "w = word_cat u v ==> size u + size v <= size w ==> word_split w = (u, v)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1069
  apply (case_tac "word_split ?w")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1070
  apply (rule trans, assumption)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1071
  apply (drule test_bit_split)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1072
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1073
   apply (rule word_eqI, clarsimp simp: test_bit_cat word_size)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1074
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1075
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1076
lemmas word_cat_bl_no_bin [simp] =
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
  1077
  word_cat_bl [where a="number_of a" 
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
  1078
                 and b="number_of b", 
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
  1079
               unfolded to_bl_no_bin, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1080
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1081
lemmas word_split_bl_no_bin [simp] =
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
  1082
  word_split_bl_eq [where c="number_of c", unfolded to_bl_no_bin, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1083
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1084
-- {* this odd result arises from the fact that the statement of the
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1085
      result implies that the decoded words are of the same type, 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1086
      and therefore of the same length, as the original word *}
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1087
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1088
lemma word_rsplit_same: "word_rsplit w = [w]"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1089
  unfolding word_rsplit_def by (simp add : bin_rsplit_all)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1090
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1091
lemma word_rsplit_empty_iff_size:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1092
  "(word_rsplit w = []) = (size w = 0)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1093
  unfolding word_rsplit_def bin_rsplit_def word_size
26289
9d2c375e242b avoid unclear fact references;
wenzelm
parents: 26086
diff changeset
  1094
  by (simp add: bin_rsplit_aux_simp_alt Let_def split: Product_Type.split_split)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1095
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1096
lemma test_bit_rsplit:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1097
  "sw = word_rsplit w ==> m < size (hd sw :: 'a :: len word) ==> 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1098
    k < length sw ==> (rev sw ! k) !! m = (w !! (k * size (hd sw) + m))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1099
  apply (unfold word_rsplit_def word_test_bit_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1100
  apply (rule trans)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1101
   apply (rule_tac f = "%x. bin_nth x m" in arg_cong)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1102
   apply (rule nth_map [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1103
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1104
  apply (rule bin_nth_rsplit)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1105
     apply simp_all
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1106
  apply (simp add : word_size rev_map map_compose [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1107
  apply (rule trans)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1108
   defer
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1109
   apply (rule map_ident [THEN fun_cong])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1110
  apply (rule refl [THEN map_cong])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1111
  apply (simp add : word_ubin.eq_norm)
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1112
  apply (erule bin_rsplit_size_sign [OF len_gt_0 refl])
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1113
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1114
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1115
lemma word_rcat_bl: "word_rcat wl == of_bl (concat (map to_bl wl))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1116
  unfolding word_rcat_def to_bl_def' of_bl_def
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1117
  by (clarsimp simp add : bin_rcat_bl map_compose)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1118
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1119
lemma size_rcat_lem':
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1120
  "size (concat (map to_bl wl)) = length wl * size (hd wl)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1121
  unfolding word_size by (induct wl) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1122
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1123
lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1124
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1125
lemmas td_gal_lt_len = len_gt_0 [THEN td_gal_lt, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1126
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1127
lemma nth_rcat_lem' [rule_format] :
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1128
  "sw = size (hd wl  :: 'a :: len word) ==> (ALL n. n < size wl * sw --> 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1129
    rev (concat (map to_bl wl)) ! n = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1130
    rev (to_bl (rev wl ! (n div sw))) ! (n mod sw))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1131
  apply (unfold word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1132
  apply (induct "wl")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1133
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1134
  apply (clarsimp simp add : nth_append size_rcat_lem)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1135
  apply (simp (no_asm_use) only:  mult_Suc [symmetric] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1136
         td_gal_lt_len less_Suc_eq_le mod_div_equality')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1137
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1138
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1139
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1140
lemmas nth_rcat_lem = refl [THEN nth_rcat_lem', unfolded word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1141
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1142
lemma test_bit_rcat:
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1143
  "sw = size (hd wl :: 'a :: len word) ==> rc = word_rcat wl ==> rc !! n = 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1144
    (n < size rc & n div sw < size wl & (rev wl) ! (n div sw) !! (n mod sw))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1145
  apply (unfold word_rcat_bl word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1146
  apply (clarsimp simp add : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1147
    test_bit_of_bl size_rcat_lem word_size td_gal_lt_len)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1148
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1149
   apply (auto simp add : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1150
    test_bit_bl word_size td_gal_lt_len [THEN iffD2, THEN nth_rcat_lem])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1151
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1152
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1153
lemma foldl_eq_foldr [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1154
  "ALL x. foldl op + x xs = foldr op + (x # xs) (0 :: 'a :: comm_monoid_add)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1155
  by (induct xs) (auto simp add : add_assoc)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1156
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1157
lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1158
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1159
lemmas test_bit_rsplit_alt = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1160
  trans [OF nth_rev_alt [THEN test_bit_cong] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1161
  test_bit_rsplit [OF refl asm_rl diff_Suc_less]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1162
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1163
-- "lazy way of expressing that u and v, and su and sv, have same types"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1164
lemma word_rsplit_len_indep':
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1165
  "[u,v] = p ==> [su,sv] = q ==> word_rsplit u = su ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1166
    word_rsplit v = sv ==> length su = length sv"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1167
  apply (unfold word_rsplit_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1168
  apply (auto simp add : bin_rsplit_len_indep)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1169
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1170
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1171
lemmas word_rsplit_len_indep = word_rsplit_len_indep' [OF refl refl refl refl]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1172
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1173
lemma length_word_rsplit_size: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1174
  "n = len_of TYPE ('a :: len) ==> 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1175
    (length (word_rsplit w :: 'a word list) <= m) = (size w <= m * n)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1176
  apply (unfold word_rsplit_def word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1177
  apply (clarsimp simp add : bin_rsplit_len_le)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1178
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1179
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1180
lemmas length_word_rsplit_lt_size = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1181
  length_word_rsplit_size [unfolded Not_eq_iff linorder_not_less [symmetric]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1182
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1183
lemma length_word_rsplit_exp_size: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1184
  "n = len_of TYPE ('a :: len) ==> 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1185
    length (word_rsplit w :: 'a word list) = (size w + n - 1) div n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1186
  unfolding word_rsplit_def by (clarsimp simp add : word_size bin_rsplit_len)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1187
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1188
lemma length_word_rsplit_even_size: 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1189
  "n = len_of TYPE ('a :: len) ==> size w = m * n ==> 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1190
    length (word_rsplit w :: 'a word list) = m"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1191
  by (clarsimp simp add : length_word_rsplit_exp_size given_quot_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1192
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1193
lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1194
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1195
(* alternative proof of word_rcat_rsplit *)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1196
lemmas tdle = iffD2 [OF split_div_lemma refl, THEN conjunct1] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1197
lemmas dtle = xtr4 [OF tdle mult_commute]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1198
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1199
lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1200
  apply (rule word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1201
  apply (clarsimp simp add : test_bit_rcat word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1202
  apply (subst refl [THEN test_bit_rsplit])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1203
    apply (simp_all add: word_size 
26072
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 26008
diff changeset
  1204
      refl [THEN length_word_rsplit_size [simplified not_less [symmetric], simplified]])
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1205
   apply safe
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1206
   apply (erule xtr7, rule len_gt_0 [THEN dtle])+
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1207
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1208
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1209
lemma size_word_rsplit_rcat_size':
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1210
  "word_rcat (ws :: 'a :: len word list) = frcw ==> 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1211
    size frcw = length ws * len_of TYPE ('a) ==> 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1212
    size (hd [word_rsplit frcw, ws]) = size ws" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1213
  apply (clarsimp simp add : word_size length_word_rsplit_exp_size')
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1214
  apply (fast intro: given_quot_alt)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1215
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1216
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1217
lemmas size_word_rsplit_rcat_size = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1218
  size_word_rsplit_rcat_size' [simplified]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1219
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1220
lemma msrevs:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1221
  fixes n::nat
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1222
  shows "0 < n \<Longrightarrow> (k * n + m) div n = m div n + k"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1223
  and   "(k * n + m) mod n = m mod n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1224
  by (auto simp: add_commute)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1225
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1226
lemma word_rsplit_rcat_size':
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1227
  "word_rcat (ws :: 'a :: len word list) = frcw ==> 
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1228
    size frcw = length ws * len_of TYPE ('a) ==> word_rsplit frcw = ws" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1229
  apply (frule size_word_rsplit_rcat_size, assumption)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1230
  apply (clarsimp simp add : word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1231
  apply (rule nth_equalityI, assumption)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1232
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1233
  apply (rule word_eqI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1234
  apply (rule trans)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1235
   apply (rule test_bit_rsplit_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1236
     apply (clarsimp simp: word_size)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1237
  apply (rule trans)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1238
  apply (rule test_bit_rcat [OF refl refl])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1239
  apply (simp add : word_size msrevs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1240
  apply (subst nth_rev)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1241
   apply arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1242
  apply (simp add : le0 [THEN [2] xtr7, THEN diff_Suc_less])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1243
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1244
  apply (simp add : diff_mult_distrib)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1245
  apply (rule mpl_lem)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1246
  apply (cases "size ws")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1247
   apply simp_all
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1248
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1249
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1250
lemmas word_rsplit_rcat_size = refl [THEN word_rsplit_rcat_size']
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1251
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1252
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
  1253
subsection "Rotation"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1254
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1255
lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1256
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1257
lemmas word_rot_defs = word_roti_def word_rotr_def word_rotl_def
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1258
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1259
lemma rotate_eq_mod: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1260
  "m mod length xs = n mod length xs ==> rotate m xs = rotate n xs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1261
  apply (rule box_equals)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1262
    defer
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1263
    apply (rule rotate_conv_mod [symmetric])+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1264
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1265
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1266
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1267
lemmas rotate_eqs [standard] = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1268
  trans [OF rotate0 [THEN fun_cong] id_apply]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1269
  rotate_rotate [symmetric] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1270
  rotate_id 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1271
  rotate_conv_mod 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1272
  rotate_eq_mod
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1273
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1274
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
  1275
subsubsection "Rotation of list to right"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1276
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1277
lemma rotate1_rl': "rotater1 (l @ [a]) = a # l"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1278
  unfolding rotater1_def by (cases l) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1279
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1280
lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1281
  apply (unfold rotater1_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1282
  apply (cases "l")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1283
  apply (case_tac [2] "list")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1284
  apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1285
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1286
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1287
lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1288
  unfolding rotater1_def by (cases l) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1289
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1290
lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1291
  apply (cases "xs")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1292
  apply (simp add : rotater1_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1293
  apply (simp add : rotate1_rl')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1294
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1295
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1296
lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1297
  unfolding rotater_def by (induct n) (auto intro: rotater1_rev')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1298
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
  1299
lemmas rotater_rev = rotater_rev' [where xs = "rev ys", simplified, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1300
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1301
lemma rotater_drop_take: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1302
  "rotater n xs = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1303
   drop (length xs - n mod length xs) xs @
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1304
   take (length xs - n mod length xs) xs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1305
  by (clarsimp simp add : rotater_rev rotate_drop_take rev_take rev_drop)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1306
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1307
lemma rotater_Suc [simp] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1308
  "rotater (Suc n) xs = rotater1 (rotater n xs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1309
  unfolding rotater_def by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1310
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1311
lemma rotate_inv_plus [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1312
  "ALL k. k = m + n --> rotater k (rotate n xs) = rotater m xs & 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1313
    rotate k (rotater n xs) = rotate m xs & 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1314
    rotater n (rotate k xs) = rotate m xs & 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1315
    rotate n (rotater k xs) = rotater m xs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1316
  unfolding rotater_def rotate_def
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1317
  by (induct n) (auto intro: funpow_swap1 [THEN trans])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1318
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1319
lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1320
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1321
lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1322
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1323
lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1324
lemmas rotate_rl [simp] =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1325
  rotate_inv_eq [THEN conjunct2, THEN conjunct1, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1326
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1327
lemma rotate_gal: "(rotater n xs = ys) = (rotate n ys = xs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1328
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1329
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1330
lemma rotate_gal': "(ys = rotater n xs) = (xs = rotate n ys)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1331
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1332
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1333
lemma length_rotater [simp]: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1334
  "length (rotater n xs) = length xs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1335
  by (simp add : rotater_rev)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1336
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1337
lemmas rrs0 = rotate_eqs [THEN restrict_to_left, 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1338
  simplified rotate_gal [symmetric] rotate_gal' [symmetric], standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1339
lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1340
lemmas rotater_eqs = rrs1 [simplified length_rotater, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1341
lemmas rotater_0 = rotater_eqs (1)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1342
lemmas rotater_add = rotater_eqs (2)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1343
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1344
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1345
subsubsection "map, map2, commuting with rotate(r)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1346
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1347
lemma last_map: "xs ~= [] ==> last (map f xs) = f (last xs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1348
  by (induct xs) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1349
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1350
lemma butlast_map:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1351
  "xs ~= [] ==> butlast (map f xs) = map f (butlast xs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1352
  by (induct xs) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1353
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1354
lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1355
  unfolding rotater1_def
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1356
  by (cases xs) (auto simp add: last_map butlast_map)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1357
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1358
lemma rotater_map:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1359
  "rotater n (map f xs) = map f (rotater n xs)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1360
  unfolding rotater_def
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1361
  by (induct n) (auto simp add : rotater1_map)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1362
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1363
lemma but_last_zip [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1364
  "ALL ys. length xs = length ys --> xs ~= [] --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1365
    last (zip xs ys) = (last xs, last ys) & 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1366
    butlast (zip xs ys) = zip (butlast xs) (butlast ys)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1367
  apply (induct "xs")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1368
  apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1369
     apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1370
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1371
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1372
lemma but_last_map2 [rule_format] :
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1373
  "ALL ys. length xs = length ys --> xs ~= [] --> 
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1374
    last (map2 f xs ys) = f (last xs) (last ys) & 
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1375
    butlast (map2 f xs ys) = map2 f (butlast xs) (butlast ys)" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1376
  apply (induct "xs")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1377
  apply auto
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1378
     apply (unfold map2_def)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1379
     apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1380
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1381
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1382
lemma rotater1_zip:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1383
  "length xs = length ys ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1384
    rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1385
  apply (unfold rotater1_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1386
  apply (cases "xs")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1387
   apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1388
   apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1389
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1390
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1391
lemma rotater1_map2:
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1392
  "length xs = length ys ==> 
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1393
    rotater1 (map2 f xs ys) = map2 f (rotater1 xs) (rotater1 ys)" 
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1394
  unfolding map2_def by (simp add: rotater1_map rotater1_zip)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1395
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1396
lemmas lrth = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1397
  box_equals [OF asm_rl length_rotater [symmetric] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1398
                 length_rotater [symmetric], 
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1399
              THEN rotater1_map2]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1400
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1401
lemma rotater_map2: 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1402
  "length xs = length ys ==> 
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1403
    rotater n (map2 f xs ys) = map2 f (rotater n xs) (rotater n ys)" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1404
  by (induct n) (auto intro!: lrth)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1405
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1406
lemma rotate1_map2:
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1407
  "length xs = length ys ==> 
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1408
    rotate1 (map2 f xs ys) = map2 f (rotate1 xs) (rotate1 ys)" 
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1409
  apply (unfold map2_def)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1410
  apply (cases xs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1411
   apply (cases ys, auto simp add : rotate1_def)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1412
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1413
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1414
lemmas lth = box_equals [OF asm_rl length_rotate [symmetric] 
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1415
  length_rotate [symmetric], THEN rotate1_map2]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1416
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1417
lemma rotate_map2: 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1418
  "length xs = length ys ==> 
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1419
    rotate n (map2 f xs ys) = map2 f (rotate n xs) (rotate n ys)" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1420
  by (induct n) (auto intro!: lth)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1421
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1422
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1423
-- "corresponding equalities for word rotation"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1424
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1425
lemma to_bl_rotl: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1426
  "to_bl (word_rotl n w) = rotate n (to_bl w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1427
  by (simp add: word_bl.Abs_inverse' word_rotl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1428
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1429
lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1430
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1431
lemmas word_rotl_eqs =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1432
  blrs0 [simplified word_bl.Rep' word_bl.Rep_inject to_bl_rotl [symmetric]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1433
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1434
lemma to_bl_rotr: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1435
  "to_bl (word_rotr n w) = rotater n (to_bl w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1436
  by (simp add: word_bl.Abs_inverse' word_rotr_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1437
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1438
lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1439
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1440
lemmas word_rotr_eqs =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1441
  brrs0 [simplified word_bl.Rep' word_bl.Rep_inject to_bl_rotr [symmetric]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1442
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1443
declare word_rotr_eqs (1) [simp]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1444
declare word_rotl_eqs (1) [simp]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1445
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1446
lemma
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1447
  word_rot_rl [simp]:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1448
  "word_rotl k (word_rotr k v) = v" and
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1449
  word_rot_lr [simp]:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1450
  "word_rotr k (word_rotl k v) = v"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1451
  by (auto simp add: to_bl_rotr to_bl_rotl word_bl.Rep_inject [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1452
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1453
lemma
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1454
  word_rot_gal:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1455
  "(word_rotr n v = w) = (word_rotl n w = v)" and
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1456
  word_rot_gal':
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1457
  "(w = word_rotr n v) = (v = word_rotl n w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1458
  by (auto simp: to_bl_rotr to_bl_rotl word_bl.Rep_inject [symmetric] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1459
           dest: sym)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1460
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1461
lemma word_rotr_rev:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1462
  "word_rotr n w = word_reverse (word_rotl n (word_reverse w))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1463
  by (simp add: word_bl.Rep_inject [symmetric] to_bl_word_rev
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1464
                to_bl_rotr to_bl_rotl rotater_rev)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1465
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1466
lemma word_roti_0 [simp]: "word_roti 0 w = w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1467
  by (unfold word_rot_defs) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1468
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1469
lemmas abl_cong = arg_cong [where f = "of_bl"]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1470
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1471
lemma word_roti_add: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1472
  "word_roti (m + n) w = word_roti m (word_roti n w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1473
proof -
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1474
  have rotater_eq_lem: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1475
    "\<And>m n xs. m = n ==> rotater m xs = rotater n xs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1476
    by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1477
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1478
  have rotate_eq_lem: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1479
    "\<And>m n xs. m = n ==> rotate m xs = rotate n xs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1480
    by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1481
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1482
  note rpts [symmetric, standard] = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1483
    rotate_inv_plus [THEN conjunct1]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1484
    rotate_inv_plus [THEN conjunct2, THEN conjunct1]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1485
    rotate_inv_plus [THEN conjunct2, THEN conjunct2, THEN conjunct1]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1486
    rotate_inv_plus [THEN conjunct2, THEN conjunct2, THEN conjunct2]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1487
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1488
  note rrp = trans [symmetric, OF rotate_rotate rotate_eq_lem]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1489
  note rrrp = trans [symmetric, OF rotater_add [symmetric] rotater_eq_lem]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1490
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1491
  show ?thesis
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1492
  apply (unfold word_rot_defs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1493
  apply (simp only: split: split_if)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1494
  apply (safe intro!: abl_cong)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1495
  apply (simp_all only: to_bl_rotl [THEN word_bl.Rep_inverse'] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1496
                    to_bl_rotl
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1497
                    to_bl_rotr [THEN word_bl.Rep_inverse']
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1498
                    to_bl_rotr)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1499
  apply (rule rrp rrrp rpts,
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1500
         simp add: nat_add_distrib [symmetric] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1501
                   nat_diff_distrib [symmetric])+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1502
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1503
qed
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1504
    
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1505
lemma word_roti_conv_mod': "word_roti n w = word_roti (n mod int (size w)) w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1506
  apply (unfold word_rot_defs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1507
  apply (cut_tac y="size w" in gt_or_eq_0)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1508
  apply (erule disjE)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1509
   apply simp_all
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1510
  apply (safe intro!: abl_cong)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1511
   apply (rule rotater_eqs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1512
   apply (simp add: word_size nat_mod_distrib)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1513
  apply (simp add: rotater_add [symmetric] rotate_gal [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1514
  apply (rule rotater_eqs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1515
  apply (simp add: word_size nat_mod_distrib)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1516
  apply (rule int_eq_0_conv [THEN iffD1])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1517
  apply (simp only: zmod_int zadd_int [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1518
  apply (simp add: rdmods)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1519
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1520
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1521
lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1522
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1523
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
  1524
subsubsection "Word rotation commutes with bit-wise operations"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1525
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1526
(* using locale to not pollute lemma namespace *)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1527
locale word_rotate 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1528
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1529
context word_rotate
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1530
begin
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1531
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1532
lemmas word_rot_defs' = to_bl_rotl to_bl_rotr
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1533
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1534
lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1535
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1536
lemmas lbl_lbl = trans [OF word_bl.Rep' word_bl.Rep' [symmetric]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1537
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1538
lemmas ths_map2 [OF lbl_lbl] = rotate_map2 rotater_map2
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1539
26936
faf8a5b5ba87 avoid undeclared variables in facts;
wenzelm
parents: 26558
diff changeset
  1540
lemmas ths_map [where xs = "to_bl v", standard] = rotate_map rotater_map
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1541
26558
7fcc10088e72 renamed app2 to map2
haftmann
parents: 26289
diff changeset
  1542
lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_map2 ths_map
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1543
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1544
lemma word_rot_logs:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1545
  "word_rotl n (NOT v) = NOT word_rotl n v"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1546
  "word_rotr n (NOT v) = NOT word_rotr n v"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1547
  "word_rotl n (x AND y) = word_rotl n x AND word_rotl n y"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1548
  "word_rotr n (x AND y) = word_rotr n x AND word_rotr n y"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1549
  "word_rotl n (x OR y) = word_rotl n x OR word_rotl n y"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1550
  "word_rotr n (x OR y) = word_rotr n x OR word_rotr n y"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1551
  "word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1552
  "word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y"  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1553
  by (rule word_bl.Rep_eqD,
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1554
      rule word_rot_defs' [THEN trans],
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1555
      simp only: blwl_syms [symmetric],
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1556
      rule th1s [THEN trans], 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1557
      rule refl)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1558
end
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1559
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1560
lemmas word_rot_logs = word_rotate.word_rot_logs
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1561
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1562
lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take,
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1563
  simplified word_bl.Rep', standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1564
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1565
lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take,
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1566
  simplified word_bl.Rep', standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1567
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1568
lemma bl_word_roti_dt': 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24408
diff changeset
  1569
  "n = nat ((- i) mod int (size (w :: 'a :: len word))) ==> 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1570
    to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1571
  apply (unfold word_roti_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1572
  apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1573
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1574
   apply (simp add: zmod_zminus1_eq_if)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1575
   apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1576
    apply (simp add: nat_mult_distrib)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1577
   apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1578
                                      [THEN conjunct2, THEN order_less_imp_le]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1579
                    nat_mod_distrib)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1580
  apply (simp add: nat_mod_distrib)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1581
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1582
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1583
lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1584
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1585
lemmas word_rotl_dt = bl_word_rotl_dt 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1586
  [THEN word_bl.Rep_inverse' [symmetric], standard] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1587
lemmas word_rotr_dt = bl_word_rotr_dt 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1588
  [THEN word_bl.Rep_inverse' [symmetric], standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1589
lemmas word_roti_dt = bl_word_roti_dt 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1590
  [THEN word_bl.Rep_inverse' [symmetric], standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1591
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1592
lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 & word_rotl i 0 = 0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1593
  by (simp add : word_rotr_dt word_rotl_dt to_bl_0 replicate_add [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1594
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1595
lemma word_roti_0' [simp] : "word_roti n 0 = 0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1596
  unfolding word_roti_def by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1597
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1598
lemmas word_rotr_dt_no_bin' [simp] = 
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
  1599
  word_rotr_dt [where w="number_of w", unfolded to_bl_no_bin, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1600
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1601
lemmas word_rotl_dt_no_bin' [simp] = 
25350
a5fcf6d12a53 eliminated illegal schematic variables in where/of;
wenzelm
parents: 24465
diff changeset
  1602
  word_rotl_dt [where w="number_of w", unfolded to_bl_no_bin, standard]
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1603
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1604
declare word_roti_def [simp]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1605
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1606
end
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1607