src/HOL/ex/Word_Type.thy
author haftmann
Sun, 16 Oct 2016 09:31:05 +0200
changeset 64242 93c6f0da5c70
parent 64114 45e065eea984
child 64593 50c715579715
permissions -rw-r--r--
more standardized theorem names for facts involving the div and mod identity
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     1
(*  Author:  Florian Haftmann, TUM
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     2
*)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     3
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     4
section \<open>Proof of concept for algebraically founded bit word types\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     5
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     6
theory Word_Type
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     7
  imports
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     8
    Main
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     9
    "~~/src/HOL/Library/Type_Length"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    10
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    11
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    12
subsection \<open>Truncating bit representations of numeric types\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    13
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    14
class semiring_bits = semiring_div_parity +
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    15
  assumes semiring_bits: "(1 + 2 * a) mod of_nat (2 * n) = 1 + 2 * (a mod of_nat n)"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    16
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    17
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    18
definition bitrunc :: "nat \<Rightarrow> 'a \<Rightarrow> 'a"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    19
  where bitrunc_eq_mod: "bitrunc n a = a mod of_nat (2 ^ n)"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    20
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    21
lemma bitrunc_bitrunc [simp]:
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    22
  "bitrunc n (bitrunc n a) = bitrunc n a"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    23
  by (simp add: bitrunc_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    24
  
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    25
lemma bitrunc_0 [simp]:
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    26
  "bitrunc 0 a = 0"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    27
  by (simp add: bitrunc_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    28
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    29
lemma bitrunc_Suc [simp]:
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    30
  "bitrunc (Suc n) a = bitrunc n (a div 2) * 2 + a mod 2"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    31
proof -
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    32
  define b and c
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    33
    where "b = a div 2" and "c = a mod 2"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    34
  then have a: "a = b * 2 + c" 
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    35
    and "c = 0 \<or> c = 1"
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64114
diff changeset
    36
    by (simp_all add: div_mult_mod_eq parity)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    37
  from \<open>c = 0 \<or> c = 1\<close>
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    38
  have "bitrunc (Suc n) (b * 2 + c) = bitrunc n b * 2 + c"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    39
  proof
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    40
    assume "c = 0"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    41
    moreover have "(2 * b) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n)"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    42
      by (simp add: mod_mult_mult1)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    43
    ultimately show ?thesis
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    44
      by (simp add: bitrunc_eq_mod ac_simps)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    45
  next
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    46
    assume "c = 1"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    47
    with semiring_bits [of b "2 ^ n"] show ?thesis
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    48
      by (simp add: bitrunc_eq_mod ac_simps)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    49
  qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    50
  with a show ?thesis
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    51
    by (simp add: b_def c_def)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    52
qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    53
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    54
lemma bitrunc_of_0 [simp]:
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    55
  "bitrunc n 0 = 0"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    56
  by (simp add: bitrunc_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    57
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    58
lemma bitrunc_plus:
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    59
  "bitrunc n (bitrunc n a + bitrunc n b) = bitrunc n (a + b)"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    60
  by (simp add: bitrunc_eq_mod mod_add_eq [symmetric])
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    61
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    62
lemma bitrunc_of_1_eq_0_iff [simp]:
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    63
  "bitrunc n 1 = 0 \<longleftrightarrow> n = 0"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    64
  by (induct n) simp_all
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    65
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    66
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    67
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    68
instance nat :: semiring_bits
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    69
  by standard (simp add: mod_Suc Suc_double_not_eq_double)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    70
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    71
instance int :: semiring_bits
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    72
  by standard (simp add: pos_zmod_mult_2)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    73
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    74
lemma bitrunc_uminus:
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    75
  fixes k :: int
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    76
  shows "bitrunc n (- (bitrunc n k)) = bitrunc n (- k)"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    77
  by (simp add: bitrunc_eq_mod mod_minus_eq [symmetric])
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    78
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    79
lemma bitrunc_minus:
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    80
  fixes k l :: int
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    81
  shows "bitrunc n (bitrunc n k - bitrunc n l) = bitrunc n (k - l)"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    82
  by (simp add: bitrunc_eq_mod mod_diff_eq [symmetric])
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    83
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    84
lemma bitrunc_nonnegative [simp]:
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    85
  fixes k :: int
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    86
  shows "bitrunc n k \<ge> 0"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    87
  by (simp add: bitrunc_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    88
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    89
definition signed_bitrunc :: "nat \<Rightarrow> int \<Rightarrow> int"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    90
  where signed_bitrunc_eq_bitrunc:
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    91
    "signed_bitrunc n k = bitrunc (Suc n) (k + 2 ^ n) - 2 ^ n"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    92
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    93
lemma signed_bitrunc_eq_bitrunc':
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    94
  assumes "n > 0"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    95
  shows "signed_bitrunc (n - Suc 0) k = bitrunc n (k + 2 ^ (n - 1)) - 2 ^ (n - 1)"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    96
  using assms by (simp add: signed_bitrunc_eq_bitrunc)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    97
  
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    98
lemma signed_bitrunc_0 [simp]:
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
    99
  "signed_bitrunc 0 k = - (k mod 2)"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   100
proof (cases "even k")
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   101
  case True
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   102
  then have "odd (k + 1)"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   103
    by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   104
  then have "(k + 1) mod 2 = 1"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   105
    by (simp add: even_iff_mod_2_eq_zero)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   106
  with True show ?thesis
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   107
    by (simp add: signed_bitrunc_eq_bitrunc)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   108
next
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   109
  case False
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   110
  then show ?thesis
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   111
    by (simp add: signed_bitrunc_eq_bitrunc odd_iff_mod_2_eq_one)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   112
qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   113
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   114
lemma signed_bitrunc_Suc [simp]:
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   115
  "signed_bitrunc (Suc n) k = signed_bitrunc n (k div 2) * 2 + k mod 2"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   116
  using zero_not_eq_two by (simp add: signed_bitrunc_eq_bitrunc algebra_simps)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   117
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   118
lemma signed_bitrunc_of_0 [simp]:
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   119
  "signed_bitrunc n 0 = 0"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   120
  by (simp add: signed_bitrunc_eq_bitrunc bitrunc_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   121
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   122
lemma signed_bitrunc_of_minus_1 [simp]:
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   123
  "signed_bitrunc n (- 1) = - 1"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   124
  by (induct n) simp_all
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   125
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   126
lemma signed_bitrunc_eq_iff_bitrunc_eq:
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   127
  assumes "n > 0"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   128
  shows "signed_bitrunc (n - Suc 0) k = signed_bitrunc (n - Suc 0) l \<longleftrightarrow> bitrunc n k = bitrunc n l" (is "?P \<longleftrightarrow> ?Q")
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   129
proof -
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   130
  from assms obtain m where m: "n = Suc m"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   131
    by (cases n) auto
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   132
  show ?thesis
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   133
  proof 
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   134
    assume ?Q
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   135
    have "bitrunc (Suc m) (k + 2 ^ m) =
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   136
      bitrunc (Suc m) (bitrunc (Suc m) k + bitrunc (Suc m) (2 ^ m))"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   137
      by (simp only: bitrunc_plus)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   138
    also have "\<dots> =
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   139
      bitrunc (Suc m) (bitrunc (Suc m) l + bitrunc (Suc m) (2 ^ m))"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   140
      by (simp only: \<open>?Q\<close> m [symmetric])
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   141
    also have "\<dots> = bitrunc (Suc m) (l + 2 ^ m)"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   142
      by (simp only: bitrunc_plus)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   143
    finally show ?P
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   144
      by (simp only: signed_bitrunc_eq_bitrunc m) simp
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   145
  next
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   146
    assume ?P
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   147
    with assms have "(k + 2 ^ (n - Suc 0)) mod 2 ^ n = (l + 2 ^ (n - Suc 0)) mod 2 ^ n"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   148
      by (simp add: signed_bitrunc_eq_bitrunc' bitrunc_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   149
    then have "(i + (k + 2 ^ (n - Suc 0))) mod 2 ^ n = (i + (l + 2 ^ (n - Suc 0))) mod 2 ^ n" for i
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   150
      by (metis mod_add_eq)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   151
    then have "k mod 2 ^ n = l mod 2 ^ n"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   152
      by (metis add_diff_cancel_right' uminus_add_conv_diff)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   153
    then show ?Q
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   154
      by (simp add: bitrunc_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   155
  qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   156
qed 
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   157
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   158
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   159
subsection \<open>Bit strings as quotient type\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   160
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   161
subsubsection \<open>Basic properties\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   162
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   163
quotient_type (overloaded) 'a word = int / "\<lambda>k l. bitrunc LENGTH('a) k = bitrunc LENGTH('a::len0) l"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   164
  by (auto intro!: equivpI reflpI sympI transpI)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   165
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   166
instantiation word :: (len0) "{semiring_numeral, comm_semiring_0, comm_ring}"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   167
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   168
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   169
lift_definition zero_word :: "'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   170
  is 0
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   171
  .
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   172
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   173
lift_definition one_word :: "'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   174
  is 1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   175
  .
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   176
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   177
lift_definition plus_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   178
  is plus
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   179
  by (subst bitrunc_plus [symmetric]) (simp add: bitrunc_plus)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   180
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   181
lift_definition uminus_word :: "'a word \<Rightarrow> 'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   182
  is uminus
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   183
  by (subst bitrunc_uminus [symmetric]) (simp add: bitrunc_uminus)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   184
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   185
lift_definition minus_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   186
  is minus
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   187
  by (subst bitrunc_minus [symmetric]) (simp add: bitrunc_minus)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   188
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   189
lift_definition times_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   190
  is times
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   191
  by (auto simp add: bitrunc_eq_mod intro: mod_mult_cong)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   192
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   193
instance
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   194
  by standard (transfer; simp add: algebra_simps)+
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   195
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   196
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   197
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   198
instance word :: (len) comm_ring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   199
  by standard (transfer; simp)+
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   200
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   201
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   202
subsubsection \<open>Conversions\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   203
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   204
lemma [transfer_rule]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   205
  "rel_fun HOL.eq pcr_word int of_nat"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   206
proof -
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   207
  note transfer_rule_of_nat [transfer_rule]
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   208
  show ?thesis by transfer_prover
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   209
qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   210
  
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   211
lemma [transfer_rule]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   212
  "rel_fun HOL.eq pcr_word (\<lambda>k. k) of_int"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   213
proof -
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   214
  note transfer_rule_of_int [transfer_rule]
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   215
  have "rel_fun HOL.eq pcr_word (of_int :: int \<Rightarrow> int) (of_int :: int \<Rightarrow> 'a word)"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   216
    by transfer_prover
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   217
  then show ?thesis by (simp add: id_def)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   218
qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   219
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   220
context semiring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   221
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   222
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   223
lift_definition unsigned :: "'b::len0 word \<Rightarrow> 'a"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   224
  is "of_nat \<circ> nat \<circ> bitrunc LENGTH('b)"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   225
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   226
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   227
lemma unsigned_0 [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   228
  "unsigned 0 = 0"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   229
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   230
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   231
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   232
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   233
context semiring_char_0
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   234
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   235
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   236
lemma word_eq_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   237
  "a = b \<longleftrightarrow> unsigned a = unsigned b"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   238
  by safe (transfer; simp add: eq_nat_nat_iff)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   239
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   240
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   241
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   242
context ring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   243
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   244
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   245
lift_definition signed :: "'b::len word \<Rightarrow> 'a"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   246
  is "of_int \<circ> signed_bitrunc (LENGTH('b) - 1)"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   247
  by (simp add: signed_bitrunc_eq_iff_bitrunc_eq [symmetric])
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   248
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   249
lemma signed_0 [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   250
  "signed 0 = 0"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   251
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   252
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   253
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   254
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   255
lemma unsigned_of_nat [simp]:
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   256
  "unsigned (of_nat n :: 'a word) = bitrunc LENGTH('a::len) n"
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   257
  by transfer (simp add: nat_eq_iff bitrunc_eq_mod zmod_int)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   258
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   259
lemma of_nat_unsigned [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   260
  "of_nat (unsigned a) = a"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   261
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   262
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   263
lemma of_int_unsigned [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   264
  "of_int (unsigned a) = a"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   265
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   266
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   267
context ring_char_0
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   268
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   269
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   270
lemma word_eq_iff_signed:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   271
  "a = b \<longleftrightarrow> signed a = signed b"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   272
  by safe (transfer; auto simp add: signed_bitrunc_eq_iff_bitrunc_eq)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   273
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   274
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   275
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   276
lemma signed_of_int [simp]:
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   277
  "signed (of_int k :: 'a word) = signed_bitrunc (LENGTH('a::len) - 1) k"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   278
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   279
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   280
lemma of_int_signed [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   281
  "of_int (signed a) = a"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   282
  by transfer (simp add: signed_bitrunc_eq_bitrunc bitrunc_eq_mod zdiff_zmod_left)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   283
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   284
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   285
subsubsection \<open>Properties\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   286
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   287
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   288
subsubsection \<open>Division\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   289
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   290
instantiation word :: (len0) modulo
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   291
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   292
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   293
lift_definition divide_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   294
  is "\<lambda>a b. bitrunc LENGTH('a) a div bitrunc LENGTH('a) b"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   295
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   296
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   297
lift_definition modulo_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   298
  is "\<lambda>a b. bitrunc LENGTH('a) a mod bitrunc LENGTH('a) b"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   299
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   300
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   301
instance ..
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   302
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   303
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   304
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   305
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   306
subsubsection \<open>Orderings\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   307
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   308
instantiation word :: (len0) linorder
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   309
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   310
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   311
lift_definition less_eq_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   312
  is "\<lambda>a b. bitrunc LENGTH('a) a \<le> bitrunc LENGTH('a) b"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   313
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   314
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   315
lift_definition less_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   316
  is "\<lambda>a b. bitrunc LENGTH('a) a < bitrunc LENGTH('a) b"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   317
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   318
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   319
instance
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   320
  by standard (transfer; auto)+
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   321
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   322
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   323
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   324
context linordered_semidom
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   325
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   326
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   327
lemma word_less_eq_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   328
  "a \<le> b \<longleftrightarrow> unsigned a \<le> unsigned b"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   329
  by (transfer fixing: less_eq) (simp add: nat_le_eq_zle)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   330
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   331
lemma word_less_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   332
  "a < b \<longleftrightarrow> unsigned a < unsigned b"
64114
45e065eea984 tuned name of bit truncating operations
haftmann
parents: 64113
diff changeset
   333
  by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF bitrunc_nonnegative])
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   334
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   335
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   336
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   337
end