| author | wenzelm | 
| Wed, 14 Oct 2015 17:24:21 +0200 | |
| changeset 61441 | 20ff1d5c74e1 | 
| parent 61424 | c3658c18b7bc | 
| child 61799 | 4cf66f21b764 | 
| permissions | -rw-r--r-- | 
| 24333 | 1 | (* | 
| 2 | Author: Jeremy Dawson, NICTA | |
| 3 | ||
| 44939 
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
 kleing parents: 
41842diff
changeset | 4 | Theorems to do with integers, expressed using Pls, Min, BIT, | 
| 24333 | 5 | theorems linking them to lists of booleans, and repeated splitting | 
| 6 | and concatenation. | |
| 7 | *) | |
| 8 | ||
| 58874 | 9 | section "Bool lists and integers" | 
| 24333 | 10 | |
| 37658 | 11 | theory Bool_List_Representation | 
| 54874 | 12 | imports Main Bits_Int | 
| 26557 | 13 | begin | 
| 24333 | 14 | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 15 | definition map2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list"
 | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 16 | where | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
58874diff
changeset | 17 | "map2 f as bs = map (case_prod f) (zip as bs)" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 18 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 19 | lemma map2_Nil [simp, code]: | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 20 | "map2 f [] ys = []" | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 21 | unfolding map2_def by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 22 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 23 | lemma map2_Nil2 [simp, code]: | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 24 | "map2 f xs [] = []" | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 25 | unfolding map2_def by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 26 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 27 | lemma map2_Cons [simp, code]: | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 28 | "map2 f (x # xs) (y # ys) = f x y # map2 f xs ys" | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 29 | unfolding map2_def by auto | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 30 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 31 | |
| 37657 | 32 | subsection {* Operations on lists of booleans *}
 | 
| 33 | ||
| 54848 | 34 | primrec bl_to_bin_aux :: "bool list \<Rightarrow> int \<Rightarrow> int" | 
| 35 | where | |
| 37657 | 36 | Nil: "bl_to_bin_aux [] w = w" | 
| 37 | | Cons: "bl_to_bin_aux (b # bs) w = | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
53438diff
changeset | 38 | bl_to_bin_aux bs (w BIT b)" | 
| 37657 | 39 | |
| 54848 | 40 | definition bl_to_bin :: "bool list \<Rightarrow> int" | 
| 41 | where | |
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 42 | bl_to_bin_def: "bl_to_bin bs = bl_to_bin_aux bs 0" | 
| 37667 | 43 | |
| 54848 | 44 | primrec bin_to_bl_aux :: "nat \<Rightarrow> int \<Rightarrow> bool list \<Rightarrow> bool list" | 
| 45 | where | |
| 37657 | 46 | Z: "bin_to_bl_aux 0 w bl = bl" | 
| 47 | | Suc: "bin_to_bl_aux (Suc n) w bl = | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
53438diff
changeset | 48 | bin_to_bl_aux n (bin_rest w) ((bin_last w) # bl)" | 
| 37657 | 49 | |
| 54848 | 50 | definition bin_to_bl :: "nat \<Rightarrow> int \<Rightarrow> bool list" | 
| 51 | where | |
| 37657 | 52 | bin_to_bl_def : "bin_to_bl n w = bin_to_bl_aux n w []" | 
| 53 | ||
| 54848 | 54 | primrec bl_of_nth :: "nat \<Rightarrow> (nat \<Rightarrow> bool) \<Rightarrow> bool list" | 
| 55 | where | |
| 37657 | 56 | Suc: "bl_of_nth (Suc n) f = f n # bl_of_nth n f" | 
| 57 | | Z: "bl_of_nth 0 f = []" | |
| 58 | ||
| 54848 | 59 | primrec takefill :: "'a \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list" | 
| 60 | where | |
| 37657 | 61 | Z: "takefill fill 0 xs = []" | 
| 62 | | Suc: "takefill fill (Suc n) xs = ( | |
| 63 | case xs of [] => fill # takefill fill n xs | |
| 64 | | y # ys => y # takefill fill n ys)" | |
| 65 | ||
| 66 | ||
| 24465 | 67 | subsection "Arithmetic in terms of bool lists" | 
| 68 | ||
| 44939 
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
 kleing parents: 
41842diff
changeset | 69 | text {* 
 | 
| 
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
 kleing parents: 
41842diff
changeset | 70 | Arithmetic operations in terms of the reversed bool list, | 
| 
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
 kleing parents: 
41842diff
changeset | 71 | assuming input list(s) the same length, and don't extend them. | 
| 
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
 kleing parents: 
41842diff
changeset | 72 | *} | 
| 24465 | 73 | |
| 54848 | 74 | primrec rbl_succ :: "bool list => bool list" | 
| 75 | where | |
| 24465 | 76 | Nil: "rbl_succ Nil = Nil" | 
| 26557 | 77 | | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)" | 
| 24465 | 78 | |
| 54848 | 79 | primrec rbl_pred :: "bool list => bool list" | 
| 80 | where | |
| 26557 | 81 | Nil: "rbl_pred Nil = Nil" | 
| 82 | | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)" | |
| 24465 | 83 | |
| 54848 | 84 | primrec rbl_add :: "bool list => bool list => bool list" | 
| 85 | where | |
| 44939 
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
 kleing parents: 
41842diff
changeset | 86 | -- "result is length of first arg, second arg may be longer" | 
| 26557 | 87 | Nil: "rbl_add Nil x = Nil" | 
| 88 | | Cons: "rbl_add (y # ys) x = (let ws = rbl_add ys (tl x) in | |
| 24465 | 89 | (y ~= hd x) # (if hd x & y then rbl_succ ws else ws))" | 
| 90 | ||
| 54848 | 91 | primrec rbl_mult :: "bool list => bool list => bool list" | 
| 92 | where | |
| 44939 
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
 kleing parents: 
41842diff
changeset | 93 | -- "result is length of first arg, second arg may be longer" | 
| 26557 | 94 | Nil: "rbl_mult Nil x = Nil" | 
| 95 | | Cons: "rbl_mult (y # ys) x = (let ws = False # rbl_mult ys x in | |
| 24465 | 96 | if y then rbl_add ws x else ws)" | 
| 24333 | 97 | |
| 98 | lemma butlast_power: | |
| 30971 | 99 | "(butlast ^^ n) bl = take (length bl - n) bl" | 
| 24333 | 100 | by (induct n) (auto simp: butlast_take) | 
| 101 | ||
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 102 | lemma bin_to_bl_aux_zero_minus_simp [simp]: | 
| 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 103 | "0 < n \<Longrightarrow> bin_to_bl_aux n 0 bl = | 
| 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 104 | bin_to_bl_aux (n - 1) 0 (False # bl)" | 
| 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 105 | by (cases n) auto | 
| 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 106 | |
| 46617 
8c5d10d41391
make bool list functions respect int/bin distinction
 huffman parents: 
46240diff
changeset | 107 | lemma bin_to_bl_aux_minus1_minus_simp [simp]: | 
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
57514diff
changeset | 108 | "0 < n ==> bin_to_bl_aux n (- 1) bl = | 
| 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
57514diff
changeset | 109 | bin_to_bl_aux (n - 1) (- 1) (True # bl)" | 
| 24333 | 110 | by (cases n) auto | 
| 111 | ||
| 46617 
8c5d10d41391
make bool list functions respect int/bin distinction
 huffman parents: 
46240diff
changeset | 112 | lemma bin_to_bl_aux_one_minus_simp [simp]: | 
| 
8c5d10d41391
make bool list functions respect int/bin distinction
 huffman parents: 
46240diff
changeset | 113 | "0 < n \<Longrightarrow> bin_to_bl_aux n 1 bl = | 
| 
8c5d10d41391
make bool list functions respect int/bin distinction
 huffman parents: 
46240diff
changeset | 114 | bin_to_bl_aux (n - 1) 0 (True # bl)" | 
| 24333 | 115 | by (cases n) auto | 
| 116 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 117 | lemma bin_to_bl_aux_Bit_minus_simp [simp]: | 
| 24333 | 118 | "0 < n ==> bin_to_bl_aux n (w BIT b) bl = | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
53438diff
changeset | 119 | bin_to_bl_aux (n - 1) w (b # bl)" | 
| 24333 | 120 | by (cases n) auto | 
| 121 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 122 | lemma bin_to_bl_aux_Bit0_minus_simp [simp]: | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46655diff
changeset | 123 | "0 < n ==> bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46655diff
changeset | 124 | bin_to_bl_aux (n - 1) (numeral w) (False # bl)" | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 125 | by (cases n) auto | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 126 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 127 | lemma bin_to_bl_aux_Bit1_minus_simp [simp]: | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46655diff
changeset | 128 | "0 < n ==> bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46655diff
changeset | 129 | bin_to_bl_aux (n - 1) (numeral w) (True # bl)" | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 130 | by (cases n) auto | 
| 24333 | 131 | |
| 44939 
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
 kleing parents: 
41842diff
changeset | 132 | text {* Link between bin and bool list. *}
 | 
| 24465 | 133 | |
| 26557 | 134 | lemma bl_to_bin_aux_append: | 
| 135 | "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)" | |
| 136 | by (induct bs arbitrary: w) auto | |
| 24465 | 137 | |
| 26557 | 138 | lemma bin_to_bl_aux_append: | 
| 139 | "bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)" | |
| 140 | by (induct n arbitrary: w bs) auto | |
| 24333 | 141 | |
| 24465 | 142 | lemma bl_to_bin_append: | 
| 26557 | 143 | "bl_to_bin (bs @ cs) = bl_to_bin_aux cs (bl_to_bin bs)" | 
| 24465 | 144 | unfolding bl_to_bin_def by (rule bl_to_bin_aux_append) | 
| 145 | ||
| 24333 | 146 | lemma bin_to_bl_aux_alt: | 
| 147 | "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" | |
| 148 | unfolding bin_to_bl_def by (simp add : bin_to_bl_aux_append) | |
| 149 | ||
| 45996 
e69d631fe9af
declare simp rules immediately, instead of using 'declare' commands
 huffman parents: 
45854diff
changeset | 150 | lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []" | 
| 24333 | 151 | unfolding bin_to_bl_def by auto | 
| 152 | ||
| 26557 | 153 | lemma size_bin_to_bl_aux: | 
| 154 | "size (bin_to_bl_aux n w bs) = n + length bs" | |
| 155 | by (induct n arbitrary: w bs) auto | |
| 24333 | 156 | |
| 45996 
e69d631fe9af
declare simp rules immediately, instead of using 'declare' commands
 huffman parents: 
45854diff
changeset | 157 | lemma size_bin_to_bl [simp]: "size (bin_to_bl n w) = n" | 
| 24333 | 158 | unfolding bin_to_bl_def by (simp add : size_bin_to_bl_aux) | 
| 159 | ||
| 26557 | 160 | lemma bin_bl_bin': | 
| 161 | "bl_to_bin (bin_to_bl_aux n w bs) = | |
| 162 | bl_to_bin_aux bs (bintrunc n w)" | |
| 163 | by (induct n arbitrary: w bs) (auto simp add : bl_to_bin_def) | |
| 24465 | 164 | |
| 45996 
e69d631fe9af
declare simp rules immediately, instead of using 'declare' commands
 huffman parents: 
45854diff
changeset | 165 | lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = bintrunc n w" | 
| 24465 | 166 | unfolding bin_to_bl_def bin_bl_bin' by auto | 
| 167 | ||
| 26557 | 168 | lemma bl_bin_bl': | 
| 169 | "bin_to_bl (n + length bs) (bl_to_bin_aux bs w) = | |
| 24465 | 170 | bin_to_bl_aux n w bs" | 
| 26557 | 171 | apply (induct bs arbitrary: w n) | 
| 24465 | 172 | apply auto | 
| 173 | apply (simp_all only : add_Suc [symmetric]) | |
| 174 | apply (auto simp add : bin_to_bl_def) | |
| 175 | done | |
| 176 | ||
| 45996 
e69d631fe9af
declare simp rules immediately, instead of using 'declare' commands
 huffman parents: 
45854diff
changeset | 177 | lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs" | 
| 24465 | 178 | unfolding bl_to_bin_def | 
| 179 | apply (rule box_equals) | |
| 180 | apply (rule bl_bin_bl') | |
| 181 | prefer 2 | |
| 182 | apply (rule bin_to_bl_aux.Z) | |
| 183 | apply simp | |
| 184 | done | |
| 185 | ||
| 186 | lemma bl_to_bin_inj: | |
| 187 | "bl_to_bin bs = bl_to_bin cs ==> length bs = length cs ==> bs = cs" | |
| 188 | apply (rule_tac box_equals) | |
| 189 | defer | |
| 190 | apply (rule bl_bin_bl) | |
| 191 | apply (rule bl_bin_bl) | |
| 192 | apply simp | |
| 193 | done | |
| 194 | ||
| 45996 
e69d631fe9af
declare simp rules immediately, instead of using 'declare' commands
 huffman parents: 
45854diff
changeset | 195 | lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl" | 
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 196 | unfolding bl_to_bin_def by auto | 
| 45996 
e69d631fe9af
declare simp rules immediately, instead of using 'declare' commands
 huffman parents: 
45854diff
changeset | 197 | |
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 198 | lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = 0" | 
| 24465 | 199 | unfolding bl_to_bin_def by auto | 
| 200 | ||
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 201 | lemma bin_to_bl_zero_aux: | 
| 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 202 | "bin_to_bl_aux n 0 bl = replicate n False @ bl" | 
| 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 203 | by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) | 
| 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 204 | |
| 46617 
8c5d10d41391
make bool list functions respect int/bin distinction
 huffman parents: 
46240diff
changeset | 205 | lemma bin_to_bl_zero: "bin_to_bl n 0 = replicate n False" | 
| 
8c5d10d41391
make bool list functions respect int/bin distinction
 huffman parents: 
46240diff
changeset | 206 | unfolding bin_to_bl_def by (simp add: bin_to_bl_zero_aux) | 
| 
8c5d10d41391
make bool list functions respect int/bin distinction
 huffman parents: 
46240diff
changeset | 207 | |
| 
8c5d10d41391
make bool list functions respect int/bin distinction
 huffman parents: 
46240diff
changeset | 208 | lemma bin_to_bl_minus1_aux: | 
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
57514diff
changeset | 209 | "bin_to_bl_aux n (- 1) bl = replicate n True @ bl" | 
| 26557 | 210 | by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) | 
| 24333 | 211 | |
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
57514diff
changeset | 212 | lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True" | 
| 46617 
8c5d10d41391
make bool list functions respect int/bin distinction
 huffman parents: 
46240diff
changeset | 213 | unfolding bin_to_bl_def by (simp add: bin_to_bl_minus1_aux) | 
| 24333 | 214 | |
| 24465 | 215 | lemma bl_to_bin_rep_F: | 
| 216 | "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" | |
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 217 | apply (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') | 
| 24465 | 218 | apply (simp add: bl_to_bin_def) | 
| 219 | done | |
| 220 | ||
| 45996 
e69d631fe9af
declare simp rules immediately, instead of using 'declare' commands
 huffman parents: 
45854diff
changeset | 221 | lemma bin_to_bl_trunc [simp]: | 
| 24465 | 222 | "n <= m ==> bin_to_bl n (bintrunc m w) = bin_to_bl n w" | 
| 223 | by (auto intro: bl_to_bin_inj) | |
| 224 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 225 | lemma bin_to_bl_aux_bintr: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 226 | "bin_to_bl_aux n (bintrunc m bin) bl = | 
| 24333 | 227 | replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 228 | apply (induct n arbitrary: m bin bl) | 
| 24333 | 229 | apply clarsimp | 
| 230 | apply clarsimp | |
| 231 | apply (case_tac "m") | |
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 232 | apply (clarsimp simp: bin_to_bl_zero_aux) | 
| 24333 | 233 | apply (erule thin_rl) | 
| 234 | apply (induct_tac n) | |
| 235 | apply auto | |
| 236 | done | |
| 237 | ||
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 238 | lemma bin_to_bl_bintr: | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 239 | "bin_to_bl n (bintrunc m bin) = | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 240 | replicate (n - m) False @ bin_to_bl (min n m) bin" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 241 | unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr) | 
| 24333 | 242 | |
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 243 | lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0" | 
| 24465 | 244 | by (induct n) auto | 
| 245 | ||
| 26557 | 246 | lemma len_bin_to_bl_aux: | 
| 247 | "length (bin_to_bl_aux n w bs) = n + length bs" | |
| 54869 | 248 | by (fact size_bin_to_bl_aux) | 
| 24333 | 249 | |
| 250 | lemma len_bin_to_bl [simp]: "length (bin_to_bl n w) = n" | |
| 46617 
8c5d10d41391
make bool list functions respect int/bin distinction
 huffman parents: 
46240diff
changeset | 251 | by (fact size_bin_to_bl) (* FIXME: duplicate *) | 
| 24333 | 252 | |
| 26557 | 253 | lemma sign_bl_bin': | 
| 254 | "bin_sign (bl_to_bin_aux bs w) = bin_sign w" | |
| 255 | by (induct bs arbitrary: w) auto | |
| 24333 | 256 | |
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 257 | lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0" | 
| 24333 | 258 | unfolding bl_to_bin_def by (simp add : sign_bl_bin') | 
| 259 | ||
| 26557 | 260 | lemma bl_sbin_sign_aux: | 
| 261 | "hd (bin_to_bl_aux (Suc n) w bs) = | |
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 262 | (bin_sign (sbintrunc n w) = -1)" | 
| 26557 | 263 | apply (induct n arbitrary: w bs) | 
| 24333 | 264 | apply clarsimp | 
| 26557 | 265 | apply (cases w rule: bin_exhaust) | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
53438diff
changeset | 266 | apply simp | 
| 24333 | 267 | done | 
| 268 | ||
| 269 | lemma bl_sbin_sign: | |
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 270 | "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = -1)" | 
| 24333 | 271 | unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) | 
| 272 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 273 | lemma bin_nth_of_bl_aux: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 274 | "bin_nth (bl_to_bin_aux bl w) n = | 
| 24333 | 275 | (n < size bl & rev bl ! n | n >= length bl & bin_nth w (n - size bl))" | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 276 | apply (induct bl arbitrary: w) | 
| 24333 | 277 | apply clarsimp | 
| 278 | apply clarsimp | |
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 279 | apply (cut_tac x=n and y="size bl" in linorder_less_linear) | 
| 24333 | 280 | apply (erule disjE, simp add: nth_append)+ | 
| 26557 | 281 | apply auto | 
| 24333 | 282 | done | 
| 283 | ||
| 45475 | 284 | lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl & rev bl ! n)" | 
| 24333 | 285 | unfolding bl_to_bin_def by (simp add : bin_nth_of_bl_aux) | 
| 286 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 287 | lemma bin_nth_bl: "n < m \<Longrightarrow> bin_nth w n = nth (rev (bin_to_bl m w)) n" | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 288 | apply (induct n arbitrary: m w) | 
| 24333 | 289 | apply clarsimp | 
| 290 | apply (case_tac m, clarsimp) | |
| 291 | apply (clarsimp simp: bin_to_bl_def) | |
| 292 | apply (simp add: bin_to_bl_aux_alt) | |
| 293 | apply clarsimp | |
| 294 | apply (case_tac m, clarsimp) | |
| 295 | apply (clarsimp simp: bin_to_bl_def) | |
| 296 | apply (simp add: bin_to_bl_aux_alt) | |
| 297 | done | |
| 298 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 299 | lemma nth_rev: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 300 | "n < length xs \<Longrightarrow> rev xs ! n = xs ! (length xs - 1 - n)" | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 301 | apply (induct xs) | 
| 24465 | 302 | apply simp | 
| 303 | apply (clarsimp simp add : nth_append nth.simps split add : nat.split) | |
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 304 | apply (rule_tac f = "\<lambda>n. xs ! n" in arg_cong) | 
| 24465 | 305 | apply arith | 
| 306 | done | |
| 307 | ||
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 308 | lemma nth_rev_alt: "n < length ys \<Longrightarrow> ys ! n = rev ys ! (length ys - Suc n)" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 309 | by (simp add: nth_rev) | 
| 24465 | 310 | |
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 311 | lemma nth_bin_to_bl_aux: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 312 | "n < m + length bl \<Longrightarrow> (bin_to_bl_aux m w bl) ! n = | 
| 24333 | 313 | (if n < m then bin_nth w (m - 1 - n) else bl ! (n - m))" | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 314 | apply (induct m arbitrary: w n bl) | 
| 24333 | 315 | apply clarsimp | 
| 316 | apply clarsimp | |
| 317 | apply (case_tac w rule: bin_exhaust) | |
| 318 | apply simp | |
| 319 | done | |
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 320 | |
| 24333 | 321 | lemma nth_bin_to_bl: "n < m ==> (bin_to_bl m w) ! n = bin_nth w (m - Suc n)" | 
| 322 | unfolding bin_to_bl_def by (simp add : nth_bin_to_bl_aux) | |
| 323 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 324 | lemma bl_to_bin_lt2p_aux: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 325 | "bl_to_bin_aux bs w < (w + 1) * (2 ^ length bs)" | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 326 | apply (induct bs arbitrary: w) | 
| 24333 | 327 | apply clarsimp | 
| 328 | apply clarsimp | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 329 | apply (drule meta_spec, erule xtrans(8) [rotated], simp add: Bit_def)+ | 
| 24333 | 330 | done | 
| 331 | ||
| 332 | lemma bl_to_bin_lt2p: "bl_to_bin bs < (2 ^ length bs)" | |
| 333 | apply (unfold bl_to_bin_def) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 334 | apply (rule xtrans(1)) | 
| 24333 | 335 | prefer 2 | 
| 336 | apply (rule bl_to_bin_lt2p_aux) | |
| 337 | apply simp | |
| 338 | done | |
| 339 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 340 | lemma bl_to_bin_ge2p_aux: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 341 | "bl_to_bin_aux bs w >= w * (2 ^ length bs)" | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 342 | apply (induct bs arbitrary: w) | 
| 24333 | 343 | apply clarsimp | 
| 344 | apply clarsimp | |
| 46652 | 345 | apply (drule meta_spec, erule order_trans [rotated], | 
| 346 | simp add: Bit_B0_2t Bit_B1_2t algebra_simps)+ | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
53438diff
changeset | 347 | apply (simp add: Bit_def) | 
| 24333 | 348 | done | 
| 349 | ||
| 350 | lemma bl_to_bin_ge0: "bl_to_bin bs >= 0" | |
| 351 | apply (unfold bl_to_bin_def) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 352 | apply (rule xtrans(4)) | 
| 24333 | 353 | apply (rule bl_to_bin_ge2p_aux) | 
| 46617 
8c5d10d41391
make bool list functions respect int/bin distinction
 huffman parents: 
46240diff
changeset | 354 | apply simp | 
| 24333 | 355 | done | 
| 356 | ||
| 357 | lemma butlast_rest_bin: | |
| 358 | "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)" | |
| 359 | apply (unfold bin_to_bl_def) | |
| 360 | apply (cases w rule: bin_exhaust) | |
| 361 | apply (cases n, clarsimp) | |
| 362 | apply clarsimp | |
| 363 | apply (auto simp add: bin_to_bl_aux_alt) | |
| 364 | done | |
| 365 | ||
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 366 | lemma butlast_bin_rest: | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 367 | "butlast bl = bin_to_bl (length bl - Suc 0) (bin_rest (bl_to_bin bl))" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 368 | using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp | 
| 24333 | 369 | |
| 26557 | 370 | lemma butlast_rest_bl2bin_aux: | 
| 371 | "bl ~= [] \<Longrightarrow> | |
| 372 | bl_to_bin_aux (butlast bl) w = bin_rest (bl_to_bin_aux bl w)" | |
| 373 | by (induct bl arbitrary: w) auto | |
| 24333 | 374 | |
| 375 | lemma butlast_rest_bl2bin: | |
| 376 | "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)" | |
| 377 | apply (unfold bl_to_bin_def) | |
| 378 | apply (cases bl) | |
| 379 | apply (auto simp add: butlast_rest_bl2bin_aux) | |
| 380 | done | |
| 381 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 382 | lemma trunc_bl2bin_aux: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 383 | "bintrunc m (bl_to_bin_aux bl w) = | 
| 26557 | 384 | bl_to_bin_aux (drop (length bl - m) bl) (bintrunc (m - length bl) w)" | 
| 53438 | 385 | proof (induct bl arbitrary: w) | 
| 386 | case Nil show ?case by simp | |
| 387 | next | |
| 388 | case (Cons b bl) show ?case | |
| 389 | proof (cases "m - length bl") | |
| 390 | case 0 then have "Suc (length bl) - m = Suc (length bl - m)" by simp | |
| 391 | with Cons show ?thesis by simp | |
| 392 | next | |
| 393 | case (Suc n) then have *: "m - Suc (length bl) = n" by simp | |
| 394 | with Suc Cons show ?thesis by simp | |
| 395 | qed | |
| 396 | qed | |
| 24333 | 397 | |
| 398 | lemma trunc_bl2bin: | |
| 399 | "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" | |
| 400 | unfolding bl_to_bin_def by (simp add : trunc_bl2bin_aux) | |
| 401 | ||
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 402 | lemma trunc_bl2bin_len [simp]: | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 403 | "bintrunc (length bl) (bl_to_bin bl) = bl_to_bin bl" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 404 | by (simp add: trunc_bl2bin) | 
| 24333 | 405 | |
| 406 | lemma bl2bin_drop: | |
| 407 | "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)" | |
| 408 | apply (rule trans) | |
| 409 | prefer 2 | |
| 410 | apply (rule trunc_bl2bin [symmetric]) | |
| 411 | apply (cases "k <= length bl") | |
| 412 | apply auto | |
| 413 | done | |
| 414 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 415 | lemma nth_rest_power_bin: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 416 | "bin_nth ((bin_rest ^^ k) w) n = bin_nth w (n + k)" | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 417 | apply (induct k arbitrary: n, clarsimp) | 
| 24333 | 418 | apply clarsimp | 
| 419 | apply (simp only: bin_nth.Suc [symmetric] add_Suc) | |
| 420 | done | |
| 421 | ||
| 422 | lemma take_rest_power_bin: | |
| 30971 | 423 | "m <= n ==> take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^^ (n - m)) w)" | 
| 24333 | 424 | apply (rule nth_equalityI) | 
| 425 | apply simp | |
| 426 | apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) | |
| 427 | done | |
| 428 | ||
| 24465 | 429 | lemma hd_butlast: "size xs > 1 ==> hd (butlast xs) = hd xs" | 
| 430 | by (cases xs) auto | |
| 24333 | 431 | |
| 26557 | 432 | lemma last_bin_last': | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
53438diff
changeset | 433 | "size xs > 0 \<Longrightarrow> last xs \<longleftrightarrow> bin_last (bl_to_bin_aux xs w)" | 
| 26557 | 434 | by (induct xs arbitrary: w) auto | 
| 24333 | 435 | |
| 436 | lemma last_bin_last: | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
53438diff
changeset | 437 | "size xs > 0 ==> last xs \<longleftrightarrow> bin_last (bl_to_bin xs)" | 
| 24333 | 438 | unfolding bl_to_bin_def by (erule last_bin_last') | 
| 439 | ||
| 440 | lemma bin_last_last: | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
53438diff
changeset | 441 | "bin_last w \<longleftrightarrow> last (bin_to_bl (Suc n) w)" | 
| 24333 | 442 | apply (unfold bin_to_bl_def) | 
| 443 | apply simp | |
| 444 | apply (auto simp add: bin_to_bl_aux_alt) | |
| 445 | done | |
| 446 | ||
| 24465 | 447 | (** links between bit-wise operations and operations on bool lists **) | 
| 448 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 449 | lemma bl_xor_aux_bin: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 450 | "map2 (%x y. x ~= y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = | 
| 26557 | 451 | bin_to_bl_aux n (v XOR w) (map2 (%x y. x ~= y) bs cs)" | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 452 | apply (induct n arbitrary: v w bs cs) | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 453 | apply simp | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 454 | apply (case_tac v rule: bin_exhaust) | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 455 | apply (case_tac w rule: bin_exhaust) | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 456 | apply clarsimp | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 457 | apply (case_tac b) | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
53438diff
changeset | 458 | apply auto | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 459 | done | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 460 | |
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 461 | lemma bl_or_aux_bin: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 462 | "map2 (op | ) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 463 | bin_to_bl_aux n (v OR w) (map2 (op | ) bs cs)" | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 464 | apply (induct n arbitrary: v w bs cs) | 
| 24333 | 465 | apply simp | 
| 466 | apply (case_tac v rule: bin_exhaust) | |
| 467 | apply (case_tac w rule: bin_exhaust) | |
| 468 | apply clarsimp | |
| 469 | done | |
| 470 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 471 | lemma bl_and_aux_bin: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 472 | "map2 (op & ) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = | 
| 26557 | 473 | bin_to_bl_aux n (v AND w) (map2 (op & ) bs cs)" | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 474 | apply (induct n arbitrary: v w bs cs) | 
| 24333 | 475 | apply simp | 
| 476 | apply (case_tac v rule: bin_exhaust) | |
| 477 | apply (case_tac w rule: bin_exhaust) | |
| 478 | apply clarsimp | |
| 479 | done | |
| 480 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 481 | lemma bl_not_aux_bin: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 482 | "map Not (bin_to_bl_aux n w cs) = | 
| 24353 | 483 | bin_to_bl_aux n (NOT w) (map Not cs)" | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 484 | apply (induct n arbitrary: w cs) | 
| 24333 | 485 | apply clarsimp | 
| 486 | apply clarsimp | |
| 487 | done | |
| 488 | ||
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 489 | lemma bl_not_bin: "map Not (bin_to_bl n w) = bin_to_bl n (NOT w)" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 490 | unfolding bin_to_bl_def by (simp add: bl_not_aux_bin) | 
| 24333 | 491 | |
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 492 | lemma bl_and_bin: | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 493 | "map2 (op \<and>) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v AND w)" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 494 | unfolding bin_to_bl_def by (simp add: bl_and_aux_bin) | 
| 24333 | 495 | |
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 496 | lemma bl_or_bin: | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 497 | "map2 (op \<or>) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v OR w)" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 498 | unfolding bin_to_bl_def by (simp add: bl_or_aux_bin) | 
| 24333 | 499 | |
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 500 | lemma bl_xor_bin: | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 501 | "map2 (\<lambda>x y. x \<noteq> y) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v XOR w)" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 502 | unfolding bin_to_bl_def by (simp only: bl_xor_aux_bin map2_Nil) | 
| 24333 | 503 | |
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 504 | lemma drop_bin2bl_aux: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 505 | "drop m (bin_to_bl_aux n bin bs) = | 
| 24333 | 506 | bin_to_bl_aux (n - m) bin (drop (m - n) bs)" | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 507 | apply (induct n arbitrary: m bin bs, clarsimp) | 
| 24333 | 508 | apply clarsimp | 
| 509 | apply (case_tac bin rule: bin_exhaust) | |
| 510 | apply (case_tac "m <= n", simp) | |
| 511 | apply (case_tac "m - n", simp) | |
| 512 | apply simp | |
| 513 | apply (rule_tac f = "%nat. drop nat bs" in arg_cong) | |
| 514 | apply simp | |
| 515 | done | |
| 516 | ||
| 517 | lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin" | |
| 518 | unfolding bin_to_bl_def by (simp add : drop_bin2bl_aux) | |
| 519 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 520 | lemma take_bin2bl_lem1: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 521 | "take m (bin_to_bl_aux m w bs) = bin_to_bl m w" | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 522 | apply (induct m arbitrary: w bs, clarsimp) | 
| 24333 | 523 | apply clarsimp | 
| 524 | apply (simp add: bin_to_bl_aux_alt) | |
| 525 | apply (simp add: bin_to_bl_def) | |
| 526 | apply (simp add: bin_to_bl_aux_alt) | |
| 527 | done | |
| 528 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 529 | lemma take_bin2bl_lem: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 530 | "take m (bin_to_bl_aux (m + n) w bs) = | 
| 24333 | 531 | take m (bin_to_bl (m + n) w)" | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 532 | apply (induct n arbitrary: w bs) | 
| 24333 | 533 | apply (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1) | 
| 534 | apply simp | |
| 535 | done | |
| 536 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 537 | lemma bin_split_take: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 538 | "bin_split n c = (a, b) \<Longrightarrow> | 
| 24333 | 539 | bin_to_bl m a = take m (bin_to_bl (m + n) c)" | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 540 | apply (induct n arbitrary: b c) | 
| 24333 | 541 | apply clarsimp | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 542 | apply (clarsimp simp: Let_def split: prod.split_asm) | 
| 24333 | 543 | apply (simp add: bin_to_bl_def) | 
| 544 | apply (simp add: take_bin2bl_lem) | |
| 545 | done | |
| 546 | ||
| 547 | lemma bin_split_take1: | |
| 548 | "k = m + n ==> bin_split n c = (a, b) ==> | |
| 549 | bin_to_bl m a = take m (bin_to_bl k c)" | |
| 550 | by (auto elim: bin_split_take) | |
| 551 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 552 | lemma nth_takefill: "m < n \<Longrightarrow> | 
| 24333 | 553 | takefill fill n l ! m = (if m < length l then l ! m else fill)" | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 554 | apply (induct n arbitrary: m l, clarsimp) | 
| 24333 | 555 | apply clarsimp | 
| 556 | apply (case_tac m) | |
| 557 | apply (simp split: list.split) | |
| 558 | apply (simp split: list.split) | |
| 559 | done | |
| 560 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 561 | lemma takefill_alt: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 562 | "takefill fill n l = take n l @ replicate (n - length l) fill" | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 563 | by (induct n arbitrary: l) (auto split: list.split) | 
| 24333 | 564 | |
| 565 | lemma takefill_replicate [simp]: | |
| 566 | "takefill fill n (replicate m fill) = replicate n fill" | |
| 567 | by (simp add : takefill_alt replicate_add [symmetric]) | |
| 568 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 569 | lemma takefill_le': | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 570 | "n = m + k \<Longrightarrow> takefill x m (takefill x n l) = takefill x m l" | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 571 | by (induct m arbitrary: l n) (auto split: list.split) | 
| 24333 | 572 | |
| 573 | lemma length_takefill [simp]: "length (takefill fill n l) = n" | |
| 574 | by (simp add : takefill_alt) | |
| 575 | ||
| 576 | lemma take_takefill': | |
| 577 | "!!w n. n = k + m ==> take k (takefill fill n w) = takefill fill k w" | |
| 578 | by (induct k) (auto split add : list.split) | |
| 579 | ||
| 580 | lemma drop_takefill: | |
| 581 | "!!w. drop k (takefill fill (m + k) w) = takefill fill m (drop k w)" | |
| 582 | by (induct k) (auto split add : list.split) | |
| 583 | ||
| 584 | lemma takefill_le [simp]: | |
| 585 | "m \<le> n \<Longrightarrow> takefill x m (takefill x n l) = takefill x m l" | |
| 586 | by (auto simp: le_iff_add takefill_le') | |
| 587 | ||
| 588 | lemma take_takefill [simp]: | |
| 589 | "m \<le> n \<Longrightarrow> take m (takefill fill n w) = takefill fill m w" | |
| 590 | by (auto simp: le_iff_add take_takefill') | |
| 591 | ||
| 592 | lemma takefill_append: | |
| 593 | "takefill fill (m + length xs) (xs @ w) = xs @ (takefill fill m w)" | |
| 594 | by (induct xs) auto | |
| 595 | ||
| 596 | lemma takefill_same': | |
| 597 | "l = length xs ==> takefill fill l xs = xs" | |
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
56073diff
changeset | 598 | by (induct xs arbitrary: l, auto) | 
| 24333 | 599 | |
| 600 | lemmas takefill_same [simp] = takefill_same' [OF refl] | |
| 601 | ||
| 602 | lemma takefill_bintrunc: | |
| 603 | "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" | |
| 604 | apply (rule nth_equalityI) | |
| 605 | apply simp | |
| 606 | apply (clarsimp simp: nth_takefill nth_rev nth_bin_to_bl bin_nth_of_bl) | |
| 607 | done | |
| 608 | ||
| 609 | lemma bl_bin_bl_rtf: | |
| 610 | "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))" | |
| 611 | by (simp add : takefill_bintrunc) | |
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 612 | |
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 613 | lemma bl_bin_bl_rep_drop: | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 614 | "bin_to_bl n (bl_to_bin bl) = | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 615 | replicate (n - length bl) False @ drop (length bl - n) bl" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 616 | by (simp add: bl_bin_bl_rtf takefill_alt rev_take) | 
| 24333 | 617 | |
| 618 | lemma tf_rev: | |
| 619 | "n + k = m + length bl ==> takefill x m (rev (takefill y n bl)) = | |
| 620 | rev (takefill y m (rev (takefill x k (rev bl))))" | |
| 621 | apply (rule nth_equalityI) | |
| 622 | apply (auto simp add: nth_takefill nth_rev) | |
| 623 | apply (rule_tac f = "%n. bl ! n" in arg_cong) | |
| 624 | apply arith | |
| 625 | done | |
| 626 | ||
| 627 | lemma takefill_minus: | |
| 628 | "0 < n ==> takefill fill (Suc (n - 1)) w = takefill fill n w" | |
| 629 | by auto | |
| 630 | ||
| 631 | lemmas takefill_Suc_cases = | |
| 45604 | 632 | list.cases [THEN takefill.Suc [THEN trans]] | 
| 24333 | 633 | |
| 634 | lemmas takefill_Suc_Nil = takefill_Suc_cases (1) | |
| 635 | lemmas takefill_Suc_Cons = takefill_Suc_cases (2) | |
| 636 | ||
| 637 | lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2] | |
| 45604 | 638 | takefill_minus [symmetric, THEN trans]] | 
| 24333 | 639 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46655diff
changeset | 640 | lemma takefill_numeral_Nil [simp]: | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47108diff
changeset | 641 | "takefill fill (numeral k) [] = fill # takefill fill (pred_numeral k) []" | 
| 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47108diff
changeset | 642 | by (simp add: numeral_eq_Suc) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46655diff
changeset | 643 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46655diff
changeset | 644 | lemma takefill_numeral_Cons [simp]: | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47108diff
changeset | 645 | "takefill fill (numeral k) (x # xs) = x # takefill fill (pred_numeral k) xs" | 
| 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47108diff
changeset | 646 | by (simp add: numeral_eq_Suc) | 
| 24333 | 647 | |
| 648 | (* links with function bl_to_bin *) | |
| 649 | ||
| 650 | lemma bl_to_bin_aux_cat: | |
| 26557 | 651 | "!!nv v. bl_to_bin_aux bs (bin_cat w nv v) = | 
| 652 | bin_cat w (nv + length bs) (bl_to_bin_aux bs v)" | |
| 24333 | 653 | apply (induct bs) | 
| 654 | apply simp | |
| 655 | apply (simp add: bin_cat_Suc_Bit [symmetric] del: bin_cat.simps) | |
| 656 | done | |
| 657 | ||
| 658 | lemma bin_to_bl_aux_cat: | |
| 659 | "!!w bs. bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs = | |
| 660 | bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" | |
| 661 | by (induct nw) auto | |
| 662 | ||
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 663 | lemma bl_to_bin_aux_alt: | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 664 | "bl_to_bin_aux bs w = bin_cat w (length bs) (bl_to_bin bs)" | 
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
45997diff
changeset | 665 | using bl_to_bin_aux_cat [where nv = "0" and v = "0"] | 
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 666 | unfolding bl_to_bin_def [symmetric] by simp | 
| 24333 | 667 | |
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 668 | lemma bin_to_bl_cat: | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 669 | "bin_to_bl (nv + nw) (bin_cat v nw w) = | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 670 | bin_to_bl_aux nv v (bin_to_bl nw w)" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 671 | unfolding bin_to_bl_def by (simp add: bin_to_bl_aux_cat) | 
| 24333 | 672 | |
| 673 | lemmas bl_to_bin_aux_app_cat = | |
| 674 | trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt] | |
| 675 | ||
| 676 | lemmas bin_to_bl_aux_cat_app = | |
| 677 | trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] | |
| 678 | ||
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 679 | lemma bl_to_bin_app_cat: | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 680 | "bl_to_bin (bsa @ bs) = bin_cat (bl_to_bin bsa) (length bs) (bl_to_bin bs)" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 681 | by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def) | 
| 24333 | 682 | |
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 683 | lemma bin_to_bl_cat_app: | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 684 | "bin_to_bl (n + nw) (bin_cat w nw wa) = bin_to_bl n w @ bin_to_bl nw wa" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 685 | by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app) | 
| 24333 | 686 | |
| 687 | (* bl_to_bin_app_cat_alt and bl_to_bin_app_cat are easily interderivable *) | |
| 688 | lemma bl_to_bin_app_cat_alt: | |
| 689 | "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)" | |
| 690 | by (simp add : bl_to_bin_app_cat) | |
| 691 | ||
| 692 | lemma mask_lem: "(bl_to_bin (True # replicate n False)) = | |
| 46645 
573aff6b9b0a
adapt lemma mask_lem to respect int/bin distinction
 huffman parents: 
46617diff
changeset | 693 | (bl_to_bin (replicate n True)) + 1" | 
| 24333 | 694 | apply (unfold bl_to_bin_def) | 
| 695 | apply (induct n) | |
| 46645 
573aff6b9b0a
adapt lemma mask_lem to respect int/bin distinction
 huffman parents: 
46617diff
changeset | 696 | apply simp | 
| 31790 | 697 | apply (simp only: Suc_eq_plus1 replicate_add | 
| 24333 | 698 | append_Cons [symmetric] bl_to_bin_aux_append) | 
| 46645 
573aff6b9b0a
adapt lemma mask_lem to respect int/bin distinction
 huffman parents: 
46617diff
changeset | 699 | apply (simp add: Bit_B0_2t Bit_B1_2t) | 
| 24333 | 700 | done | 
| 701 | ||
| 24465 | 702 | (* function bl_of_nth *) | 
| 24333 | 703 | lemma length_bl_of_nth [simp]: "length (bl_of_nth n f) = n" | 
| 704 | by (induct n) auto | |
| 705 | ||
| 706 | lemma nth_bl_of_nth [simp]: | |
| 707 | "m < n \<Longrightarrow> rev (bl_of_nth n f) ! m = f m" | |
| 708 | apply (induct n) | |
| 709 | apply simp | |
| 710 | apply (clarsimp simp add : nth_append) | |
| 711 | apply (rule_tac f = "f" in arg_cong) | |
| 712 | apply simp | |
| 713 | done | |
| 714 | ||
| 715 | lemma bl_of_nth_inj: | |
| 716 | "(!!k. k < n ==> f k = g k) ==> bl_of_nth n f = bl_of_nth n g" | |
| 717 | by (induct n) auto | |
| 718 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 719 | lemma bl_of_nth_nth_le: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 720 | "n \<le> length xs \<Longrightarrow> bl_of_nth n (nth (rev xs)) = drop (length xs - n) xs" | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 721 | apply (induct n arbitrary: xs, clarsimp) | 
| 24333 | 722 | apply clarsimp | 
| 723 | apply (rule trans [OF _ hd_Cons_tl]) | |
| 724 | apply (frule Suc_le_lessD) | |
| 725 | apply (simp add: nth_rev trans [OF drop_Suc drop_tl, symmetric]) | |
| 726 | apply (subst hd_drop_conv_nth) | |
| 727 | apply force | |
| 728 | apply simp_all | |
| 729 | apply (rule_tac f = "%n. drop n xs" in arg_cong) | |
| 730 | apply simp | |
| 731 | done | |
| 732 | ||
| 45854 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 733 | lemma bl_of_nth_nth [simp]: "bl_of_nth (length xs) (op ! (rev xs)) = xs" | 
| 
40554613b4f0
replace many uses of 'lemmas' with explicit 'lemma'
 huffman parents: 
45847diff
changeset | 734 | by (simp add: bl_of_nth_nth_le) | 
| 24333 | 735 | |
| 736 | lemma size_rbl_pred: "length (rbl_pred bl) = length bl" | |
| 737 | by (induct bl) auto | |
| 738 | ||
| 739 | lemma size_rbl_succ: "length (rbl_succ bl) = length bl" | |
| 740 | by (induct bl) auto | |
| 741 | ||
| 742 | lemma size_rbl_add: | |
| 743 | "!!cl. length (rbl_add bl cl) = length bl" | |
| 744 | by (induct bl) (auto simp: Let_def size_rbl_succ) | |
| 745 | ||
| 746 | lemma size_rbl_mult: | |
| 747 | "!!cl. length (rbl_mult bl cl) = length bl" | |
| 748 | by (induct bl) (auto simp add : Let_def size_rbl_add) | |
| 749 | ||
| 750 | lemmas rbl_sizes [simp] = | |
| 751 | size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult | |
| 752 | ||
| 753 | lemmas rbl_Nils = | |
| 754 | rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil | |
| 755 | ||
| 46653 | 756 | lemma rbl_pred: | 
| 757 | "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))" | |
| 758 | apply (induct n arbitrary: bin, simp) | |
| 24333 | 759 | apply (unfold bin_to_bl_def) | 
| 760 | apply clarsimp | |
| 761 | apply (case_tac bin rule: bin_exhaust) | |
| 762 | apply (case_tac b) | |
| 46653 | 763 | apply (clarsimp simp: bin_to_bl_aux_alt)+ | 
| 24333 | 764 | done | 
| 765 | ||
| 766 | lemma rbl_succ: | |
| 46653 | 767 | "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" | 
| 768 | apply (induct n arbitrary: bin, simp) | |
| 24333 | 769 | apply (unfold bin_to_bl_def) | 
| 770 | apply clarsimp | |
| 771 | apply (case_tac bin rule: bin_exhaust) | |
| 772 | apply (case_tac b) | |
| 46653 | 773 | apply (clarsimp simp: bin_to_bl_aux_alt)+ | 
| 24333 | 774 | done | 
| 775 | ||
| 776 | lemma rbl_add: | |
| 777 | "!!bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = | |
| 778 | rev (bin_to_bl n (bina + binb))" | |
| 779 | apply (induct n, simp) | |
| 780 | apply (unfold bin_to_bl_def) | |
| 781 | apply clarsimp | |
| 782 | apply (case_tac bina rule: bin_exhaust) | |
| 783 | apply (case_tac binb rule: bin_exhaust) | |
| 784 | apply (case_tac b) | |
| 785 | apply (case_tac [!] "ba") | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57492diff
changeset | 786 | apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps) | 
| 24333 | 787 | done | 
| 788 | ||
| 789 | lemma rbl_add_app2: | |
| 790 | "!!blb. length blb >= length bla ==> | |
| 791 | rbl_add bla (blb @ blc) = rbl_add bla blb" | |
| 792 | apply (induct bla, simp) | |
| 793 | apply clarsimp | |
| 794 | apply (case_tac blb, clarsimp) | |
| 795 | apply (clarsimp simp: Let_def) | |
| 796 | done | |
| 797 | ||
| 798 | lemma rbl_add_take2: | |
| 799 | "!!blb. length blb >= length bla ==> | |
| 800 | rbl_add bla (take (length bla) blb) = rbl_add bla blb" | |
| 801 | apply (induct bla, simp) | |
| 802 | apply clarsimp | |
| 803 | apply (case_tac blb, clarsimp) | |
| 804 | apply (clarsimp simp: Let_def) | |
| 805 | done | |
| 806 | ||
| 807 | lemma rbl_add_long: | |
| 808 | "m >= n ==> rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = | |
| 809 | rev (bin_to_bl n (bina + binb))" | |
| 810 | apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) | |
| 811 | apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) | |
| 812 | apply (rule rev_swap [THEN iffD1]) | |
| 813 | apply (simp add: rev_take drop_bin2bl) | |
| 814 | apply simp | |
| 815 | done | |
| 816 | ||
| 817 | lemma rbl_mult_app2: | |
| 818 | "!!blb. length blb >= length bla ==> | |
| 819 | rbl_mult bla (blb @ blc) = rbl_mult bla blb" | |
| 820 | apply (induct bla, simp) | |
| 821 | apply clarsimp | |
| 822 | apply (case_tac blb, clarsimp) | |
| 823 | apply (clarsimp simp: Let_def rbl_add_app2) | |
| 824 | done | |
| 825 | ||
| 826 | lemma rbl_mult_take2: | |
| 827 | "length blb >= length bla ==> | |
| 828 | rbl_mult bla (take (length bla) blb) = rbl_mult bla blb" | |
| 829 | apply (rule trans) | |
| 830 | apply (rule rbl_mult_app2 [symmetric]) | |
| 831 | apply simp | |
| 832 | apply (rule_tac f = "rbl_mult bla" in arg_cong) | |
| 833 | apply (rule append_take_drop_id) | |
| 834 | done | |
| 835 | ||
| 836 | lemma rbl_mult_gt1: | |
| 837 | "m >= length bl ==> rbl_mult bl (rev (bin_to_bl m binb)) = | |
| 838 | rbl_mult bl (rev (bin_to_bl (length bl) binb))" | |
| 839 | apply (rule trans) | |
| 840 | apply (rule rbl_mult_take2 [symmetric]) | |
| 841 | apply simp_all | |
| 842 | apply (rule_tac f = "rbl_mult bl" in arg_cong) | |
| 843 | apply (rule rev_swap [THEN iffD1]) | |
| 844 | apply (simp add: rev_take drop_bin2bl) | |
| 845 | done | |
| 846 | ||
| 847 | lemma rbl_mult_gt: | |
| 848 | "m > n ==> rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = | |
| 849 | rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" | |
| 850 | by (auto intro: trans [OF rbl_mult_gt1]) | |
| 851 | ||
| 852 | lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] | |
| 853 | ||
| 854 | lemma rbbl_Cons: | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
53438diff
changeset | 855 | "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (x BIT b))" | 
| 24333 | 856 | apply (unfold bin_to_bl_def) | 
| 857 | apply simp | |
| 858 | apply (simp add: bin_to_bl_aux_alt) | |
| 859 | done | |
| 46653 | 860 | |
| 24333 | 861 | lemma rbl_mult: "!!bina binb. | 
| 862 | rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = | |
| 863 | rev (bin_to_bl n (bina * binb))" | |
| 864 | apply (induct n) | |
| 865 | apply simp | |
| 866 | apply (unfold bin_to_bl_def) | |
| 867 | apply clarsimp | |
| 868 | apply (case_tac bina rule: bin_exhaust) | |
| 869 | apply (case_tac binb rule: bin_exhaust) | |
| 870 | apply (case_tac b) | |
| 871 | apply (case_tac [!] "ba") | |
| 46653 | 872 | apply (auto simp: bin_to_bl_aux_alt Let_def) | 
| 873 | apply (auto simp: rbbl_Cons rbl_mult_Suc rbl_add) | |
| 24333 | 874 | done | 
| 875 | ||
| 876 | lemma rbl_add_split: | |
| 877 | "P (rbl_add (y # ys) (x # xs)) = | |
| 878 | (ALL ws. length ws = length ys --> ws = rbl_add ys xs --> | |
| 26008 | 879 | (y --> ((x --> P (False # rbl_succ ws)) & (~ x --> P (True # ws)))) & | 
| 24333 | 880 | (~ y --> P (x # ws)))" | 
| 881 | apply (auto simp add: Let_def) | |
| 882 | apply (case_tac [!] "y") | |
| 883 | apply auto | |
| 884 | done | |
| 885 | ||
| 886 | lemma rbl_mult_split: | |
| 887 | "P (rbl_mult (y # ys) xs) = | |
| 888 | (ALL ws. length ws = Suc (length ys) --> ws = False # rbl_mult ys xs --> | |
| 889 | (y --> P (rbl_add ws xs)) & (~ y --> P ws))" | |
| 890 | by (clarsimp simp add : Let_def) | |
| 891 | ||
| 892 | ||
| 24350 | 893 | subsection "Repeated splitting or concatenation" | 
| 24333 | 894 | |
| 895 | lemma sclem: | |
| 896 | "size (concat (map (bin_to_bl n) xs)) = length xs * n" | |
| 897 | by (induct xs) auto | |
| 898 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 899 | lemma bin_cat_foldl_lem: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 900 | "foldl (%u. bin_cat u n) x xs = | 
| 24333 | 901 | bin_cat x (size xs * n) (foldl (%u. bin_cat u n) y xs)" | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 902 | apply (induct xs arbitrary: x) | 
| 24333 | 903 | apply simp | 
| 904 | apply (simp (no_asm)) | |
| 905 | apply (frule asm_rl) | |
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 906 | apply (drule meta_spec) | 
| 24333 | 907 | apply (erule trans) | 
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 908 | apply (drule_tac x = "bin_cat y n a" in meta_spec) | 
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
54854diff
changeset | 909 | apply (simp add : bin_cat_assoc_sym min.absorb2) | 
| 24333 | 910 | done | 
| 911 | ||
| 912 | lemma bin_rcat_bl: | |
| 913 | "(bin_rcat n wl) = bl_to_bin (concat (map (bin_to_bl n) wl))" | |
| 914 | apply (unfold bin_rcat_def) | |
| 915 | apply (rule sym) | |
| 916 | apply (induct wl) | |
| 917 | apply (auto simp add : bl_to_bin_append) | |
| 918 | apply (simp add : bl_to_bin_aux_alt sclem) | |
| 919 | apply (simp add : bin_cat_foldl_lem [symmetric]) | |
| 920 | done | |
| 921 | ||
| 922 | lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps | |
| 923 | lemmas rsplit_aux_simps = bin_rsplit_aux_simps | |
| 924 | ||
| 45604 | 925 | lemmas th_if_simp1 = split_if [where P = "op = l", THEN iffD1, THEN conjunct1, THEN mp] for l | 
| 926 | lemmas th_if_simp2 = split_if [where P = "op = l", THEN iffD1, THEN conjunct2, THEN mp] for l | |
| 24333 | 927 | |
| 928 | lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] | |
| 929 | ||
| 930 | lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] | |
| 931 | (* these safe to [simp add] as require calculating m - n *) | |
| 932 | lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] | |
| 933 | lemmas rbscl = bin_rsplit_aux_simp2s (2) | |
| 934 | ||
| 935 | lemmas rsplit_aux_0_simps [simp] = | |
| 936 | rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] | |
| 937 | ||
| 938 | lemma bin_rsplit_aux_append: | |
| 26557 | 939 | "bin_rsplit_aux n m c (bs @ cs) = bin_rsplit_aux n m c bs @ cs" | 
| 940 | apply (induct n m c bs rule: bin_rsplit_aux.induct) | |
| 24333 | 941 | apply (subst bin_rsplit_aux.simps) | 
| 942 | apply (subst bin_rsplit_aux.simps) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 943 | apply (clarsimp split: prod.split) | 
| 24333 | 944 | done | 
| 945 | ||
| 946 | lemma bin_rsplitl_aux_append: | |
| 26557 | 947 | "bin_rsplitl_aux n m c (bs @ cs) = bin_rsplitl_aux n m c bs @ cs" | 
| 948 | apply (induct n m c bs rule: bin_rsplitl_aux.induct) | |
| 24333 | 949 | apply (subst bin_rsplitl_aux.simps) | 
| 950 | apply (subst bin_rsplitl_aux.simps) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 951 | apply (clarsimp split: prod.split) | 
| 24333 | 952 | done | 
| 953 | ||
| 954 | lemmas rsplit_aux_apps [where bs = "[]"] = | |
| 955 | bin_rsplit_aux_append bin_rsplitl_aux_append | |
| 956 | ||
| 957 | lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def | |
| 958 | ||
| 959 | lemmas rsplit_aux_alts = rsplit_aux_apps | |
| 960 | [unfolded append_Nil rsplit_def_auxs [symmetric]] | |
| 961 | ||
| 962 | lemma bin_split_minus: "0 < n ==> bin_split (Suc (n - 1)) w = bin_split n w" | |
| 963 | by auto | |
| 964 | ||
| 965 | lemmas bin_split_minus_simp = | |
| 45604 | 966 | bin_split.Suc [THEN [2] bin_split_minus [symmetric, THEN trans]] | 
| 24333 | 967 | |
| 968 | lemma bin_split_pred_simp [simp]: | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46655diff
changeset | 969 | "(0::nat) < numeral bin \<Longrightarrow> | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46655diff
changeset | 970 | bin_split (numeral bin) w = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46655diff
changeset | 971 | (let (w1, w2) = bin_split (numeral bin - 1) (bin_rest w) | 
| 24333 | 972 | in (w1, w2 BIT bin_last w))" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46655diff
changeset | 973 | by (simp only: bin_split_minus_simp) | 
| 24333 | 974 | |
| 975 | lemma bin_rsplit_aux_simp_alt: | |
| 26557 | 976 | "bin_rsplit_aux n m c bs = | 
| 24333 | 977 | (if m = 0 \<or> n = 0 | 
| 978 | then bs | |
| 979 | else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" | |
| 26557 | 980 | unfolding bin_rsplit_aux.simps [of n m c bs] | 
| 981 | apply simp | |
| 982 | apply (subst rsplit_aux_alts) | |
| 983 | apply (simp add: bin_rsplit_def) | |
| 24333 | 984 | done | 
| 985 | ||
| 986 | lemmas bin_rsplit_simp_alt = | |
| 45604 | 987 | trans [OF bin_rsplit_def bin_rsplit_aux_simp_alt] | 
| 24333 | 988 | |
| 989 | lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] | |
| 990 | ||
| 991 | lemma bin_rsplit_size_sign' [rule_format] : | |
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 992 | "\<lbrakk>n > 0; rev sw = bin_rsplit n (nw, w)\<rbrakk> \<Longrightarrow> | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 993 | (ALL v: set sw. bintrunc n v = v)" | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 994 | apply (induct sw arbitrary: nw w v) | 
| 24333 | 995 | apply clarsimp | 
| 996 | apply clarsimp | |
| 997 | apply (drule bthrs) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 998 | apply (simp (no_asm_use) add: Let_def split: prod.split_asm split_if_asm) | 
| 24333 | 999 | apply clarify | 
| 1000 | apply (drule split_bintrunc) | |
| 1001 | apply simp | |
| 1002 | done | |
| 1003 | ||
| 1004 | lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl | |
| 45604 | 1005 | rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]]] | 
| 24333 | 1006 | |
| 1007 | lemma bin_nth_rsplit [rule_format] : | |
| 1008 | "n > 0 ==> m < n ==> (ALL w k nw. rev sw = bin_rsplit n (nw, w) --> | |
| 1009 | k < size sw --> bin_nth (sw ! k) m = bin_nth w (k * n + m))" | |
| 1010 | apply (induct sw) | |
| 1011 | apply clarsimp | |
| 1012 | apply clarsimp | |
| 1013 | apply (drule bthrs) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 1014 | apply (simp (no_asm_use) add: Let_def split: prod.split_asm split_if_asm) | 
| 24333 | 1015 | apply clarify | 
| 1016 | apply (erule allE, erule impE, erule exI) | |
| 1017 | apply (case_tac k) | |
| 1018 | apply clarsimp | |
| 1019 | prefer 2 | |
| 1020 | apply clarsimp | |
| 1021 | apply (erule allE) | |
| 1022 | apply (erule (1) impE) | |
| 1023 | apply (drule bin_nth_split, erule conjE, erule allE, | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57492diff
changeset | 1024 | erule trans, simp add : ac_simps)+ | 
| 24333 | 1025 | done | 
| 1026 | ||
| 1027 | lemma bin_rsplit_all: | |
| 1028 | "0 < nw ==> nw <= n ==> bin_rsplit n (nw, w) = [bintrunc n w]" | |
| 1029 | unfolding bin_rsplit_def | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 1030 | by (clarsimp dest!: split_bintrunc simp: rsplit_aux_simp2ls split: prod.split) | 
| 24333 | 1031 | |
| 1032 | lemma bin_rsplit_l [rule_format] : | |
| 1033 | "ALL bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, bintrunc m bin)" | |
| 1034 | apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) | |
| 1035 | apply (simp (no_asm) add : bin_rsplitl_def bin_rsplit_def) | |
| 1036 | apply (rule allI) | |
| 1037 | apply (subst bin_rsplitl_aux.simps) | |
| 1038 | apply (subst bin_rsplit_aux.simps) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 1039 | apply (clarsimp simp: Let_def split: prod.split) | 
| 24333 | 1040 | apply (drule bin_split_trunc) | 
| 1041 | apply (drule sym [THEN trans], assumption) | |
| 26557 | 1042 | apply (subst rsplit_aux_alts(1)) | 
| 1043 | apply (subst rsplit_aux_alts(2)) | |
| 1044 | apply clarsimp | |
| 1045 | unfolding bin_rsplit_def bin_rsplitl_def | |
| 1046 | apply simp | |
| 24333 | 1047 | done | 
| 26557 | 1048 | |
| 24333 | 1049 | lemma bin_rsplit_rcat [rule_format] : | 
| 1050 | "n > 0 --> bin_rsplit n (n * size ws, bin_rcat n ws) = map (bintrunc n) ws" | |
| 1051 | apply (unfold bin_rsplit_def bin_rcat_def) | |
| 1052 | apply (rule_tac xs = "ws" in rev_induct) | |
| 1053 | apply clarsimp | |
| 1054 | apply clarsimp | |
| 26557 | 1055 | apply (subst rsplit_aux_alts) | 
| 1056 | unfolding bin_split_cat | |
| 1057 | apply simp | |
| 24333 | 1058 | done | 
| 1059 | ||
| 1060 | lemma bin_rsplit_aux_len_le [rule_format] : | |
| 26557 | 1061 | "\<forall>ws m. n \<noteq> 0 \<longrightarrow> ws = bin_rsplit_aux n nw w bs \<longrightarrow> | 
| 1062 | length ws \<le> m \<longleftrightarrow> nw + length bs * n \<le> m * n" | |
| 54871 | 1063 | proof - | 
| 1064 |   { fix i j j' k k' m :: nat and R
 | |
| 1065 | assume d: "(i::nat) \<le> j \<or> m < j'" | |
| 1066 | assume R1: "i * k \<le> j * k \<Longrightarrow> R" | |
| 1067 | assume R2: "Suc m * k' \<le> j' * k' \<Longrightarrow> R" | |
| 1068 | have "R" using d | |
| 1069 | apply safe | |
| 1070 | apply (rule R1, erule mult_le_mono1) | |
| 1071 | apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]]) | |
| 1072 | done | |
| 1073 | } note A = this | |
| 1074 |   { fix sc m n lb :: nat
 | |
| 1075 | have "(0::nat) < sc \<Longrightarrow> sc - n + (n + lb * n) \<le> m * n \<longleftrightarrow> sc + lb * n \<le> m * n" | |
| 1076 | apply safe | |
| 1077 | apply arith | |
| 1078 | apply (case_tac "sc >= n") | |
| 1079 | apply arith | |
| 1080 | apply (insert linorder_le_less_linear [of m lb]) | |
| 1081 | apply (erule_tac k2=n and k'2=n in A) | |
| 1082 | apply arith | |
| 1083 | apply simp | |
| 1084 | done | |
| 1085 | } note B = this | |
| 1086 | show ?thesis | |
| 1087 | apply (induct n nw w bs rule: bin_rsplit_aux.induct) | |
| 1088 | apply (subst bin_rsplit_aux.simps) | |
| 1089 | apply (simp add: B Let_def split: prod.split) | |
| 1090 | done | |
| 1091 | qed | |
| 24333 | 1092 | |
| 1093 | lemma bin_rsplit_len_le: | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1094 | "n \<noteq> 0 --> ws = bin_rsplit n (nw, w) --> (length ws <= m) = (nw <= m * n)" | 
| 24333 | 1095 | unfolding bin_rsplit_def by (clarsimp simp add : bin_rsplit_aux_len_le) | 
| 1096 | ||
| 45997 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 1097 | lemma bin_rsplit_aux_len: | 
| 
13392893ea12
use 'induct arbitrary' instead of 'rule_format' attribute
 huffman parents: 
45996diff
changeset | 1098 | "n \<noteq> 0 \<Longrightarrow> length (bin_rsplit_aux n nw w cs) = | 
| 24333 | 1099 | (nw + n - 1) div n + length cs" | 
| 26557 | 1100 | apply (induct n nw w cs rule: bin_rsplit_aux.induct) | 
| 24333 | 1101 | apply (subst bin_rsplit_aux.simps) | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 1102 | apply (clarsimp simp: Let_def split: prod.split) | 
| 24333 | 1103 | apply (erule thin_rl) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27105diff
changeset | 1104 | apply (case_tac m) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27105diff
changeset | 1105 | apply simp | 
| 24333 | 1106 | apply (case_tac "m <= n") | 
| 27677 | 1107 | apply auto | 
| 24333 | 1108 | done | 
| 1109 | ||
| 1110 | lemma bin_rsplit_len: | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1111 | "n\<noteq>0 ==> length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" | 
| 24333 | 1112 | unfolding bin_rsplit_def by (clarsimp simp add : bin_rsplit_aux_len) | 
| 1113 | ||
| 26557 | 1114 | lemma bin_rsplit_aux_len_indep: | 
| 1115 | "n \<noteq> 0 \<Longrightarrow> length bs = length cs \<Longrightarrow> | |
| 1116 | length (bin_rsplit_aux n nw v bs) = | |
| 1117 | length (bin_rsplit_aux n nw w cs)" | |
| 1118 | proof (induct n nw w cs arbitrary: v bs rule: bin_rsplit_aux.induct) | |
| 1119 | case (1 n m w cs v bs) show ?case | |
| 1120 | proof (cases "m = 0") | |
| 28298 | 1121 | case True then show ?thesis using `length bs = length cs` by simp | 
| 26557 | 1122 | next | 
| 1123 | case False | |
| 1124 | from "1.hyps" `m \<noteq> 0` `n \<noteq> 0` have hyp: "\<And>v bs. length bs = Suc (length cs) \<Longrightarrow> | |
| 1125 | length (bin_rsplit_aux n (m - n) v bs) = | |
| 1126 | length (bin_rsplit_aux n (m - n) (fst (bin_split n w)) (snd (bin_split n w) # cs))" | |
| 1127 | by auto | |
| 1128 | show ?thesis using `length bs = length cs` `n \<noteq> 0` | |
| 1129 | by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 1130 | split: prod.split) | 
| 26557 | 1131 | qed | 
| 1132 | qed | |
| 24333 | 1133 | |
| 1134 | lemma bin_rsplit_len_indep: | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1135 | "n\<noteq>0 ==> length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" | 
| 24333 | 1136 | apply (unfold bin_rsplit_def) | 
| 26557 | 1137 | apply (simp (no_asm)) | 
| 24333 | 1138 | apply (erule bin_rsplit_aux_len_indep) | 
| 1139 | apply (rule refl) | |
| 1140 | done | |
| 1141 | ||
| 54874 | 1142 | |
| 1143 | text {* Even more bit operations *}
 | |
| 1144 | ||
| 1145 | instantiation int :: bitss | |
| 1146 | begin | |
| 1147 | ||
| 1148 | definition [iff]: | |
| 1149 | "i !! n \<longleftrightarrow> bin_nth i n" | |
| 1150 | ||
| 1151 | definition | |
| 1152 | "lsb i = (i :: int) !! 0" | |
| 1153 | ||
| 1154 | definition | |
| 1155 | "set_bit i n b = bin_sc n b i" | |
| 1156 | ||
| 1157 | definition | |
| 1158 | "set_bits f = | |
| 1159 | (if \<exists>n. \<forall>n'\<ge>n. \<not> f n' then | |
| 1160 | let n = LEAST n. \<forall>n'\<ge>n. \<not> f n' | |
| 1161 | in bl_to_bin (rev (map f [0..<n])) | |
| 1162 | else if \<exists>n. \<forall>n'\<ge>n. f n' then | |
| 1163 | let n = LEAST n. \<forall>n'\<ge>n. f n' | |
| 1164 | in sbintrunc n (bl_to_bin (True # rev (map f [0..<n]))) | |
| 1165 | else 0 :: int)" | |
| 1166 | ||
| 1167 | definition | |
| 1168 | "shiftl x n = (x :: int) * 2 ^ n" | |
| 1169 | ||
| 1170 | definition | |
| 1171 | "shiftr x n = (x :: int) div 2 ^ n" | |
| 1172 | ||
| 1173 | definition | |
| 1174 | "msb x \<longleftrightarrow> (x :: int) < 0" | |
| 1175 | ||
| 1176 | instance .. | |
| 1177 | ||
| 24333 | 1178 | end | 
| 54874 | 1179 | |
| 1180 | end |