src/HOL/Word/BinBoolList.thy
author huffman
Wed, 22 Aug 2007 16:44:35 +0200
changeset 24399 371f8c6b2101
parent 24396 c1e20c65a3be
child 24465 70f0214b3ecc
permissions -rw-r--r--
move if_simps from BinBoolList to Num_Lemmas
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
  ID:     $Id$
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     3
  Author: Jeremy Dawson, NICTA
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     4
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     5
  contains theorems to do with integers, expressed using Pls, Min, BIT,
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     6
  theorems linking them to lists of booleans, and repeated splitting 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     7
  and concatenation.
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     8
*) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     9
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    10
header "Bool lists and integers"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    11
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    12
theory BinBoolList imports BinOperations begin
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    13
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    14
subsection "Lemmas about standard list operations"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    15
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    16
lemma tl_take: "tl (take n l) = take (n - 1) (tl l)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    17
  apply (cases n, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    18
  apply (cases l, auto)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    19
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    20
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    21
lemma take_butlast [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    22
  "ALL n. n < length l --> take n (butlast l) = take n l"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    23
  apply (induct l, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    24
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    25
  apply (case_tac n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    26
  apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    27
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    28
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    29
lemma butlast_take [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    30
  "ALL n. n <= length l --> butlast (take n l) = take (n - 1) l"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    31
  apply (induct l, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    32
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    33
  apply (case_tac "n")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    34
   apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    35
   apply simp_all
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    36
  apply (case_tac "nat")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    37
   apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    38
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    39
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    40
lemma butlast_drop [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    41
  "ALL n. butlast (drop n l) = drop n (butlast l)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    42
  apply (induct l)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    43
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    44
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    45
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    46
   apply ((case_tac n, auto)[1])+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    47
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    48
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    49
lemma butlast_power:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    50
  "(butlast ^ n) bl = take (length bl - n) bl"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    51
  by (induct n) (auto simp: butlast_take)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    52
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    53
lemma nth_rev [rule_format] : 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    54
  "n < length xs --> rev xs ! n = xs ! (length xs - 1 - n)"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    55
  apply (induct_tac "xs")
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    56
   apply simp
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    57
  apply (clarsimp simp add : nth_append nth.simps split add : nat.split)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    58
  apply (rule_tac f = "%n. list ! n" in arg_cong) 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    59
  apply arith
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    60
  done
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    61
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    62
lemmas nth_rev_alt = nth_rev [where xs = "rev ?ys", simplified]
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    63
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    64
lemma hd_butlast: "size xs > 1 ==> hd (butlast xs) = hd xs"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    65
  by (cases xs) auto
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    66
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    67
subsection "From integers to bool lists"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    68
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    69
consts
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    70
  bin_to_bl :: "nat => int => bool list"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    71
  bin_to_bl_aux :: "nat => int => bool list => bool list"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    72
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    73
primrec
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    74
  Z : "bin_to_bl_aux 0 w bl = bl"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    75
  Suc : "bin_to_bl_aux (Suc n) w bl =
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    76
    bin_to_bl_aux n (bin_rest w) ((bin_last w = bit.B1) # bl)"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    77
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    78
defs
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    79
  bin_to_bl_def : "bin_to_bl n w == bin_to_bl_aux n w []"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
    80
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    81
lemma bin_to_bl_aux_Pls_minus_simp:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    82
  "0 < n ==> bin_to_bl_aux n Numeral.Pls bl = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    83
    bin_to_bl_aux (n - 1) Numeral.Pls (False # bl)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    84
  by (cases n) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    85
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    86
lemma bin_to_bl_aux_Min_minus_simp:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    87
  "0 < n ==> bin_to_bl_aux n Numeral.Min bl = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    88
    bin_to_bl_aux (n - 1) Numeral.Min (True # bl)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    89
  by (cases n) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    90
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    91
lemma bin_to_bl_aux_Bit_minus_simp:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    92
  "0 < n ==> bin_to_bl_aux n (w BIT b) bl = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    93
    bin_to_bl_aux (n - 1) w ((b = bit.B1) # bl)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    94
  by (cases n) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    95
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    96
declare bin_to_bl_aux_Pls_minus_simp [simp]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    97
  bin_to_bl_aux_Min_minus_simp [simp]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    98
  bin_to_bl_aux_Bit_minus_simp [simp]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    99
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   100
lemma bin_to_bl_aux_append [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   101
  "ALL w bs. bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   102
  by (induct n) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   103
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   104
lemma bin_to_bl_aux_alt: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   105
  "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   106
  unfolding bin_to_bl_def by (simp add : bin_to_bl_aux_append)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   107
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   108
lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   109
  unfolding bin_to_bl_def by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   110
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   111
lemma size_bin_to_bl_aux [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   112
  "ALL w bs. size (bin_to_bl_aux n w bs) = n + length bs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   113
  by (induct n) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   114
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   115
lemma size_bin_to_bl [simp]: "size (bin_to_bl n w) = n" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   116
  unfolding bin_to_bl_def by (simp add : size_bin_to_bl_aux)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   117
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   118
lemma bin_to_bl_Pls_aux [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   119
  "ALL bl. bin_to_bl_aux n Numeral.Pls bl = replicate n False @ bl"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   120
  by (induct n) (auto simp: replicate_app_Cons_same)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   121
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   122
lemma bin_to_bl_Pls: "bin_to_bl n Numeral.Pls = replicate n False"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   123
  unfolding bin_to_bl_def by (simp add : bin_to_bl_Pls_aux)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   124
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   125
lemma bin_to_bl_Min_aux [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   126
  "ALL bl. bin_to_bl_aux n Numeral.Min bl = replicate n True @ bl"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   127
  by (induct n) (auto simp: replicate_app_Cons_same)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   128
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   129
lemma bin_to_bl_Min: "bin_to_bl n Numeral.Min = replicate n True"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   130
  unfolding bin_to_bl_def by (simp add : bin_to_bl_Min_aux)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   131
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   132
lemma bin_to_bl_aux_bintr [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   133
  "ALL m bin bl. bin_to_bl_aux n (bintrunc m bin) bl = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   134
    replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   135
  apply (induct_tac "n")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   136
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   137
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   138
  apply (case_tac "m")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   139
   apply (clarsimp simp: bin_to_bl_Pls_aux) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   140
   apply (erule thin_rl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   141
   apply (induct_tac n)   
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
lemmas bin_to_bl_bintr = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   146
  bin_to_bl_aux_bintr [where bl = "[]", folded bin_to_bl_def]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   147
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   148
lemma len_bin_to_bl_aux [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   149
  "ALL w bs. length (bin_to_bl_aux n w bs) = n + length bs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   150
  by (induct "n") auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   151
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   152
lemma len_bin_to_bl [simp]: "length (bin_to_bl n w) = n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   153
  unfolding bin_to_bl_def len_bin_to_bl_aux by auto
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   154
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   155
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   156
subsection "From bool lists to integers"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   157
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   158
consts
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   159
  bl_to_bin :: "bool list => int"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   160
  bl_to_bin_aux :: "int => bool list => int"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   161
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   162
primrec
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   163
  Nil : "bl_to_bin_aux w [] = w"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   164
  Cons : "bl_to_bin_aux w (b # bs) = 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   165
      bl_to_bin_aux (w BIT (if b then bit.B1 else bit.B0)) bs"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   166
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   167
defs
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   168
  bl_to_bin_def : "bl_to_bin bs == bl_to_bin_aux Numeral.Pls bs"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   169
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   170
(** link between bin and bool list **)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   171
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   172
lemma bl_to_bin_aux_append [rule_format] : 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   173
  "ALL w. bl_to_bin_aux w (bs @ cs) = bl_to_bin_aux (bl_to_bin_aux w bs) cs"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   174
  by (induct bs) auto
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   175
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   176
lemma bl_to_bin_append: 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   177
  "bl_to_bin (bs @ cs) = bl_to_bin_aux (bl_to_bin bs) cs"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   178
  unfolding bl_to_bin_def by (rule bl_to_bin_aux_append)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   179
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   180
lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   181
  unfolding bl_to_bin_def by auto
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   182
  
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   183
lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = Numeral.Pls"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   184
  unfolding bl_to_bin_def by auto
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   185
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   186
lemma bl_to_bin_rep_F: 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   187
  "bl_to_bin (replicate n False @ bl) = bl_to_bin bl"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   188
  by (induct n) auto
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   189
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   190
lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = Numeral.Pls"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   191
  by (induct n) auto
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   192
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   193
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   194
subsection "Converting between bool lists and integers"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   195
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   196
lemma bin_bl_bin' [rule_format] : 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   197
  "ALL w bs. bl_to_bin (bin_to_bl_aux n w bs) = 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   198
    bl_to_bin_aux (bintrunc n w) bs"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   199
  by (induct n) (auto simp add : bl_to_bin_def)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   200
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   201
lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = bintrunc n w"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   202
  unfolding bin_to_bl_def bin_bl_bin' by auto
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   203
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   204
lemma bl_bin_bl' [rule_format] :
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   205
  "ALL w n. bin_to_bl (n + length bs) (bl_to_bin_aux w bs) = 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   206
    bin_to_bl_aux n w bs"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   207
  apply (induct "bs")
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   208
   apply auto
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   209
    apply (simp_all only : add_Suc [symmetric])
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   210
    apply (auto simp add : bin_to_bl_def)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   211
  done
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   212
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   213
lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   214
  unfolding bl_to_bin_def
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   215
  apply (rule box_equals)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   216
    apply (rule bl_bin_bl')
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   217
   prefer 2
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   218
   apply (rule bin_to_bl_aux.Z)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   219
  apply simp
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   220
  done
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   221
  
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   222
lemma bl_to_bin_inj:
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   223
  "bl_to_bin bs = bl_to_bin cs ==> length bs = length cs ==> bs = cs"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   224
  apply (rule_tac box_equals)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   225
    defer
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   226
    apply (rule bl_bin_bl)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   227
   apply (rule bl_bin_bl)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   228
  apply simp
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   229
  done
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   230
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   231
lemma bin_to_bl_trunc [simp]:
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   232
  "n <= m ==> bin_to_bl n (bintrunc m w) = bin_to_bl n w"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   233
  by (auto intro: bl_to_bin_inj)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   234
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   235
subsection "@{term bin_sign} with bool list operations"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   236
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   237
lemma sign_bl_bin' [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   238
  "ALL w. bin_sign (bl_to_bin_aux w bs) = bin_sign w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   239
  by (induct bs) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   240
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   241
lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = Numeral.Pls"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   242
  unfolding bl_to_bin_def by (simp add : sign_bl_bin')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   243
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   244
lemma bl_sbin_sign_aux [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   245
  "ALL w bs. hd (bin_to_bl_aux (Suc n) w bs) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   246
    (bin_sign (sbintrunc n w) = Numeral.Min)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   247
  apply (induct "n")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   248
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   249
   apply (case_tac w rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   250
   apply (simp split add : bit.split)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   251
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   252
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   253
    
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   254
lemma bl_sbin_sign: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   255
  "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = Numeral.Min)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   256
  unfolding bin_to_bl_def by (rule bl_sbin_sign_aux)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   257
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   258
subsection "@{term bin_nth} with bool list operations"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   259
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   260
lemma bin_nth_of_bl_aux [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   261
  "ALL w. bin_nth (bl_to_bin_aux w bl) n = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   262
    (n < size bl & rev bl ! n | n >= length bl & bin_nth w (n - size bl))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   263
  apply (induct_tac bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   264
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   265
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   266
  apply (cut_tac x=n and y="size list" in linorder_less_linear)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   267
  apply (erule disjE, simp add: nth_append)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   268
  apply (simp add: nth_append)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   269
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   270
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   271
lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl & rev bl ! n)";
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   272
  unfolding bl_to_bin_def by (simp add : bin_nth_of_bl_aux)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   273
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   274
lemma bin_nth_bl [rule_format] : "ALL m w. n < m --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   275
    bin_nth w n = nth (rev (bin_to_bl m w)) n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   276
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   277
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   278
   apply (case_tac m, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   279
   apply (clarsimp simp: bin_to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   280
   apply (simp add: bin_to_bl_aux_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   281
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   282
  apply (case_tac m, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   283
  apply (clarsimp simp: bin_to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   284
  apply (simp add: bin_to_bl_aux_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   285
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   286
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   287
lemma nth_bin_to_bl_aux [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   288
  "ALL w n bl. n < m + length bl --> (bin_to_bl_aux m w bl) ! n = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   289
    (if n < m then bin_nth w (m - 1 - n) else bl ! (n - m))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   290
  apply (induct_tac "m")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   291
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   292
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   293
  apply (case_tac w rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   294
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   295
  apply (case_tac "na - n")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   296
   apply arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   297
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   298
  apply (rule_tac f = "%n. bl ! n" in arg_cong) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   299
  apply arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   300
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   301
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   302
lemma nth_bin_to_bl: "n < m ==> (bin_to_bl m w) ! n = bin_nth w (m - Suc n)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   303
  unfolding bin_to_bl_def by (simp add : nth_bin_to_bl_aux)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   304
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   305
subsection "Ordering with bool list operations"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   306
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   307
lemma bl_to_bin_lt2p_aux [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   308
  "ALL w. bl_to_bin_aux w bs < (w + 1) * (2 ^ length bs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   309
  apply (induct "bs")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   310
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   311
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   312
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   313
   apply (erule allE, erule xtr8 [rotated],
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   314
          simp add: Bit_def ring_simps cong add : number_of_False_cong)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   315
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   316
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   317
lemma bl_to_bin_lt2p: "bl_to_bin bs < (2 ^ length bs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   318
  apply (unfold bl_to_bin_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   319
  apply (rule xtr1)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   320
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   321
   apply (rule bl_to_bin_lt2p_aux)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   322
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   323
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   324
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   325
lemma bl_to_bin_ge2p_aux [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   326
  "ALL w. bl_to_bin_aux w bs >= w * (2 ^ length bs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   327
  apply (induct bs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   328
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   329
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   330
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   331
   apply (erule allE, erule less_eq_less.order_trans [rotated],
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   332
          simp add: Bit_def ring_simps cong add : number_of_False_cong)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   333
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   334
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   335
lemma bl_to_bin_ge0: "bl_to_bin bs >= 0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   336
  apply (unfold bl_to_bin_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   337
  apply (rule xtr4)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   338
   apply (rule bl_to_bin_ge2p_aux)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   339
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   340
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   341
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   342
subsection "@{term butlast} with bool list operations"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   343
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   344
lemma butlast_rest_bin: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   345
  "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   346
  apply (unfold bin_to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   347
  apply (cases w rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   348
  apply (cases n, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   349
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   350
  apply (auto simp add: bin_to_bl_aux_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   351
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   352
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   353
lemmas butlast_bin_rest = butlast_rest_bin
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   354
  [where w="bl_to_bin ?bl" and n="length ?bl", simplified]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   355
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   356
lemma butlast_rest_bl2bin_aux [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   357
  "ALL w. bl ~= [] --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   358
    bl_to_bin_aux w (butlast bl) = bin_rest (bl_to_bin_aux w bl)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   359
  by (induct bl) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   360
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   361
lemma butlast_rest_bl2bin: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   362
  "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   363
  apply (unfold bl_to_bin_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   364
  apply (cases bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   365
   apply (auto simp add: butlast_rest_bl2bin_aux)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   366
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   367
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   368
subsection "Truncation"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   369
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   370
lemma trunc_bl2bin_aux [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   371
  "ALL w. bintrunc m (bl_to_bin_aux w bl) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   372
    bl_to_bin_aux (bintrunc (m - length bl) w) (drop (length bl - m) bl)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   373
  apply (induct_tac bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   374
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   375
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   376
  apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   377
   apply (case_tac "m - size list")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   378
    apply (simp add : diff_is_0_eq [THEN iffD1, THEN Suc_diff_le])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   379
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   380
   apply (rule_tac f = "%nat. bl_to_bin_aux (bintrunc nat w BIT bit.B1) list" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   381
                   in arg_cong) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   382
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   383
  apply (case_tac "m - size list")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   384
   apply (simp add: diff_is_0_eq [THEN iffD1, THEN Suc_diff_le])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   385
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   386
  apply (rule_tac f = "%nat. bl_to_bin_aux (bintrunc nat w BIT bit.B0) list" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   387
                  in arg_cong) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   388
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   389
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   390
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   391
lemma trunc_bl2bin: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   392
  "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   393
  unfolding bl_to_bin_def by (simp add : trunc_bl2bin_aux)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   394
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   395
lemmas trunc_bl2bin_len [simp] =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   396
  trunc_bl2bin [of "length bl" bl, simplified, standard]  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   397
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   398
lemma bl2bin_drop: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   399
  "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   400
  apply (rule trans)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   401
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   402
   apply (rule trunc_bl2bin [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   403
  apply (cases "k <= length bl")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   404
   apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   405
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   406
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   407
(* TODO: This is not related to bool lists; should be moved *)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   408
lemma nth_rest_power_bin [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   409
  "ALL n. bin_nth ((bin_rest ^ k) w) n = bin_nth w (n + k)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   410
  apply (induct k, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   411
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   412
  apply (simp only: bin_nth.Suc [symmetric] add_Suc)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   413
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   414
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   415
lemma take_rest_power_bin:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   416
  "m <= n ==> take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^ (n - m)) w)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   417
  apply (rule nth_equalityI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   418
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   419
  apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   420
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   421
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   422
subsection "@{term last} with bool list operations"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   423
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   424
lemma last_bin_last' [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   425
  "ALL w. size xs > 0 --> last xs = (bin_last (bl_to_bin_aux w xs) = bit.B1)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   426
  by (induct xs) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   427
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   428
lemma last_bin_last: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   429
  "size xs > 0 ==> last xs = (bin_last (bl_to_bin xs) = bit.B1)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   430
  unfolding bl_to_bin_def by (erule last_bin_last')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   431
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   432
lemma bin_last_last: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   433
  "bin_last w = (if last (bin_to_bl (Suc n) w) then bit.B1 else bit.B0)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   434
  apply (unfold bin_to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   435
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   436
  apply (auto simp add: bin_to_bl_aux_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   437
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   438
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   439
subsection "Bit-wise operations on bool lists"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   440
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   441
consts
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   442
  app2 :: "('a => 'b => 'c) => 'a list => 'b list => 'c list"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   443
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   444
defs
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   445
  app2_def : "app2 f as bs == map (split f) (zip as bs)"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   446
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   447
lemma app2_Nil [simp]: "app2 f [] ys = []"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   448
  unfolding app2_def by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   449
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   450
lemma app2_Cons [simp]:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   451
  "app2 f (x # xs) (y # ys) = f x y # app2 f xs ys"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   452
  unfolding app2_def by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   453
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   454
lemma bl_xor_aux_bin [rule_format] : "ALL v w bs cs. 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   455
    app2 (%x y. x ~= y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = 
24353
9a7a9b19e925 use overloaded bitwise operators at type int
huffman
parents: 24350
diff changeset
   456
    bin_to_bl_aux n (v XOR w) (app2 (%x y. x ~= y) bs cs)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   457
  apply (induct_tac n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   458
   apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   459
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   460
  apply (case_tac v rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   461
  apply (case_tac w rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   462
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   463
  apply (case_tac b)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   464
  apply (case_tac ba, safe, simp_all)+
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
lemma bl_or_aux_bin [rule_format] : "ALL v w bs cs. 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   468
    app2 (op | ) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = 
24353
9a7a9b19e925 use overloaded bitwise operators at type int
huffman
parents: 24350
diff changeset
   469
    bin_to_bl_aux n (v OR w) (app2 (op | ) bs cs)" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   470
  apply (induct_tac n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   471
   apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   472
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   473
  apply (case_tac v rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   474
  apply (case_tac w rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   475
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   476
  apply (case_tac b)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   477
  apply (case_tac ba, safe, simp_all)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   478
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   479
    
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   480
lemma bl_and_aux_bin [rule_format] : "ALL v w bs cs. 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   481
    app2 (op & ) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = 
24353
9a7a9b19e925 use overloaded bitwise operators at type int
huffman
parents: 24350
diff changeset
   482
    bin_to_bl_aux n (v AND w) (app2 (op & ) bs cs)" 
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   483
  apply (induct_tac n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   484
   apply safe
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   485
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   486
  apply (case_tac v rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   487
  apply (case_tac w rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   488
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   489
  apply (case_tac b)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   490
  apply (case_tac ba, safe, simp_all)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   491
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   492
    
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   493
lemma bl_not_aux_bin [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   494
  "ALL w cs. map Not (bin_to_bl_aux n w cs) = 
24353
9a7a9b19e925 use overloaded bitwise operators at type int
huffman
parents: 24350
diff changeset
   495
    bin_to_bl_aux n (NOT w) (map Not cs)"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   496
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   497
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   498
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   499
  apply (case_tac w rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   500
  apply (case_tac b)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   501
   apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   502
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   503
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   504
lemmas bl_not_bin = bl_not_aux_bin
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   505
  [where cs = "[]", unfolded bin_to_bl_def [symmetric] map.simps]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   506
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   507
lemmas bl_and_bin = bl_and_aux_bin [where bs="[]" and cs="[]", 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   508
                                    unfolded app2_Nil, folded bin_to_bl_def]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   509
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   510
lemmas bl_or_bin = bl_or_aux_bin [where bs="[]" and cs="[]", 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   511
                                  unfolded app2_Nil, folded bin_to_bl_def]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   512
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   513
lemmas bl_xor_bin = bl_xor_aux_bin [where bs="[]" and cs="[]", 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   514
                                    unfolded app2_Nil, folded bin_to_bl_def]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   515
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   516
subsection "@{term drop} with bool list operations"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   517
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   518
lemma drop_bin2bl_aux [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   519
  "ALL m bin bs. drop m (bin_to_bl_aux n bin bs) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   520
    bin_to_bl_aux (n - m) bin (drop (m - n) bs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   521
  apply (induct n, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   522
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   523
  apply (case_tac bin rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   524
  apply (case_tac "m <= n", simp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   525
  apply (case_tac "m - n", simp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   526
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   527
  apply (rule_tac f = "%nat. drop nat bs" in arg_cong) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   528
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   529
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   530
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   531
lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   532
  unfolding bin_to_bl_def by (simp add : drop_bin2bl_aux)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   533
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   534
lemma take_bin2bl_lem1 [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   535
  "ALL w bs. take m (bin_to_bl_aux m w bs) = bin_to_bl m w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   536
  apply (induct m, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   537
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   538
  apply (simp add: bin_to_bl_aux_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   539
  apply (simp add: bin_to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   540
  apply (simp add: bin_to_bl_aux_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   541
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   542
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   543
lemma take_bin2bl_lem [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   544
  "ALL w bs. take m (bin_to_bl_aux (m + n) w bs) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   545
    take m (bin_to_bl (m + n) w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   546
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   547
   apply clarify
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   548
   apply (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   549
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   550
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   551
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   552
subsection "@{term bin_split} with bool list operations"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   553
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   554
lemma bin_split_take [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   555
  "ALL b c. bin_split n c = (a, b) --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   556
    bin_to_bl m a = take m (bin_to_bl (m + n) c)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   557
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   558
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   559
  apply (clarsimp simp: Let_def split: ls_splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   560
  apply (simp add: bin_to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   561
  apply (simp add: take_bin2bl_lem)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   562
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   563
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   564
lemma bin_split_take1: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   565
  "k = m + n ==> bin_split n c = (a, b) ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   566
    bin_to_bl m a = take m (bin_to_bl k c)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   567
  by (auto elim: bin_split_take)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   568
  
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   569
subsection "@{term takefill}"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   570
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   571
consts
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   572
  takefill :: "'a => nat => 'a list => 'a list"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   573
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   574
-- "takefill - like take but if argument list too short,"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   575
-- "extends result to get requested length"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   576
primrec
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   577
  Z : "takefill fill 0 xs = []"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   578
  Suc : "takefill fill (Suc n) xs = (
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   579
    case xs of [] => fill # takefill fill n xs
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   580
      | y # ys => y # takefill fill n ys)"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   581
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   582
lemma nth_takefill [rule_format] : "ALL m l. m < n --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   583
    takefill fill n l ! m = (if m < length l then l ! m else fill)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   584
  apply (induct n, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   585
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   586
  apply (case_tac m)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   587
   apply (simp split: list.split)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   588
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   589
  apply (erule allE)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   590
  apply (erule (1) impE)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   591
  apply (simp split: list.split)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   592
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   593
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   594
lemma takefill_alt [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   595
  "ALL l. takefill fill n l = take n l @ replicate (n - length l) fill"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   596
  by (induct n) (auto split: list.split)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   597
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   598
lemma takefill_replicate [simp]:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   599
  "takefill fill n (replicate m fill) = replicate n fill"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   600
  by (simp add : takefill_alt replicate_add [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   601
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   602
lemma takefill_le' [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   603
  "ALL l n. n = m + k --> takefill x m (takefill x n l) = takefill x m l"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   604
  by (induct m) (auto split: list.split)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   605
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   606
lemma length_takefill [simp]: "length (takefill fill n l) = n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   607
  by (simp add : takefill_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   608
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   609
lemma take_takefill':
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   610
  "!!w n.  n = k + m ==> take k (takefill fill n w) = takefill fill k w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   611
  by (induct k) (auto split add : list.split) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   612
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   613
lemma drop_takefill:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   614
  "!!w. drop k (takefill fill (m + k) w) = takefill fill m (drop k w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   615
  by (induct k) (auto split add : list.split) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   616
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   617
lemma takefill_le [simp]:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   618
  "m \<le> n \<Longrightarrow> takefill x m (takefill x n l) = takefill x m l"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   619
  by (auto simp: le_iff_add takefill_le')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   620
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   621
lemma take_takefill [simp]:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   622
  "m \<le> n \<Longrightarrow> take m (takefill fill n w) = takefill fill m w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   623
  by (auto simp: le_iff_add take_takefill')
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   624
 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   625
lemma takefill_append:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   626
  "takefill fill (m + length xs) (xs @ w) = xs @ (takefill fill m w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   627
  by (induct xs) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   628
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   629
lemma takefill_same': 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   630
  "l = length xs ==> takefill fill l xs = xs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   631
  by clarify (induct xs, auto)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   632
 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   633
lemmas takefill_same [simp] = takefill_same' [OF refl]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   634
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   635
lemma takefill_bintrunc:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   636
  "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   637
  apply (rule nth_equalityI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   638
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   639
  apply (clarsimp simp: nth_takefill nth_rev nth_bin_to_bl bin_nth_of_bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   640
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   641
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   642
lemma bl_bin_bl_rtf:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   643
  "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   644
  by (simp add : takefill_bintrunc)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   645
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   646
lemmas bl_bin_bl_rep_drop = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   647
  bl_bin_bl_rtf [simplified takefill_alt,
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   648
                 simplified, simplified rev_take, simplified]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   649
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   650
lemma tf_rev:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   651
  "n + k = m + length bl ==> takefill x m (rev (takefill y n bl)) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   652
    rev (takefill y m (rev (takefill x k (rev bl))))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   653
  apply (rule nth_equalityI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   654
   apply (auto simp add: nth_takefill nth_rev)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   655
  apply (rule_tac f = "%n. bl ! n" in arg_cong) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   656
  apply arith 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   657
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   658
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   659
lemma takefill_minus:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   660
  "0 < n ==> takefill fill (Suc (n - 1)) w = takefill fill n w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   661
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   662
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   663
lemmas takefill_Suc_cases = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   664
  list.cases [THEN takefill.Suc [THEN trans], standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   665
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   666
lemmas takefill_Suc_Nil = takefill_Suc_cases (1)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   667
lemmas takefill_Suc_Cons = takefill_Suc_cases (2)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   668
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   669
lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2] 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   670
  takefill_minus [symmetric, THEN trans], standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   671
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   672
lemmas takefill_pred_simps [simp] =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   673
  takefill_minus_simps [where n="number_of bin", simplified nobm1, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   674
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   675
subsection "@{term bin_cat} with bool list operations"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   676
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   677
(* links with function bl_to_bin *)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   678
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   679
lemma bl_to_bin_aux_cat: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   680
  "!!nv v. bl_to_bin_aux (bin_cat w nv v) bs = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   681
    bin_cat w (nv + length bs) (bl_to_bin_aux v bs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   682
  apply (induct bs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   683
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   684
  apply (simp add: bin_cat_Suc_Bit [symmetric] del: bin_cat.simps)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   685
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   686
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   687
lemma bin_to_bl_aux_cat: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   688
  "!!w bs. bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   689
    bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   690
  by (induct nw) auto 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   691
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   692
lemmas bl_to_bin_aux_alt = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   693
  bl_to_bin_aux_cat [where nv = "0" and v = "Numeral.Pls", 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   694
    simplified bl_to_bin_def [symmetric], simplified]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   695
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   696
lemmas bin_to_bl_cat =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   697
  bin_to_bl_aux_cat [where bs = "[]", folded bin_to_bl_def]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   698
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   699
lemmas bl_to_bin_aux_app_cat = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   700
  trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   701
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   702
lemmas bin_to_bl_aux_cat_app =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   703
  trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   704
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   705
lemmas bl_to_bin_app_cat = bl_to_bin_aux_app_cat
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   706
  [where w = "Numeral.Pls", folded bl_to_bin_def]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   707
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   708
lemmas bin_to_bl_cat_app = bin_to_bl_aux_cat_app
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   709
  [where bs = "[]", folded bin_to_bl_def]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   710
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   711
(* bl_to_bin_app_cat_alt and bl_to_bin_app_cat are easily interderivable *)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   712
lemma bl_to_bin_app_cat_alt: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   713
  "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   714
  by (simp add : bl_to_bin_app_cat)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   715
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   716
lemma mask_lem: "(bl_to_bin (True # replicate n False)) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   717
    Numeral.succ (bl_to_bin (replicate n True))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   718
  apply (unfold bl_to_bin_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   719
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   720
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   721
  apply (simp only: Suc_eq_add_numeral_1 replicate_add
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   722
                    append_Cons [symmetric] bl_to_bin_aux_append)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   723
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   724
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   725
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   726
subsection "function @{term bl_of_nth}"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   727
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   728
consts
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   729
  bl_of_nth :: "nat => (nat => bool) => bool list"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   730
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   731
primrec
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   732
  Suc : "bl_of_nth (Suc n) f = f n # bl_of_nth n f"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   733
  Z : "bl_of_nth 0 f = []"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   734
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   735
lemma length_bl_of_nth [simp]: "length (bl_of_nth n f) = n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   736
  by (induct n)  auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   737
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   738
lemma nth_bl_of_nth [simp]:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   739
  "m < n \<Longrightarrow> rev (bl_of_nth n f) ! m = f m"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   740
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   741
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   742
  apply (clarsimp simp add : nth_append)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   743
  apply (rule_tac f = "f" in arg_cong) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   744
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   745
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   746
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   747
lemma bl_of_nth_inj: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   748
  "(!!k. k < n ==> f k = g k) ==> bl_of_nth n f = bl_of_nth n g"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   749
  by (induct n)  auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   750
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   751
lemma bl_of_nth_nth_le [rule_format] : "ALL xs. 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   752
    length xs >= n --> bl_of_nth n (nth (rev xs)) = drop (length xs - n) xs";
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   753
  apply (induct n, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   754
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   755
  apply (rule trans [OF _ hd_Cons_tl])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   756
   apply (frule Suc_le_lessD)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   757
   apply (simp add: nth_rev trans [OF drop_Suc drop_tl, symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   758
   apply (subst hd_drop_conv_nth)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   759
     apply force
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   760
    apply simp_all
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   761
  apply (rule_tac f = "%n. drop n xs" in arg_cong) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   762
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   763
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   764
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   765
lemmas bl_of_nth_nth [simp] = order_refl [THEN bl_of_nth_nth_le, simplified]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   766
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   767
subsection "Arithmetic in terms of bool lists"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   768
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   769
consts (* arithmetic operations in terms of the reversed bool list,
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   770
  assuming input list(s) the same length, and don't extend them *)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   771
  rbl_succ :: "bool list => bool list"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   772
  rbl_pred :: "bool list => bool list"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   773
  rbl_add :: "bool list => bool list => bool list"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   774
  rbl_mult :: "bool list => bool list => bool list"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   775
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   776
primrec 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   777
  Nil: "rbl_succ Nil = Nil"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   778
  Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   779
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   780
primrec 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   781
  Nil : "rbl_pred Nil = Nil"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   782
  Cons : "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   783
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   784
primrec (* result is length of first arg, second arg may be longer *)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   785
  Nil : "rbl_add Nil x = Nil"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   786
  Cons : "rbl_add (y # ys) x = (let ws = rbl_add ys (tl x) in 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   787
    (y ~= hd x) # (if hd x & y then rbl_succ ws else ws))"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   788
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   789
primrec (* result is length of first arg, second arg may be longer *)
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   790
  Nil : "rbl_mult Nil x = Nil"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   791
  Cons : "rbl_mult (y # ys) x = (let ws = False # rbl_mult ys x in 
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   792
    if y then rbl_add ws x else ws)"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   793
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   794
lemma size_rbl_pred: "length (rbl_pred bl) = length bl"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   795
  by (induct bl) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   796
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   797
lemma size_rbl_succ: "length (rbl_succ bl) = length bl"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   798
  by (induct bl) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   799
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   800
lemma size_rbl_add:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   801
  "!!cl. length (rbl_add bl cl) = length bl"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   802
  by (induct bl) (auto simp: Let_def size_rbl_succ)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   803
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   804
lemma size_rbl_mult: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   805
  "!!cl. length (rbl_mult bl cl) = length bl"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   806
  by (induct bl) (auto simp add : Let_def size_rbl_add)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   807
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   808
lemmas rbl_sizes [simp] = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   809
  size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   810
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   811
lemmas rbl_Nils =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   812
  rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   813
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   814
lemma rbl_pred: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   815
  "!!bin. rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (Numeral.pred bin))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   816
  apply (induct n, simp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   817
  apply (unfold bin_to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   818
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   819
  apply (case_tac bin rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   820
  apply (case_tac b)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   821
   apply (clarsimp simp: bin_to_bl_aux_alt)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   822
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   823
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   824
lemma rbl_succ: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   825
  "!!bin. rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (Numeral.succ bin))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   826
  apply (induct n, simp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   827
  apply (unfold bin_to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   828
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   829
  apply (case_tac bin rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   830
  apply (case_tac b)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   831
   apply (clarsimp simp: bin_to_bl_aux_alt)+
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
lemma rbl_add: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   835
  "!!bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   836
    rev (bin_to_bl n (bina + binb))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   837
  apply (induct n, simp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   838
  apply (unfold bin_to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   839
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   840
  apply (case_tac bina rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   841
  apply (case_tac binb rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   842
  apply (case_tac b)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   843
   apply (case_tac [!] "ba")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   844
     apply (auto simp: rbl_succ succ_def bin_to_bl_aux_alt Let_def add_ac)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   845
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   846
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   847
lemma rbl_add_app2: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   848
  "!!blb. length blb >= length bla ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   849
    rbl_add bla (blb @ blc) = rbl_add bla blb"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   850
  apply (induct bla, simp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   851
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   852
  apply (case_tac blb, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   853
  apply (clarsimp simp: Let_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   854
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   855
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   856
lemma rbl_add_take2: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   857
  "!!blb. length blb >= length bla ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   858
    rbl_add bla (take (length bla) blb) = rbl_add bla blb"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   859
  apply (induct bla, simp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   860
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   861
  apply (case_tac blb, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   862
  apply (clarsimp simp: Let_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   863
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   864
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   865
lemma rbl_add_long: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   866
  "m >= n ==> rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   867
    rev (bin_to_bl n (bina + binb))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   868
  apply (rule box_equals [OF _ rbl_add_take2 rbl_add])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   869
   apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   870
   apply (rule rev_swap [THEN iffD1])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   871
   apply (simp add: rev_take drop_bin2bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   872
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   873
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   874
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   875
lemma rbl_mult_app2:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   876
  "!!blb. length blb >= length bla ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   877
    rbl_mult bla (blb @ blc) = rbl_mult bla blb"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   878
  apply (induct bla, simp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   879
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   880
  apply (case_tac blb, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   881
  apply (clarsimp simp: Let_def rbl_add_app2)
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
lemma rbl_mult_take2: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   885
  "length blb >= length bla ==> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   886
    rbl_mult bla (take (length bla) blb) = rbl_mult bla blb"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   887
  apply (rule trans)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   888
   apply (rule rbl_mult_app2 [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   889
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   890
  apply (rule_tac f = "rbl_mult bla" in arg_cong) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   891
  apply (rule append_take_drop_id)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   892
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   893
    
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   894
lemma rbl_mult_gt1: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   895
  "m >= length bl ==> rbl_mult bl (rev (bin_to_bl m binb)) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   896
    rbl_mult bl (rev (bin_to_bl (length bl) binb))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   897
  apply (rule trans)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   898
   apply (rule rbl_mult_take2 [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   899
   apply simp_all
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   900
  apply (rule_tac f = "rbl_mult bl" in arg_cong) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   901
  apply (rule rev_swap [THEN iffD1])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   902
  apply (simp add: rev_take drop_bin2bl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   903
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   904
    
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   905
lemma rbl_mult_gt: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   906
  "m > n ==> rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   907
    rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   908
  by (auto intro: trans [OF rbl_mult_gt1])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   909
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   910
lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   911
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   912
lemma rbbl_Cons: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   913
  "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (x BIT If b bit.B1 bit.B0))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   914
  apply (unfold bin_to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   915
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   916
  apply (simp add: bin_to_bl_aux_alt)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   917
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   918
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   919
lemma rbl_mult: "!!bina binb. 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   920
    rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   921
    rev (bin_to_bl n (bina * binb))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   922
  apply (induct n)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   923
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   924
  apply (unfold bin_to_bl_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   925
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   926
  apply (case_tac bina rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   927
  apply (case_tac binb rule: bin_exhaust)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   928
  apply (case_tac b)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   929
   apply (case_tac [!] "ba")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   930
     apply (auto simp: bin_to_bl_aux_alt Let_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   931
     apply (auto simp: rbbl_Cons rbl_mult_Suc rbl_add)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   932
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   933
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   934
lemma rbl_add_split: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   935
  "P (rbl_add (y # ys) (x # xs)) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   936
    (ALL ws. length ws = length ys --> ws = rbl_add ys xs --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   937
    (y --> ((x --> P (False # rbl_succ ws)) & (~ x -->  P (True # ws)))) & \ 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   938
    (~ y --> P (x # ws)))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   939
  apply (auto simp add: Let_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   940
   apply (case_tac [!] "y")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   941
     apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   942
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   943
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   944
lemma rbl_mult_split: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   945
  "P (rbl_mult (y # ys) xs) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   946
    (ALL ws. length ws = Suc (length ys) --> ws = False # rbl_mult ys xs --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   947
    (y --> P (rbl_add ws xs)) & (~ y -->  P ws))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   948
  by (clarsimp simp add : Let_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   949
  
24396
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   950
subsection "Miscellaneous lemmas"
c1e20c65a3be move bool list stuff from BinOperations to BinBoolList;
huffman
parents: 24353
diff changeset
   951
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   952
lemma and_len: "xs = ys ==> xs = ys & length xs = length ys"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   953
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   954
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   955
lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   956
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   957
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   958
lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   959
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   960
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   961
lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   962
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   963
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   964
lemma if_same_and: "(If p x y & If p u v) = (if p then x & u else y & v)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   965
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   966
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   967
lemma if_same_eq: "(If p x y  = (If p u v)) = (if p then x = (u) else y = (v))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   968
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   969
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   970
lemma if_same_eq_not:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   971
  "(If p x y  = (~ If p u v)) = (if p then x = (~u) else y = (~v))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   972
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   973
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   974
(* note - if_Cons can cause blowup in the size, if p is complex,
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   975
  so make a simproc *)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   976
lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   977
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   978
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   979
lemma if_single:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   980
  "(if xc then [xab] else [an]) = [if xc then xab else an]"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   981
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   982
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   983
lemmas seqr = eq_reflection [where x = "size ?w"]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   984
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   985
lemmas tl_Nil = tl.simps (1)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   986
lemmas tl_Cons = tl.simps (2)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   987
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   988
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
   989
subsection "Repeated splitting or concatenation"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   990
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   991
lemma sclem:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   992
  "size (concat (map (bin_to_bl n) xs)) = length xs * n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   993
  by (induct xs) auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   994
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   995
lemma bin_cat_foldl_lem [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   996
  "ALL x. foldl (%u. bin_cat u n) x xs = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   997
    bin_cat x (size xs * n) (foldl (%u. bin_cat u n) y xs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   998
  apply (induct xs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   999
   apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1000
  apply clarify
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1001
  apply (simp (no_asm))
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1002
  apply (frule asm_rl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1003
  apply (drule spec)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1004
  apply (erule trans)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1005
  apply (drule_tac x = "bin_cat y n a" in spec) 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1006
  apply (simp add : bin_cat_assoc_sym min_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1007
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1008
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1009
lemma bin_rcat_bl:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1010
  "(bin_rcat n wl) = bl_to_bin (concat (map (bin_to_bl n) wl))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1011
  apply (unfold bin_rcat_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1012
  apply (rule sym)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1013
  apply (induct wl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1014
   apply (auto simp add : bl_to_bin_append)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1015
  apply (simp add : bl_to_bin_aux_alt sclem)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1016
  apply (simp add : bin_cat_foldl_lem [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1017
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1018
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1019
lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1020
lemmas rsplit_aux_simps = bin_rsplit_aux_simps
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1021
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1022
lemmas th_if_simp1 = split_if [where P = "op = ?l",
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1023
  THEN iffD1, THEN conjunct1, THEN mp, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1024
lemmas th_if_simp2 = split_if [where P = "op = ?l",
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1025
  THEN iffD1, THEN conjunct2, THEN mp, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1026
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1027
lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1028
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1029
lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1030
(* these safe to [simp add] as require calculating m - n *)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1031
lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1032
lemmas rbscl = bin_rsplit_aux_simp2s (2)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1033
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1034
lemmas rsplit_aux_0_simps [simp] = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1035
  rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1036
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1037
lemma bin_rsplit_aux_append:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1038
  "bin_rsplit_aux (n, bs @ cs, m, c) = bin_rsplit_aux (n, bs, m, c) @ cs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1039
  apply (rule_tac u=n and v=bs and w=m and x=c in bin_rsplit_aux.induct)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1040
  apply (subst bin_rsplit_aux.simps)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1041
  apply (subst bin_rsplit_aux.simps)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1042
  apply (clarsimp split: ls_splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1043
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1044
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1045
lemma bin_rsplitl_aux_append:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1046
  "bin_rsplitl_aux (n, bs @ cs, m, c) = bin_rsplitl_aux (n, bs, m, c) @ cs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1047
  apply (rule_tac u=n and v=bs and w=m and x=c in bin_rsplitl_aux.induct)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1048
  apply (subst bin_rsplitl_aux.simps)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1049
  apply (subst bin_rsplitl_aux.simps)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1050
  apply (clarsimp split: ls_splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1051
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1052
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1053
lemmas rsplit_aux_apps [where bs = "[]"] =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1054
  bin_rsplit_aux_append bin_rsplitl_aux_append
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1055
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1056
lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1057
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1058
lemmas rsplit_aux_alts = rsplit_aux_apps 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1059
  [unfolded append_Nil rsplit_def_auxs [symmetric]]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1060
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1061
lemma bin_split_minus: "0 < n ==> bin_split (Suc (n - 1)) w = bin_split n w"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1062
  by auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1063
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1064
lemmas bin_split_minus_simp =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1065
  bin_split.Suc [THEN [2] bin_split_minus [symmetric, THEN trans], standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1066
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1067
lemma bin_split_pred_simp [simp]: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1068
  "(0::nat) < number_of bin \<Longrightarrow>
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1069
  bin_split (number_of bin) w =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1070
  (let (w1, w2) = bin_split (number_of (Numeral.pred bin)) (bin_rest w)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1071
   in (w1, w2 BIT bin_last w))" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1072
  by (simp only: nobm1 bin_split_minus_simp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1073
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1074
lemma bin_rsplit_aux_simp_alt:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1075
  "bin_rsplit_aux (n, bs, m, c) =
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1076
   (if m = 0 \<or> n = 0 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1077
   then bs
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1078
   else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1079
  apply (rule trans)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1080
   apply (subst bin_rsplit_aux.simps, rule refl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1081
  apply (simp add: rsplit_aux_alts)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1082
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1083
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1084
lemmas bin_rsplit_simp_alt = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1085
  trans [OF bin_rsplit_def [THEN meta_eq_to_obj_eq]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1086
            bin_rsplit_aux_simp_alt, standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1087
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1088
lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1089
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1090
lemma bin_rsplit_size_sign' [rule_format] : 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1091
  "n > 0 ==> (ALL nw w. rev sw = bin_rsplit n (nw, w) --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1092
    (ALL v: set sw. bintrunc n v = v))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1093
  apply (induct sw)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1094
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1095
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1096
  apply (drule bthrs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1097
  apply (simp (no_asm_use) add: Let_def split: ls_splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1098
  apply clarify
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1099
  apply (erule impE, rule exI, erule exI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1100
  apply (drule split_bintrunc)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1101
  apply simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1102
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1103
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1104
lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1105
  rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]],
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1106
  standard]
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1107
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1108
lemma bin_nth_rsplit [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1109
  "n > 0 ==> m < n ==> (ALL w k nw. rev sw = bin_rsplit n (nw, w) --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1110
       k < size sw --> bin_nth (sw ! k) m = bin_nth w (k * n + m))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1111
  apply (induct sw)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1112
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1113
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1114
  apply (drule bthrs)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1115
  apply (simp (no_asm_use) add: Let_def split: ls_splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1116
  apply clarify
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1117
  apply (erule allE, erule impE, erule exI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1118
  apply (case_tac k)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1119
   apply clarsimp   
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1120
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1121
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1122
   apply (erule allE)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1123
   apply (erule (1) impE)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1124
   apply (drule bin_nth_split, erule conjE, erule allE,
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1125
          erule trans, simp add : add_ac)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1126
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1127
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1128
lemma bin_rsplit_all:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1129
  "0 < nw ==> nw <= n ==> bin_rsplit n (nw, w) = [bintrunc n w]"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1130
  unfolding bin_rsplit_def
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1131
  by (clarsimp dest!: split_bintrunc simp: rsplit_aux_simp2ls split: ls_splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1132
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1133
lemma bin_rsplit_l [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1134
  "ALL bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, bintrunc m bin)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1135
  apply (rule_tac a = "m" in wf_less_than [THEN wf_induct])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1136
  apply (simp (no_asm) add : bin_rsplitl_def bin_rsplit_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1137
  apply (rule allI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1138
  apply (subst bin_rsplitl_aux.simps)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1139
  apply (subst bin_rsplit_aux.simps)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1140
  apply (clarsimp simp: rsplit_aux_alts Let_def split: ls_splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1141
  apply (drule bin_split_trunc)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1142
  apply (drule sym [THEN trans], assumption)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1143
  apply fast
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1144
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1145
    
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1146
lemma bin_rsplit_rcat [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1147
  "n > 0 --> bin_rsplit n (n * size ws, bin_rcat n ws) = map (bintrunc n) ws"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1148
  apply (unfold bin_rsplit_def bin_rcat_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1149
  apply (rule_tac xs = "ws" in rev_induct)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1150
   apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1151
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1152
  apply (clarsimp simp add: bin_split_cat rsplit_aux_alts)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1153
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1154
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1155
lemma bin_rsplit_aux_len_le [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1156
  "ALL ws m. n > 0 --> ws = bin_rsplit_aux (n, bs, nw, w) --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1157
    (length ws <= m) = (nw + length bs * n <= m * n)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1158
  apply (rule_tac u=n and v=bs and w=nw and x=w in bin_rsplit_aux.induct)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1159
  apply (subst bin_rsplit_aux.simps)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1160
  apply (clarsimp simp: Let_def split: ls_splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1161
  apply (erule lrlem)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1162
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1163
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1164
lemma bin_rsplit_len_le: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1165
  "n > 0 --> ws = bin_rsplit n (nw, w) --> (length ws <= m) = (nw <= m * n)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1166
  unfolding bin_rsplit_def by (clarsimp simp add : bin_rsplit_aux_len_le)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1167
 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1168
lemma bin_rsplit_aux_len [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1169
  "0 < n --> length (bin_rsplit_aux (n, cs, nw, w)) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1170
    (nw + n - 1) div n + length cs"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1171
  apply (rule_tac u=n and v=cs and w=nw and x=w in bin_rsplit_aux.induct)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1172
  apply (subst bin_rsplit_aux.simps)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1173
  apply (clarsimp simp: Let_def split: ls_splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1174
  apply (erule thin_rl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1175
  apply (case_tac "m <= n")
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1176
   prefer 2
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1177
   apply (simp add: div_add_self2 [symmetric])
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1178
  apply (case_tac m, clarsimp)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1179
  apply (simp add: div_add_self2)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1180
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1181
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1182
lemma bin_rsplit_len: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1183
  "0 < n ==> length (bin_rsplit n (nw, w)) = (nw + n - 1) div n"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1184
  unfolding bin_rsplit_def by (clarsimp simp add : bin_rsplit_aux_len)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1185
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1186
lemma bin_rsplit_aux_len_indep [rule_format] :
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1187
  "0 < n ==> (ALL v bs. length bs = length cs --> 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1188
    length (bin_rsplit_aux (n, bs, nw, v)) = 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1189
    length (bin_rsplit_aux (n, cs, nw, w)))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1190
  apply (rule_tac u=n and v=cs and w=nw and x=w in bin_rsplit_aux.induct)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1191
  apply clarsimp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1192
  apply (simp (no_asm_simp) add: bin_rsplit_aux_simp_alt Let_def 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1193
                            split: ls_splits)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1194
  apply clarify 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1195
  apply (erule allE)+
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1196
  apply (erule impE)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1197
   apply (fast elim!: sym)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1198
  apply (simp (no_asm_use) add: rsplit_aux_alts)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1199
  apply (erule impE)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1200
  apply (rule_tac x="ba # bs" in exI)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1201
  apply auto
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1202
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1203
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1204
lemma bin_rsplit_len_indep: 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1205
  "0 < n ==> length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1206
  apply (unfold bin_rsplit_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1207
  apply (erule bin_rsplit_aux_len_indep)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1208
  apply (rule refl)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1209
  done
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1210
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1211
end
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
  1212