src/HOL/ex/Bit_Lists.thy
author haftmann
Sun, 16 Jun 2019 16:40:57 +0000
changeset 70353 7aa64296b9b0
parent 70340 7383930fc946
child 70912 8c2bef3df488
permissions -rw-r--r--
even more appropriate fact name
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     1
(*  Author:  Florian Haftmann, TUM
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     2
*)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     3
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     4
section \<open>Proof of concept for algebraically founded lists of bits\<close>
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     5
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     6
theory Bit_Lists
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     7
  imports Main
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     8
begin
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     9
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    10
context comm_semiring_1
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    11
begin
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    12
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    13
primrec of_unsigned :: "bool list \<Rightarrow> 'a"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    14
  where "of_unsigned [] = 0"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    15
  | "of_unsigned (b # bs) = of_bool b + 2 * of_unsigned bs"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    16
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    17
end
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    18
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    19
context comm_ring_1
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    20
begin
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    21
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    22
definition of_signed :: "bool list \<Rightarrow> 'a"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    23
  where "of_signed bs = (if bs = [] then 0 else if last bs
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    24
    then - (of_unsigned (map Not bs) + 1) else of_unsigned bs)"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    25
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    26
end
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    27
70340
7383930fc946 slightly more specialized name for type class
haftmann
parents: 70339
diff changeset
    28
class semiring_bits = unique_euclidean_semiring_with_nat +
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    29
  assumes half_measure: "a div 2 \<noteq> a \<Longrightarrow> euclidean_size (a div 2) < euclidean_size a"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    30
  \<comment> \<open>It is not clear whether this could be derived from already existing assumptions.\<close>
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    31
begin
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    32
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    33
function bits_of :: "'a \<Rightarrow> bool list"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    34
  where "bits_of a = odd a # (let b = a div 2 in if a = b then [] else bits_of b)"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    35
  by auto
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    36
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    37
termination
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    38
  by (relation "measure euclidean_size") (auto intro: half_measure)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    39
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    40
lemma bits_of_not_empty [simp]:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    41
  "bits_of a \<noteq> []"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    42
  by (induction a rule: bits_of.induct) simp_all
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    43
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    44
lemma bits_of_0 [simp]:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    45
  "bits_of 0 = [False]"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    46
  by simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    47
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    48
lemma bits_of_1 [simp]:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    49
  "bits_of 1 = [True, False]"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    50
  by simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    51
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    52
lemma bits_of_double [simp]:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    53
  "bits_of (a * 2) = False # (if a = 0 then [] else bits_of a)"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    54
  by simp (simp add: mult_2_right)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    55
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    56
lemma bits_of_add_1_double [simp]:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    57
  "bits_of (1 + a * 2) = True # (if a + 1 = 0 then [] else bits_of a)"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    58
  by simp (simp add: mult_2_right algebra_simps)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    59
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    60
declare bits_of.simps [simp del]
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    61
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    62
lemma not_last_bits_of_nat [simp]:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    63
  "\<not> last (bits_of (of_nat n))"
70353
7aa64296b9b0 even more appropriate fact name
haftmann
parents: 70340
diff changeset
    64
  by (induction n rule: nat_bit_induct)
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    65
    (use of_nat_neq_0 in \<open>simp_all add: algebra_simps\<close>)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    66
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    67
lemma of_unsigned_bits_of_nat:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    68
  "of_unsigned (bits_of (of_nat n)) = of_nat n"
70353
7aa64296b9b0 even more appropriate fact name
haftmann
parents: 70340
diff changeset
    69
  by (induction n rule: nat_bit_induct)
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    70
    (use of_nat_neq_0 in \<open>simp_all add: algebra_simps\<close>)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    71
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    72
end
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    73
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    74
instance nat :: semiring_bits
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    75
  by standard simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    76
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    77
lemma bits_of_Suc_double [simp]:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    78
  "bits_of (Suc (n * 2)) = True # bits_of n"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    79
  using bits_of_add_1_double [of n] by simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    80
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    81
lemma of_unsigned_bits_of:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    82
  "of_unsigned (bits_of n) = n" for n :: nat
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    83
  using of_unsigned_bits_of_nat [of n, where ?'a = nat] by simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    84
70340
7383930fc946 slightly more specialized name for type class
haftmann
parents: 70339
diff changeset
    85
class ring_bits = unique_euclidean_ring_with_nat + semiring_bits
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    86
begin
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    87
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    88
lemma bits_of_minus_1 [simp]:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    89
  "bits_of (- 1) = [True]"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    90
  using bits_of.simps [of "- 1"] by simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    91
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    92
lemma bits_of_double [simp]:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    93
  "bits_of (- (a * 2)) = False # (if a = 0 then [] else bits_of (- a))"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    94
  using bits_of.simps [of "- (a * 2)"] nonzero_mult_div_cancel_right [of 2 "- a"]
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    95
  by simp (simp add: mult_2_right)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    96
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    97
lemma bits_of_minus_1_diff_double [simp]:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    98
  "bits_of (- 1 - a * 2) = True # (if a = 0 then [] else bits_of (- 1 - a))"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    99
proof -
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   100
  have [simp]: "- 1 - a = - a - 1"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   101
    by (simp add: algebra_simps)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   102
  show ?thesis
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   103
    using bits_of.simps [of "- 1 - a * 2"] div_mult_self1 [of 2 "- 1" "- a"]
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   104
    by simp (simp add: mult_2_right algebra_simps)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   105
qed
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   106
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   107
lemma last_bits_of_neg_of_nat [simp]:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   108
  "last (bits_of (- 1 - of_nat n))"
70353
7aa64296b9b0 even more appropriate fact name
haftmann
parents: 70340
diff changeset
   109
proof (induction n rule: nat_bit_induct)
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   110
  case zero
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   111
  then show ?case
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   112
    by simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   113
next
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   114
  case (even n)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   115
  then show ?case
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   116
    by (simp add: algebra_simps)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   117
next
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   118
  case (odd n)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   119
  then have "last (bits_of ((- 1 - of_nat n) * 2))"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   120
    by auto
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   121
  also have "(- 1 - of_nat n) * 2 = - 1 - (1 + 2 * of_nat n)"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   122
    by (simp add: algebra_simps)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   123
  finally show ?case
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   124
    by simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   125
qed
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   126
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   127
lemma of_signed_bits_of_nat:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   128
  "of_signed (bits_of (of_nat n)) = of_nat n"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   129
  by (simp add: of_signed_def of_unsigned_bits_of_nat)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   130
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   131
lemma of_signed_bits_neg_of_nat:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   132
  "of_signed (bits_of (- 1 - of_nat n)) = - 1 - of_nat n"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   133
proof -
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   134
  have "of_unsigned (map Not (bits_of (- 1 - of_nat n))) = of_nat n"
70353
7aa64296b9b0 even more appropriate fact name
haftmann
parents: 70340
diff changeset
   135
  proof (induction n rule: nat_bit_induct)
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   136
    case zero
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   137
    then show ?case
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   138
      by simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   139
  next
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   140
    case (even n)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   141
    then show ?case
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   142
      by (simp add: algebra_simps)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   143
  next
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   144
    case (odd n)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   145
    have *: "- 1 - (1 + of_nat n * 2) = - 2 - of_nat n * 2"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   146
      by (simp add: algebra_simps) (metis add_assoc one_add_one)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   147
    from odd show ?case
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   148
      using bits_of_double [of "of_nat (Suc n)"] of_nat_neq_0
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   149
      by (simp add: algebra_simps *)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   150
  qed
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   151
  then show ?thesis
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   152
    by (simp add: of_signed_def algebra_simps)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   153
qed
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   154
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   155
lemma of_signed_bits_of_int:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   156
  "of_signed (bits_of (of_int k)) = of_int k"
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   157
  by (cases k rule: int_cases)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   158
    (simp_all add: of_signed_bits_of_nat of_signed_bits_neg_of_nat)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   159
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   160
end
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   161
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   162
instance int :: ring_bits
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   163
  by standard auto
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   164
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   165
lemma of_signed_bits_of:
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   166
  "of_signed (bits_of k) = k" for k :: int
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   167
  using of_signed_bits_of_int [of k, where ?'a = int] by simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   168
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   169
end