src/HOL/ex/Word.thy
author haftmann
Thu, 04 Jun 2020 19:38:50 +0000
changeset 71921 a238074c5a9d
parent 71854 6a51e64ba13d
child 71954 13bb3f5cdc5b
permissions -rw-r--r--
avoid overaggressive default simp rules
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
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
     6
theory Word
64015
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
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 64593
diff changeset
     9
    "HOL-Library.Type_Length"
71095
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
    10
    "HOL-ex.Bit_Operations"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    11
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    12
70925
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
    13
subsection \<open>Preliminaries\<close>
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
    14
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    15
definition signed_take_bit :: "nat \<Rightarrow> int \<Rightarrow> int"
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    16
  where signed_take_bit_eq_take_bit:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    17
    "signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    18
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    19
lemma signed_take_bit_eq_take_bit':
70171
haftmann
parents: 67961
diff changeset
    20
  "signed_take_bit (n - Suc 0) k = take_bit n (k + 2 ^ (n - 1)) - 2 ^ (n - 1)" if "n > 0"
haftmann
parents: 67961
diff changeset
    21
  using that by (simp add: signed_take_bit_eq_take_bit)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
    22
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    23
lemma signed_take_bit_0 [simp]:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    24
  "signed_take_bit 0 k = - (k mod 2)"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    25
proof (cases "even k")
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    26
  case True
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    27
  then have "odd (k + 1)"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    28
    by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    29
  then have "(k + 1) mod 2 = 1"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    30
    by (simp add: even_iff_mod_2_eq_zero)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    31
  with True show ?thesis
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71443
diff changeset
    32
    by (simp add: signed_take_bit_eq_take_bit take_bit_Suc)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    33
next
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    34
  case False
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    35
  then show ?thesis
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71443
diff changeset
    36
    by (simp add: signed_take_bit_eq_take_bit odd_iff_mod_2_eq_one take_bit_Suc)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    37
qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    38
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71443
diff changeset
    39
lemma signed_take_bit_Suc:
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    40
  "signed_take_bit (Suc n) k = signed_take_bit n (k div 2) * 2 + k mod 2"
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71443
diff changeset
    41
  by (simp add: odd_iff_mod_2_eq_one signed_take_bit_eq_take_bit algebra_simps take_bit_Suc)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    42
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    43
lemma signed_take_bit_of_0 [simp]:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    44
  "signed_take_bit n 0 = 0"
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    45
  by (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    46
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    47
lemma signed_take_bit_of_minus_1 [simp]:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    48
  "signed_take_bit n (- 1) = - 1"
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71443
diff changeset
    49
  by (induct n) (simp_all add: signed_take_bit_Suc)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    50
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    51
lemma signed_take_bit_eq_iff_take_bit_eq:
70171
haftmann
parents: 67961
diff changeset
    52
  "signed_take_bit (n - Suc 0) k = signed_take_bit (n - Suc 0) l \<longleftrightarrow> take_bit n k = take_bit n l" (is "?P \<longleftrightarrow> ?Q")
haftmann
parents: 67961
diff changeset
    53
  if "n > 0"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    54
proof -
70171
haftmann
parents: 67961
diff changeset
    55
  from that obtain m where m: "n = Suc m"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    56
    by (cases n) auto
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    57
  show ?thesis
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
    58
  proof
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    59
    assume ?Q
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    60
    have "take_bit (Suc m) (k + 2 ^ m) =
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    61
      take_bit (Suc m) (take_bit (Suc m) k + take_bit (Suc m) (2 ^ m))"
67961
9c31678d2139 even more on bit operations
haftmann
parents: 67907
diff changeset
    62
      by (simp only: take_bit_add)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    63
    also have "\<dots> =
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    64
      take_bit (Suc m) (take_bit (Suc m) l + take_bit (Suc m) (2 ^ m))"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    65
      by (simp only: \<open>?Q\<close> m [symmetric])
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    66
    also have "\<dots> = take_bit (Suc m) (l + 2 ^ m)"
67961
9c31678d2139 even more on bit operations
haftmann
parents: 67907
diff changeset
    67
      by (simp only: take_bit_add)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    68
    finally show ?P
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    69
      by (simp only: signed_take_bit_eq_take_bit m) simp
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    70
  next
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    71
    assume ?P
70171
haftmann
parents: 67961
diff changeset
    72
    with that have "(k + 2 ^ (n - Suc 0)) mod 2 ^ n = (l + 2 ^ (n - Suc 0)) mod 2 ^ n"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    73
      by (simp add: signed_take_bit_eq_take_bit' take_bit_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    74
    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
    75
      by (metis mod_add_eq)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    76
    then have "k mod 2 ^ n = l mod 2 ^ n"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    77
      by (metis add_diff_cancel_right' uminus_add_conv_diff)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    78
    then show ?Q
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    79
      by (simp add: take_bit_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    80
  qed
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
    81
qed
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    82
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    83
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    84
subsection \<open>Bit strings as quotient type\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    85
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    86
subsubsection \<open>Basic properties\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    87
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    88
quotient_type (overloaded) 'a word = int / "\<lambda>k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len0) l"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    89
  by (auto intro!: equivpI reflpI sympI transpI)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    90
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    91
instantiation word :: (len0) "{semiring_numeral, comm_semiring_0, comm_ring}"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    92
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    93
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    94
lift_definition zero_word :: "'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    95
  is 0
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    96
  .
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    97
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    98
lift_definition one_word :: "'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    99
  is 1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   100
  .
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   101
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   102
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
   103
  is plus
67961
9c31678d2139 even more on bit operations
haftmann
parents: 67907
diff changeset
   104
  by (subst take_bit_add [symmetric]) (simp add: take_bit_add)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   105
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   106
lift_definition uminus_word :: "'a word \<Rightarrow> 'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   107
  is uminus
71424
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   108
  by (subst take_bit_minus [symmetric]) (simp add: take_bit_minus)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   109
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   110
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
   111
  is minus
71424
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   112
  by (subst take_bit_diff [symmetric]) (simp add: take_bit_diff)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   113
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   114
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
   115
  is times
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   116
  by (auto simp add: take_bit_eq_mod intro: mod_mult_cong)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   117
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   118
instance
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   119
  by standard (transfer; simp add: algebra_simps)+
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   120
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   121
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   122
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   123
instance word :: (len) comm_ring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   124
  by standard (transfer; simp)+
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   125
70903
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   126
quickcheck_generator word
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   127
  constructors:
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   128
    "zero_class.zero :: ('a::len0) word",
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   129
    "numeral :: num \<Rightarrow> ('a::len0) word",
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   130
    "uminus :: ('a::len0) word \<Rightarrow> ('a::len0) word"
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   131
70973
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   132
context
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   133
  includes lifting_syntax
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   134
  notes power_transfer [transfer_rule]
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   135
begin
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   136
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   137
lemma power_transfer_word [transfer_rule]:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   138
  \<open>(pcr_word ===> (=) ===> pcr_word) (^) (^)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   139
  by transfer_prover
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   140
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   141
end
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   142
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   143
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   144
subsubsection \<open>Conversions\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   145
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   146
context
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   147
  includes lifting_syntax
71182
410935efbf5c characterization of singleton bit operation
haftmann
parents: 71181
diff changeset
   148
  notes 
410935efbf5c characterization of singleton bit operation
haftmann
parents: 71181
diff changeset
   149
    transfer_rule_of_bool [transfer_rule]
410935efbf5c characterization of singleton bit operation
haftmann
parents: 71181
diff changeset
   150
    transfer_rule_numeral [transfer_rule]
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   151
    transfer_rule_of_nat [transfer_rule]
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   152
    transfer_rule_of_int [transfer_rule]
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   153
begin
70348
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   154
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   155
lemma [transfer_rule]:
71182
410935efbf5c characterization of singleton bit operation
haftmann
parents: 71181
diff changeset
   156
  "((=) ===> (pcr_word :: int \<Rightarrow> 'a::len word \<Rightarrow> bool)) of_bool of_bool"
410935efbf5c characterization of singleton bit operation
haftmann
parents: 71181
diff changeset
   157
  by transfer_prover
410935efbf5c characterization of singleton bit operation
haftmann
parents: 71181
diff changeset
   158
410935efbf5c characterization of singleton bit operation
haftmann
parents: 71181
diff changeset
   159
lemma [transfer_rule]:
71760
e4e05fcd8937 generalized
haftmann
parents: 71535
diff changeset
   160
  "((=) ===> (pcr_word :: int \<Rightarrow> 'a::len0 word \<Rightarrow> bool)) numeral numeral"
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   161
  by transfer_prover
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   162
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   163
lemma [transfer_rule]:
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   164
  "((=) ===> pcr_word) int of_nat"
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   165
  by transfer_prover
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   166
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   167
lemma [transfer_rule]:
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   168
  "((=) ===> pcr_word) (\<lambda>k. k) of_int"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   169
proof -
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   170
  have "((=) ===> pcr_word) of_int of_int"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   171
    by transfer_prover
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   172
  then show ?thesis by (simp add: id_def)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   173
qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   174
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   175
end
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   176
70973
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   177
lemma abs_word_eq:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   178
  "abs_word = of_int"
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   179
  by (rule ext) (transfer, rule)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   180
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   181
context semiring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   182
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   183
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   184
lift_definition unsigned :: "'b::len0 word \<Rightarrow> 'a"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   185
  is "of_nat \<circ> nat \<circ> take_bit LENGTH('b)"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   186
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   187
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   188
lemma unsigned_0 [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   189
  "unsigned 0 = 0"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   190
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   191
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   192
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   193
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   194
context semiring_char_0
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   195
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   196
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   197
lemma word_eq_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   198
  "a = b \<longleftrightarrow> unsigned a = unsigned b"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   199
  by safe (transfer; simp add: eq_nat_nat_iff)
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
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   202
70903
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   203
instantiation word :: (len0) equal
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   204
begin
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   205
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   206
definition equal_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   207
  where "equal_word a b \<longleftrightarrow> (unsigned a :: int) = unsigned b"
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   208
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   209
instance proof
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   210
  fix a b :: "'a word"
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   211
  show "HOL.equal a b \<longleftrightarrow> a = b"
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   212
    using word_eq_iff_unsigned [of a b] by (auto simp add: equal_word_def)
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   213
qed
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   214
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   215
end
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   216
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   217
context ring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   218
begin
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
lift_definition signed :: "'b::len word \<Rightarrow> 'a"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   221
  is "of_int \<circ> signed_take_bit (LENGTH('b) - 1)"
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   222
  by (simp add: signed_take_bit_eq_iff_take_bit_eq [symmetric])
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   223
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   224
lemma signed_0 [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   225
  "signed 0 = 0"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   226
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   227
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   228
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   229
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   230
lemma unsigned_of_nat [simp]:
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   231
  "unsigned (of_nat n :: 'a word) = take_bit LENGTH('a::len) n"
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   232
  by transfer (simp add: nat_eq_iff take_bit_eq_mod zmod_int)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   233
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   234
lemma of_nat_unsigned [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   235
  "of_nat (unsigned a) = a"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   236
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   237
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   238
lemma of_int_unsigned [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   239
  "of_int (unsigned a) = a"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   240
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   241
70973
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   242
lemma unsigned_nat_less:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   243
  \<open>unsigned a < (2 ^ LENGTH('a) :: nat)\<close> for a :: \<open>'a::len0 word\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   244
  by transfer (simp add: take_bit_eq_mod)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   245
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   246
lemma unsigned_int_less:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   247
  \<open>unsigned a < (2 ^ LENGTH('a) :: int)\<close> for a :: \<open>'a::len0 word\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   248
  by transfer (simp add: take_bit_eq_mod)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   249
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   250
context ring_char_0
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   251
begin
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
lemma word_eq_iff_signed:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   254
  "a = b \<longleftrightarrow> signed a = signed b"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   255
  by safe (transfer; auto simp add: signed_take_bit_eq_iff_take_bit_eq)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   256
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   257
end
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 signed_of_int [simp]:
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   260
  "signed (of_int k :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) k"
64015
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_signed [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   264
  "of_int (signed a) = a"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   265
  by transfer (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod mod_simps)
64015
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
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   268
subsubsection \<open>Properties\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   269
71196
29e7c6d11cf1 more rules
haftmann
parents: 71195
diff changeset
   270
lemma exp_eq_zero_iff:
29e7c6d11cf1 more rules
haftmann
parents: 71195
diff changeset
   271
  \<open>(2 :: 'a::len word) ^ n = 0 \<longleftrightarrow> LENGTH('a) \<le> n\<close>
29e7c6d11cf1 more rules
haftmann
parents: 71195
diff changeset
   272
  by transfer simp
29e7c6d11cf1 more rules
haftmann
parents: 71195
diff changeset
   273
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   274
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   275
subsubsection \<open>Division\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   276
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   277
instantiation word :: (len0) modulo
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   278
begin
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
lift_definition divide_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   281
  is "\<lambda>a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   282
  by simp
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
lift_definition modulo_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   285
  is "\<lambda>a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   286
  by simp
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
instance ..
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
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   291
70973
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   292
lemma zero_word_div_eq [simp]:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   293
  \<open>0 div a = 0\<close> for a :: \<open>'a::len0 word\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   294
  by transfer simp
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   295
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   296
lemma div_zero_word_eq [simp]:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   297
  \<open>a div 0 = 0\<close> for a :: \<open>'a::len0 word\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   298
  by transfer simp
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   299
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   300
context
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   301
  includes lifting_syntax
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   302
begin
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   303
70348
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   304
lemma [transfer_rule]:
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   305
  "(pcr_word ===> (\<longleftrightarrow>)) even ((dvd) 2 :: 'a::len word \<Rightarrow> bool)"
70348
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   306
proof -
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   307
  have even_word_unfold: "even k \<longleftrightarrow> (\<exists>l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \<longleftrightarrow> ?Q")
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   308
    for k :: int
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   309
  proof
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   310
    assume ?P
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   311
    then show ?Q
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   312
      by auto
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   313
  next
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   314
    assume ?Q
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   315
    then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" ..
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   316
    then have "even (take_bit LENGTH('a) k)"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   317
      by simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   318
    then show ?P
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   319
      by simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   320
  qed
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   321
  show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def])
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   322
    transfer_prover
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   323
qed
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   324
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   325
end
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   326
70348
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   327
instance word :: (len) semiring_modulo
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   328
proof
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   329
  show "a div b * b + a mod b = a" for a b :: "'a word"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   330
  proof transfer
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   331
    fix k l :: int
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   332
    define r :: int where "r = 2 ^ LENGTH('a)"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   333
    then have r: "take_bit LENGTH('a) k = k mod r" for k
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   334
      by (simp add: take_bit_eq_mod)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   335
    have "k mod r = ((k mod r) div (l mod r) * (l mod r)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   336
      + (k mod r) mod (l mod r)) mod r"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   337
      by (simp add: div_mult_mod_eq)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   338
    also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   339
      + (k mod r) mod (l mod r)) mod r"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   340
      by (simp add: mod_add_left_eq)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   341
    also have "... = (((k mod r) div (l mod r) * l) mod r
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   342
      + (k mod r) mod (l mod r)) mod r"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   343
      by (simp add: mod_mult_right_eq)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   344
    finally have "k mod r = ((k mod r) div (l mod r) * l
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   345
      + (k mod r) mod (l mod r)) mod r"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   346
      by (simp add: mod_simps)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   347
    with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   348
      + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   349
      by simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   350
  qed
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   351
qed
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   352
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   353
instance word :: (len) semiring_parity
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   354
proof
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   355
  show "\<not> 2 dvd (1::'a word)"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   356
    by transfer simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   357
  show even_iff_mod_2_eq_0: "2 dvd a \<longleftrightarrow> a mod 2 = 0"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   358
    for a :: "'a word"
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71443
diff changeset
   359
    by transfer (simp_all add: mod_2_eq_odd take_bit_Suc)
70348
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   360
  show "\<not> 2 dvd a \<longleftrightarrow> a mod 2 = 1"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   361
    for a :: "'a word"
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71443
diff changeset
   362
    by transfer (simp_all add: mod_2_eq_odd take_bit_Suc)
70348
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   363
qed
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   364
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   365
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   366
subsubsection \<open>Orderings\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   367
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   368
instantiation word :: (len0) linorder
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   369
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   370
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   371
lift_definition less_eq_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   372
  is "\<lambda>a b. take_bit LENGTH('a) a \<le> take_bit LENGTH('a) b"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   373
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   374
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   375
lift_definition less_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   376
  is "\<lambda>a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   377
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   378
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   379
instance
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   380
  by standard (transfer; auto)+
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   381
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   382
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   383
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   384
context linordered_semidom
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   385
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   386
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   387
lemma word_less_eq_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   388
  "a \<le> b \<longleftrightarrow> unsigned a \<le> unsigned b"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   389
  by (transfer fixing: less_eq) (simp add: nat_le_eq_zle)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   390
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   391
lemma word_less_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   392
  "a < b \<longleftrightarrow> unsigned a < unsigned b"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   393
  by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative])
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   394
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   395
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   396
70973
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   397
lemma word_greater_zero_iff:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   398
  \<open>a > 0 \<longleftrightarrow> a \<noteq> 0\<close> for a :: \<open>'a::len0 word\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   399
  by transfer (simp add: less_le)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   400
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   401
lemma of_nat_word_eq_iff:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   402
  \<open>of_nat m = (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m = take_bit LENGTH('a) n\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   403
  by transfer (simp add: take_bit_of_nat)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   404
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   405
lemma of_nat_word_less_eq_iff:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   406
  \<open>of_nat m \<le> (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m \<le> take_bit LENGTH('a) n\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   407
  by transfer (simp add: take_bit_of_nat)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   408
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   409
lemma of_nat_word_less_iff:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   410
  \<open>of_nat m < (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m < take_bit LENGTH('a) n\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   411
  by transfer (simp add: take_bit_of_nat)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   412
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   413
lemma of_nat_word_eq_0_iff:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   414
  \<open>of_nat n = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd n\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   415
  using of_nat_word_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   416
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   417
lemma of_int_word_eq_iff:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   418
  \<open>of_int k = (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   419
  by transfer rule
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   420
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   421
lemma of_int_word_less_eq_iff:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   422
  \<open>of_int k \<le> (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k \<le> take_bit LENGTH('a) l\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   423
  by transfer rule
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   424
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   425
lemma of_int_word_less_iff:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   426
  \<open>of_int k < (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k < take_bit LENGTH('a) l\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   427
  by transfer rule
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   428
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   429
lemma of_int_word_eq_0_iff:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   430
  \<open>of_int k = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd k\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   431
  using of_int_word_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   432
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   433
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   434
subsection \<open>Bit structure on \<^typ>\<open>'a word\<close>\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   435
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   436
lemma word_bit_induct [case_names zero even odd]:
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   437
  \<open>P a\<close> if word_zero: \<open>P 0\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   438
    and word_even: \<open>\<And>a. P a \<Longrightarrow> 0 < a \<Longrightarrow> a < 2 ^ (LENGTH('a) - 1) \<Longrightarrow> P (2 * a)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   439
    and word_odd: \<open>\<And>a. P a \<Longrightarrow> a < 2 ^ (LENGTH('a) - 1) \<Longrightarrow> P (1 + 2 * a)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   440
  for P and a :: \<open>'a::len word\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   441
proof -
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   442
  define m :: nat where \<open>m = LENGTH('a) - 1\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   443
  then have l: \<open>LENGTH('a) = Suc m\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   444
    by simp
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   445
  define n :: nat where \<open>n = unsigned a\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   446
  then have \<open>n < 2 ^ LENGTH('a)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   447
    by (simp add: unsigned_nat_less)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   448
  then have \<open>n < 2 * 2 ^ m\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   449
    by (simp add: l)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   450
  then have \<open>P (of_nat n)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   451
  proof (induction n rule: nat_bit_induct)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   452
    case zero
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   453
    show ?case
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   454
      by simp (rule word_zero)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   455
  next
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   456
    case (even n)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   457
    then have \<open>n < 2 ^ m\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   458
      by simp
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   459
    with even.IH have \<open>P (of_nat n)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   460
      by simp
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   461
    moreover from \<open>n < 2 ^ m\<close> even.hyps have \<open>0 < (of_nat n :: 'a word)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   462
      by (auto simp add: word_greater_zero_iff of_nat_word_eq_0_iff l)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   463
    moreover from \<open>n < 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   464
      using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>]
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   465
      by (cases \<open>m = 0\<close>) (simp_all add: not_less take_bit_eq_self ac_simps l)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   466
    ultimately have \<open>P (2 * of_nat n)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   467
      by (rule word_even)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   468
    then show ?case
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   469
      by simp
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   470
  next
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   471
    case (odd n)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   472
    then have \<open>Suc n \<le> 2 ^ m\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   473
      by simp
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   474
    with odd.IH have \<open>P (of_nat n)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   475
      by simp
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   476
    moreover from \<open>Suc n \<le> 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   477
      using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>]
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   478
      by (cases \<open>m = 0\<close>) (simp_all add: not_less take_bit_eq_self ac_simps l)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   479
    ultimately have \<open>P (1 + 2 * of_nat n)\<close>
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   480
      by (rule word_odd)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   481
    then show ?case
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   482
      by simp
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   483
  qed
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   484
  then show ?thesis
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   485
    by (simp add: n_def)
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   486
qed
a7a52ba0717d more lemmas
haftmann
parents: 70927
diff changeset
   487
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   488
lemma bit_word_half_eq:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   489
  \<open>(of_bool b + a * 2) div 2 = a\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   490
    if \<open>a < 2 ^ (LENGTH('a) - Suc 0)\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   491
    for a :: \<open>'a::len word\<close>
71195
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   492
proof (cases \<open>2 \<le> LENGTH('a::len)\<close>)
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   493
  case False
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   494
  have \<open>of_bool (odd k) < (1 :: int) \<longleftrightarrow> even k\<close> for k :: int
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   495
    by auto
71195
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   496
  with False that show ?thesis
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71443
diff changeset
   497
    by transfer (simp add: eq_iff)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   498
next
71195
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   499
  case True
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   500
  obtain n where length: \<open>LENGTH('a) = Suc n\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   501
    by (cases \<open>LENGTH('a)\<close>) simp_all
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   502
  show ?thesis proof (cases b)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   503
    case False
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   504
    moreover have \<open>a * 2 div 2 = a\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   505
    using that proof transfer
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   506
      fix k :: int
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   507
      from length have \<open>k * 2 mod 2 ^ LENGTH('a) = (k mod 2 ^ n) * 2\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   508
        by simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   509
      moreover assume \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   510
      with \<open>LENGTH('a) = Suc n\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   511
      have \<open>k mod 2 ^ LENGTH('a) = k mod 2 ^ n\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   512
        by (simp add: take_bit_eq_mod divmod_digit_0)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   513
      ultimately have \<open>take_bit LENGTH('a) (k * 2) = take_bit LENGTH('a) k * 2\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   514
        by (simp add: take_bit_eq_mod)
71195
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   515
      with True show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (k * 2) div take_bit LENGTH('a) 2)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   516
        = take_bit LENGTH('a) k\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   517
        by simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   518
    qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   519
    ultimately show ?thesis
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   520
      by simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   521
  next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   522
    case True
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   523
    moreover have \<open>(1 + a * 2) div 2 = a\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   524
    using that proof transfer
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   525
      fix k :: int
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   526
      from length have \<open>(1 + k * 2) mod 2 ^ LENGTH('a) = 1 + (k mod 2 ^ n) * 2\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   527
        using pos_zmod_mult_2 [of \<open>2 ^ n\<close> k] by (simp add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   528
      moreover assume \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   529
      with \<open>LENGTH('a) = Suc n\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   530
      have \<open>k mod 2 ^ LENGTH('a) = k mod 2 ^ n\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   531
        by (simp add: take_bit_eq_mod divmod_digit_0)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   532
      ultimately have \<open>take_bit LENGTH('a) (1 + k * 2) = 1 + take_bit LENGTH('a) k * 2\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   533
        by (simp add: take_bit_eq_mod)
71195
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   534
      with True show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (1 + k * 2) div take_bit LENGTH('a) 2)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   535
        = take_bit LENGTH('a) k\<close>
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71443
diff changeset
   536
        by (auto simp add: take_bit_Suc)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   537
    qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   538
    ultimately show ?thesis
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents: 70973
diff changeset
   539
      by simp
70925
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   540
  qed
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   541
qed
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   542
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71196
diff changeset
   543
lemma even_mult_exp_div_word_iff:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71196
diff changeset
   544
  \<open>even (a * 2 ^ m div 2 ^ n) \<longleftrightarrow> \<not> (
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71196
diff changeset
   545
    m \<le> n \<and>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71196
diff changeset
   546
    n < LENGTH('a) \<and> odd (a div 2 ^ (n - m)))\<close> for a :: \<open>'a::len word\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71196
diff changeset
   547
  by transfer
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71196
diff changeset
   548
    (auto simp flip: drop_bit_eq_div simp add: even_drop_bit_iff_not_bit bit_take_bit_iff,
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   549
      simp_all flip: push_bit_eq_mult add: bit_push_bit_iff_int)
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71196
diff changeset
   550
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   551
instance word :: (len) semiring_bits
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   552
proof
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   553
  show \<open>P a\<close> if stable: \<open>\<And>a. a div 2 = a \<Longrightarrow> P a\<close>
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   554
    and rec: \<open>\<And>a b. P a \<Longrightarrow> (of_bool b + 2 * a) div 2 = a \<Longrightarrow> P (of_bool b + 2 * a)\<close>
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   555
  for P and a :: \<open>'a word\<close>
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   556
  proof (induction a rule: word_bit_induct)
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   557
    case zero
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   558
    from stable [of 0] show ?case
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   559
      by simp
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   560
  next
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   561
    case (even a)
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   562
    with rec [of a False] show ?case
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   563
      using bit_word_half_eq [of a False] by (simp add: ac_simps)
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   564
  next
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   565
    case (odd a)
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   566
    with rec [of a True] show ?case
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   567
      using bit_word_half_eq [of a True] by (simp add: ac_simps)
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   568
  qed
71138
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   569
  show \<open>0 div a = 0\<close>
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   570
    for a :: \<open>'a word\<close>
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   571
    by transfer simp
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   572
  show \<open>a div 1 = a\<close>
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   573
    for a :: \<open>'a word\<close>
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   574
    by transfer simp
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   575
  show \<open>a mod b div b = 0\<close>
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   576
    for a b :: \<open>'a word\<close>
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   577
    apply transfer
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   578
    apply (simp add: take_bit_eq_mod)
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   579
    apply (subst (3) mod_pos_pos_trivial [of _ \<open>2 ^ LENGTH('a)\<close>])
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   580
      apply simp_all
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   581
     apply (metis le_less mod_by_0 pos_mod_conj zero_less_numeral zero_less_power)
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   582
    using pos_mod_bound [of \<open>2 ^ LENGTH('a)\<close>] apply simp
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   583
  proof -
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   584
    fix aa :: int and ba :: int
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   585
    have f1: "\<And>i n. (i::int) mod 2 ^ n = 0 \<or> 0 < i mod 2 ^ n"
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   586
      by (metis le_less take_bit_eq_mod take_bit_nonnegative)
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   587
    have "(0::int) < 2 ^ len_of (TYPE('a)::'a itself) \<and> ba mod 2 ^ len_of (TYPE('a)::'a itself) \<noteq> 0 \<or> aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)"
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   588
      by (metis (no_types) mod_by_0 unique_euclidean_semiring_numeral_class.pos_mod_bound zero_less_numeral zero_less_power)
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   589
    then show "aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)"
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   590
      using f1 by (meson le_less less_le_trans unique_euclidean_semiring_numeral_class.pos_mod_bound)
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   591
  qed
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   592
  show \<open>(1 + a) div 2 = a div 2\<close>
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   593
    if \<open>even a\<close>
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   594
    for a :: \<open>'a word\<close>
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71760
diff changeset
   595
    using that by transfer
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71760
diff changeset
   596
      (auto dest: le_Suc_ex simp add: mod_2_eq_odd take_bit_Suc elim!: evenE)
71182
410935efbf5c characterization of singleton bit operation
haftmann
parents: 71181
diff changeset
   597
  show \<open>(2 :: 'a word) ^ m div 2 ^ n = of_bool ((2 :: 'a word) ^ m \<noteq> 0 \<and> n \<le> m) * 2 ^ (m - n)\<close>
410935efbf5c characterization of singleton bit operation
haftmann
parents: 71181
diff changeset
   598
    for m n :: nat
410935efbf5c characterization of singleton bit operation
haftmann
parents: 71181
diff changeset
   599
    by transfer (simp, simp add: exp_div_exp_eq)
71138
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   600
  show "a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)"
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   601
    for a :: "'a word" and m n :: nat
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   602
    apply transfer
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   603
    apply (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: drop_bit_eq_div)
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   604
    apply (simp add: drop_bit_take_bit)
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   605
    done
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   606
  show "a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n"
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   607
    for a :: "'a word" and m n :: nat
71195
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   608
    by transfer (auto simp flip: take_bit_eq_mod simp add: ac_simps)
71138
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   609
  show \<open>a * 2 ^ m mod 2 ^ n = a mod 2 ^ (n - m) * 2 ^ m\<close>
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   610
    if \<open>m \<le> n\<close> for a :: "'a word" and m n :: nat
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   611
    using that apply transfer
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   612
    apply (auto simp flip: take_bit_eq_mod)
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   613
           apply (auto simp flip: push_bit_eq_mult simp add: push_bit_take_bit split: split_min_lin)
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   614
    done
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   615
  show \<open>a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\<close>
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71095
diff changeset
   616
    for a :: "'a word" and m n :: nat
71195
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   617
    by transfer (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: take_bit_eq_mod drop_bit_eq_div split: split_min_lin)
71413
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   618
  show \<open>even ((2 ^ m - 1) div (2::'a word) ^ n) \<longleftrightarrow> 2 ^ n = (0::'a word) \<or> m \<le> n\<close>
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   619
    for m n :: nat
65ffe9e910d4 more specific class assumptions
haftmann
parents: 71412
diff changeset
   620
    by transfer (auto simp add: take_bit_of_mask even_mask_div_iff)
71424
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   621
  show \<open>even (a * 2 ^ m div 2 ^ n) \<longleftrightarrow> n < m \<or> (2::'a word) ^ n = 0 \<or> m \<le> n \<and> even (a div 2 ^ (n - m))\<close>
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   622
    for a :: \<open>'a word\<close> and m n :: nat
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   623
  proof transfer
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   624
    show \<open>even (take_bit LENGTH('a) (k * 2 ^ m) div take_bit LENGTH('a) (2 ^ n)) \<longleftrightarrow>
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   625
      n < m
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   626
      \<or> take_bit LENGTH('a) ((2::int) ^ n) = take_bit LENGTH('a) 0
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   627
      \<or> (m \<le> n \<and> even (take_bit LENGTH('a) k div take_bit LENGTH('a) (2 ^ (n - m))))\<close>
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   628
    for m n :: nat and k l :: int
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   629
      by (auto simp flip: take_bit_eq_mod drop_bit_eq_div push_bit_eq_mult
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   630
        simp add: div_push_bit_of_1_eq_drop_bit drop_bit_take_bit drop_bit_push_bit_int [of n m])
e83fe2c31088 rule concerning bit (push_bit ...)
haftmann
parents: 71418
diff changeset
   631
  qed
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   632
qed
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   633
71183
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   634
context
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   635
  includes lifting_syntax
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   636
begin
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   637
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71183
diff changeset
   638
lemma transfer_rule_bit_word [transfer_rule]:
71183
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   639
  \<open>((pcr_word :: int \<Rightarrow> 'a::len word \<Rightarrow> bool) ===> (=)) (\<lambda>k n. n < LENGTH('a) \<and> bit k n) bit\<close>
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   640
proof -
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   641
  let ?t = \<open>\<lambda>a n. odd (take_bit LENGTH('a) a div take_bit LENGTH('a) ((2::int) ^ n))\<close>
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   642
  have \<open>((pcr_word :: int \<Rightarrow> 'a word \<Rightarrow> bool) ===> (=)) ?t bit\<close>
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   643
    by (unfold bit_def) transfer_prover
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   644
  also have \<open>?t = (\<lambda>k n. n < LENGTH('a) \<and> bit k n)\<close>
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   645
    by (simp add: fun_eq_iff bit_take_bit_iff flip: bit_def)
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   646
  finally show ?thesis .
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   647
qed
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   648
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   649
end
eda1ef0090ef transfer rule for bit operation
haftmann
parents: 71182
diff changeset
   650
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   651
instantiation word :: (len) semiring_bit_shifts
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   652
begin
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   653
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   654
lift_definition push_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   655
  is push_bit
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   656
proof -
71195
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   657
  show \<open>take_bit LENGTH('a) (push_bit n k) = take_bit LENGTH('a) (push_bit n l)\<close>
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   658
    if \<open>take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close> for k l :: int and n :: nat
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   659
  proof -
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   660
    from that
71195
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   661
    have \<open>take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k)
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   662
      = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)\<close>
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   663
      by simp
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   664
    moreover have \<open>min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n\<close>
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   665
      by simp
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   666
    ultimately show ?thesis
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   667
      by (simp add: take_bit_push_bit)
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   668
  qed
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   669
qed
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   670
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   671
lift_definition drop_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   672
  is \<open>\<lambda>n. drop_bit n \<circ> take_bit LENGTH('a)\<close>
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   673
  by (simp add: take_bit_eq_mod)
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   674
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   675
instance proof
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   676
  show \<open>push_bit n a = a * 2 ^ n\<close> for n :: nat and a :: "'a word"
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   677
    by transfer (simp add: push_bit_eq_mult)
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   678
  show \<open>drop_bit n a = a div 2 ^ n\<close> for n :: nat and a :: "'a word"
71195
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   679
    by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit)
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   680
qed
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   681
70925
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   682
end
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   683
71095
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   684
instantiation word :: (len) ring_bit_operations
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   685
begin
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   686
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   687
lift_definition not_word :: "'a word \<Rightarrow> 'a word"
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   688
  is not
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   689
  by (simp add: take_bit_not_iff)
71095
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   690
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   691
lift_definition and_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71196
diff changeset
   692
  is \<open>and\<close>
71095
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   693
  by simp
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   694
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   695
lift_definition or_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   696
  is or
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   697
  by simp
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   698
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   699
lift_definition xor_word ::  "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   700
  is xor
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   701
  by simp
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   702
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   703
instance proof
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71183
diff changeset
   704
  fix a b :: \<open>'a word\<close> and n :: nat
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71196
diff changeset
   705
  show \<open>- a = NOT (a - 1)\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71196
diff changeset
   706
    by transfer (simp add: minus_eq_not_minus_1)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71183
diff changeset
   707
  show \<open>bit (NOT a) n \<longleftrightarrow> (2 :: 'a word) ^ n \<noteq> 0 \<and> \<not> bit a n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71183
diff changeset
   708
    by transfer (simp add: bit_not_iff)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71183
diff changeset
   709
  show \<open>bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71183
diff changeset
   710
    by transfer (auto simp add: bit_and_iff)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71183
diff changeset
   711
  show \<open>bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71183
diff changeset
   712
    by transfer (auto simp add: bit_or_iff)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71183
diff changeset
   713
  show \<open>bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71183
diff changeset
   714
    by transfer (auto simp add: bit_xor_iff)
71095
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   715
qed
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   716
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71093
diff changeset
   717
end
71095
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   718
71854
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   719
definition even_word :: \<open>'a::len word \<Rightarrow> bool\<close>
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   720
  where [code_abbrev]: \<open>even_word = even\<close>
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   721
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   722
lemma even_word_iff [code]:
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   723
  \<open>even_word a \<longleftrightarrow> a AND 1 = 0\<close>
71921
a238074c5a9d avoid overaggressive default simp rules
haftmann
parents: 71854
diff changeset
   724
  by (simp add: even_word_def and_one_eq even_iff_mod_2_eq_zero)
71854
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   725
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   726
definition bit_word :: \<open>'a::len word \<Rightarrow> nat \<Rightarrow> bool\<close>
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   727
  where [code_abbrev]: \<open>bit_word = bit\<close>
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   728
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   729
lemma bit_word_iff [code]:
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   730
  \<open>bit_word a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close>
71921
a238074c5a9d avoid overaggressive default simp rules
haftmann
parents: 71854
diff changeset
   731
  by (simp add: bit_word_def bit_iff_odd_drop_bit and_one_eq odd_iff_mod_2_eq_one)
71854
6a51e64ba13d slightly more specific implementations
haftmann
parents: 71822
diff changeset
   732
71443
ff6394cfc05c canonical approach towards lifting
haftmann
parents: 71424
diff changeset
   733
lifting_update word.lifting
ff6394cfc05c canonical approach towards lifting
haftmann
parents: 71424
diff changeset
   734
lifting_forget word.lifting
ff6394cfc05c canonical approach towards lifting
haftmann
parents: 71424
diff changeset
   735
71095
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
   736
end