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