src/HOL/ex/Word_Type.thy
author haftmann
Wed, 23 Oct 2019 16:09:23 +0000
changeset 70927 cc204e10385c
parent 70925 525853e4ec80
child 70973 a7a52ba0717d
permissions -rw-r--r--
tuned syntax
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     1
(*  Author:  Florian Haftmann, TUM
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     2
*)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     3
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     4
section \<open>Proof of concept for algebraically founded bit word types\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     5
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     6
theory Word_Type
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     7
  imports
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
     8
    Main
70925
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
     9
    "HOL-ex.Bit_Lists"
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 64593
diff changeset
    10
    "HOL-Library.Type_Length"
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
lemma take_bit_uminus:
70171
haftmann
parents: 67961
diff changeset
    16
  "take_bit n (- (take_bit n k)) = take_bit n (- k)" for k :: int
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    17
  by (simp add: take_bit_eq_mod mod_minus_eq)
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 take_bit_minus:
70171
haftmann
parents: 67961
diff changeset
    20
  "take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)" for k l :: int
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    21
  by (simp add: take_bit_eq_mod mod_diff_eq)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    22
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    23
lemma take_bit_nonnegative [simp]:
70171
haftmann
parents: 67961
diff changeset
    24
  "take_bit n k \<ge> 0" for k :: int
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    25
  by (simp add: take_bit_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    26
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    27
definition signed_take_bit :: "nat \<Rightarrow> int \<Rightarrow> int"
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    28
  where signed_take_bit_eq_take_bit:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    29
    "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
    30
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    31
lemma signed_take_bit_eq_take_bit':
70171
haftmann
parents: 67961
diff changeset
    32
  "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
    33
  using that by (simp add: signed_take_bit_eq_take_bit)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    34
  
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    35
lemma signed_take_bit_0 [simp]:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    36
  "signed_take_bit 0 k = - (k mod 2)"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    37
proof (cases "even k")
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    38
  case True
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    39
  then have "odd (k + 1)"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    40
    by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    41
  then have "(k + 1) mod 2 = 1"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    42
    by (simp add: even_iff_mod_2_eq_zero)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    43
  with True show ?thesis
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    44
    by (simp add: signed_take_bit_eq_take_bit)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    45
next
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    46
  case False
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    47
  then show ?thesis
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    48
    by (simp add: signed_take_bit_eq_take_bit odd_iff_mod_2_eq_one)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    49
qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    50
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    51
lemma signed_take_bit_Suc [simp]:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    52
  "signed_take_bit (Suc n) k = signed_take_bit n (k div 2) * 2 + k mod 2"
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    53
  by (simp add: odd_iff_mod_2_eq_one signed_take_bit_eq_take_bit algebra_simps)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    54
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    55
lemma signed_take_bit_of_0 [simp]:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    56
  "signed_take_bit n 0 = 0"
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    57
  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
    58
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    59
lemma signed_take_bit_of_minus_1 [simp]:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    60
  "signed_take_bit n (- 1) = - 1"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    61
  by (induct n) simp_all
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    62
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    63
lemma signed_take_bit_eq_iff_take_bit_eq:
70171
haftmann
parents: 67961
diff changeset
    64
  "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
    65
  if "n > 0"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    66
proof -
70171
haftmann
parents: 67961
diff changeset
    67
  from that obtain m where m: "n = Suc m"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    68
    by (cases n) auto
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    69
  show ?thesis
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    70
  proof 
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    71
    assume ?Q
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    72
    have "take_bit (Suc m) (k + 2 ^ m) =
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    73
      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
    74
      by (simp only: take_bit_add)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    75
    also have "\<dots> =
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    76
      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
    77
      by (simp only: \<open>?Q\<close> m [symmetric])
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    78
    also have "\<dots> = take_bit (Suc m) (l + 2 ^ m)"
67961
9c31678d2139 even more on bit operations
haftmann
parents: 67907
diff changeset
    79
      by (simp only: take_bit_add)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    80
    finally show ?P
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    81
      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
    82
  next
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    83
    assume ?P
70171
haftmann
parents: 67961
diff changeset
    84
    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
    85
      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
    86
    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
    87
      by (metis mod_add_eq)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    88
    then have "k mod 2 ^ n = l mod 2 ^ n"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    89
      by (metis add_diff_cancel_right' uminus_add_conv_diff)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    90
    then show ?Q
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    91
      by (simp add: take_bit_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    92
  qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    93
qed 
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    94
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    95
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    96
subsection \<open>Bit strings as quotient type\<close>
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
subsubsection \<open>Basic properties\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    99
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   100
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
   101
  by (auto intro!: equivpI reflpI sympI transpI)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   102
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   103
instantiation word :: (len0) "{semiring_numeral, comm_semiring_0, comm_ring}"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   104
begin
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 zero_word :: "'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   107
  is 0
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   108
  .
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 one_word :: "'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   111
  is 1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   112
  .
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 plus_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   115
  is plus
67961
9c31678d2139 even more on bit operations
haftmann
parents: 67907
diff changeset
   116
  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
   117
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   118
lift_definition uminus_word :: "'a word \<Rightarrow> 'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   119
  is uminus
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   120
  by (subst take_bit_uminus [symmetric]) (simp add: take_bit_uminus)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   121
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   122
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
   123
  is minus
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   124
  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
   125
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   126
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
   127
  is times
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   128
  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
   129
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   130
instance
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   131
  by standard (transfer; simp add: algebra_simps)+
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   132
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   133
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   134
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   135
instance word :: (len) comm_ring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   136
  by standard (transfer; simp)+
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   137
70903
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   138
quickcheck_generator word
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   139
  constructors:
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   140
    "zero_class.zero :: ('a::len0) word",
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   141
    "numeral :: num \<Rightarrow> ('a::len0) word",
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   142
    "uminus :: ('a::len0) word \<Rightarrow> ('a::len0) word"
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   143
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   144
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   145
subsubsection \<open>Conversions\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   146
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   147
context
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   148
  includes lifting_syntax
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   149
  notes transfer_rule_numeral [transfer_rule]
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   150
    transfer_rule_of_nat [transfer_rule]
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   151
    transfer_rule_of_int [transfer_rule]
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   152
begin
70348
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   153
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   154
lemma [transfer_rule]:
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   155
  "((=) ===> (pcr_word :: int \<Rightarrow> 'a::len word \<Rightarrow> bool)) numeral numeral"
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   156
  by transfer_prover
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   157
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   158
lemma [transfer_rule]:
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   159
  "((=) ===> pcr_word) int of_nat"
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   160
  by transfer_prover
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   161
  
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   162
lemma [transfer_rule]:
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   163
  "((=) ===> pcr_word) (\<lambda>k. k) of_int"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   164
proof -
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   165
  have "((=) ===> pcr_word) of_int of_int"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   166
    by transfer_prover
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   167
  then show ?thesis by (simp add: id_def)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   168
qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   169
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   170
end
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   171
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   172
context semiring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   173
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   174
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   175
lift_definition unsigned :: "'b::len0 word \<Rightarrow> 'a"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   176
  is "of_nat \<circ> nat \<circ> take_bit LENGTH('b)"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   177
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   178
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   179
lemma unsigned_0 [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   180
  "unsigned 0 = 0"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   181
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   182
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   183
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   184
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   185
context semiring_char_0
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   186
begin
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 word_eq_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   189
  "a = b \<longleftrightarrow> unsigned a = unsigned b"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   190
  by safe (transfer; simp add: eq_nat_nat_iff)
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
70903
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   194
instantiation word :: (len0) equal
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   195
begin
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   196
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   197
definition equal_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   198
  where "equal_word a b \<longleftrightarrow> (unsigned a :: int) = unsigned b"
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   199
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   200
instance proof
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   201
  fix a b :: "'a word"
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   202
  show "HOL.equal a b \<longleftrightarrow> a = b"
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   203
    using word_eq_iff_unsigned [of a b] by (auto simp add: equal_word_def)
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   204
qed
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   205
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   206
end
c550368a4e29 added quickcheck setup
haftmann
parents: 70348
diff changeset
   207
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   208
context ring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   209
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   210
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   211
lift_definition signed :: "'b::len word \<Rightarrow> 'a"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   212
  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
   213
  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
   214
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   215
lemma signed_0 [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   216
  "signed 0 = 0"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   217
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   218
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   219
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   220
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   221
lemma unsigned_of_nat [simp]:
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   222
  "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
   223
  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
   224
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   225
lemma of_nat_unsigned [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   226
  "of_nat (unsigned a) = a"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   227
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   228
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   229
lemma of_int_unsigned [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   230
  "of_int (unsigned a) = a"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   231
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   232
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   233
context ring_char_0
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   234
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   235
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   236
lemma word_eq_iff_signed:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   237
  "a = b \<longleftrightarrow> signed a = signed b"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   238
  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
   239
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   240
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   241
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   242
lemma signed_of_int [simp]:
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   243
  "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
   244
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   245
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   246
lemma of_int_signed [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   247
  "of_int (signed a) = a"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   248
  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
   249
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   250
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   251
subsubsection \<open>Properties\<close>
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
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   254
subsubsection \<open>Division\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   255
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   256
instantiation word :: (len0) modulo
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   257
begin
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
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
   260
  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
   261
  by 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
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
   264
  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
   265
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   266
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   267
instance ..
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   268
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   269
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   270
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   271
context
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   272
  includes lifting_syntax
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   273
begin
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   274
70348
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   275
lemma [transfer_rule]:
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   276
  "(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
   277
proof -
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   278
  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
   279
    for k :: int
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   280
  proof
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   281
    assume ?P
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   282
    then show ?Q
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   283
      by auto
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   284
  next
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   285
    assume ?Q
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   286
    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
   287
    then have "even (take_bit LENGTH('a) k)"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   288
      by simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   289
    then show ?P
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   290
      by simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   291
  qed
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   292
  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
   293
    transfer_prover
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   294
qed
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   295
70927
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   296
end
cc204e10385c tuned syntax
haftmann
parents: 70925
diff changeset
   297
70348
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   298
instance word :: (len) semiring_modulo
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   299
proof
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   300
  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
   301
  proof transfer
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   302
    fix k l :: int
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   303
    define r :: int where "r = 2 ^ LENGTH('a)"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   304
    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
   305
      by (simp add: take_bit_eq_mod)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   306
    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
   307
      + (k mod r) mod (l mod r)) mod r"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   308
      by (simp add: div_mult_mod_eq)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   309
    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
   310
      + (k mod r) mod (l mod r)) mod r"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   311
      by (simp add: mod_add_left_eq)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   312
    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
   313
      + (k mod r) mod (l mod r)) mod r"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   314
      by (simp add: mod_mult_right_eq)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   315
    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
   316
      + (k mod r) mod (l mod r)) mod r"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   317
      by (simp add: mod_simps)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   318
    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
   319
      + 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
   320
      by simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   321
  qed
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   322
qed
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   323
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   324
instance word :: (len) semiring_parity
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   325
proof
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   326
  show "\<not> 2 dvd (1::'a word)"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   327
    by transfer simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   328
  consider (triv) "LENGTH('a) = 1" "take_bit LENGTH('a) 2 = (0 :: int)"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   329
    | (take_bit_2) "take_bit LENGTH('a) 2 = (2 :: int)"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   330
  proof (cases "LENGTH('a) \<ge> 2")
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   331
    case False
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   332
    then have "LENGTH('a) = 1"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   333
      by (auto simp add: not_le dest: less_2_cases)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   334
    then have "take_bit LENGTH('a) 2 = (0 :: int)"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   335
      by simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   336
    with \<open>LENGTH('a) = 1\<close> triv show ?thesis
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   337
      by simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   338
  next
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   339
    case True
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   340
    then obtain n where "LENGTH('a) = Suc (Suc n)"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   341
      by (auto dest: le_Suc_ex)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   342
    then have "take_bit LENGTH('a) 2 = (2 :: int)"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   343
      by simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   344
    with take_bit_2 show ?thesis
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   345
      by simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   346
  qed
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   347
  note * = this
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   348
  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
   349
    for a :: "'a word"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   350
    by (transfer; cases rule: *) (simp_all add: mod_2_eq_odd)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   351
  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
   352
    for a :: "'a word"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   353
    by (transfer; cases rule: *) (simp_all add: mod_2_eq_odd)
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   354
qed
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 70171
diff changeset
   355
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   356
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   357
subsubsection \<open>Orderings\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   358
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   359
instantiation word :: (len0) linorder
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   360
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   361
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   362
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
   363
  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
   364
  by simp
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
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
   367
  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
   368
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   369
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   370
instance
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   371
  by standard (transfer; auto)+
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   372
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   373
end
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
context linordered_semidom
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   376
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   377
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   378
lemma word_less_eq_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   379
  "a \<le> b \<longleftrightarrow> unsigned a \<le> unsigned b"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   380
  by (transfer fixing: less_eq) (simp add: nat_le_eq_zle)
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
lemma word_less_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   383
  "a < b \<longleftrightarrow> unsigned a < unsigned b"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   384
  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
   385
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   386
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   387
70925
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   388
subsection \<open>Bit operation on \<^typ>\<open>'a word\<close>\<close>
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   389
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   390
context unique_euclidean_semiring_with_nat
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   391
begin
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   392
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   393
primrec n_bits_of :: "nat \<Rightarrow> 'a \<Rightarrow> bool list"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   394
  where
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   395
    "n_bits_of 0 a = []"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   396
  | "n_bits_of (Suc n) a = odd a # n_bits_of n (a div 2)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   397
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   398
lemma n_bits_of_eq_iff:
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   399
  "n_bits_of n a = n_bits_of n b \<longleftrightarrow> take_bit n a = take_bit n b"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   400
  apply (induction n arbitrary: a b)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   401
   apply auto
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   402
   apply (metis local.dvd_add_times_triv_left_iff local.dvd_triv_right local.odd_one)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   403
  apply (metis local.dvd_add_times_triv_left_iff local.dvd_triv_right local.odd_one)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   404
  done
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   405
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   406
lemma take_n_bits_of [simp]:
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   407
  "take m (n_bits_of n a) = n_bits_of (min m n) a"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   408
proof -
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   409
  define q and v and w where "q = min m n" and "v = m - q" and "w = n - q"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   410
  then have "v = 0 \<or> w = 0"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   411
    by auto
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   412
  then have "take (q + v) (n_bits_of (q + w) a) = n_bits_of q a"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   413
    by (induction q arbitrary: a) auto
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   414
  with q_def v_def w_def show ?thesis
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   415
    by simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   416
qed
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   417
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   418
lemma unsigned_of_bits_n_bits_of [simp]:
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   419
  "unsigned_of_bits (n_bits_of n a) = take_bit n a"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   420
  by (induction n arbitrary: a) (simp_all add: ac_simps)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   421
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   422
end
70925
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   423
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   424
lemma unsigned_of_bits_eq_of_bits:
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   425
  "unsigned_of_bits bs = (of_bits (bs @ [False]) :: int)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   426
  by (simp add: of_bits_int_def)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   427
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   428
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   429
instantiation word :: (len) bit_representation
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   430
begin
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   431
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   432
lift_definition bits_of_word :: "'a word \<Rightarrow> bool list"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   433
  is "n_bits_of LENGTH('a)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   434
  by (simp add: n_bits_of_eq_iff)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   435
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   436
lift_definition of_bits_word :: "bool list \<Rightarrow> 'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   437
  is unsigned_of_bits .
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   438
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   439
instance proof
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   440
  fix a :: "'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   441
  show "of_bits (bits_of a) = a"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   442
    by transfer simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   443
qed
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   444
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   445
end
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   446
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   447
lemma take_bit_complement_iff:
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   448
  "take_bit n (complement k) = take_bit n (complement l) \<longleftrightarrow> take_bit n k = take_bit n l"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   449
  for k l :: int
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   450
  by (simp add: take_bit_eq_mod mod_eq_dvd_iff dvd_diff_commute)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   451
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   452
lemma take_bit_not_iff:
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   453
  "take_bit n (NOT k) = take_bit n (NOT l) \<longleftrightarrow> take_bit n k = take_bit n l"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   454
  for k l :: int
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   455
  by (simp add: not_int_def take_bit_complement_iff)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   456
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   457
lemma n_bits_of_not:
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   458
  "n_bits_of n (NOT k) = map Not (n_bits_of n k)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   459
  for k :: int
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   460
  by (induction n arbitrary: k) (simp_all add: not_div_2)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   461
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   462
lemma take_bit_and [simp]:
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   463
  "take_bit n (k AND l) = take_bit n k AND take_bit n l"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   464
  for k l :: int
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   465
  apply (induction n arbitrary: k l)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   466
   apply simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   467
  apply (subst and_int.rec)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   468
  apply (subst (2) and_int.rec)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   469
  apply simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   470
  done
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   471
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   472
lemma take_bit_or [simp]:
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   473
  "take_bit n (k OR l) = take_bit n k OR take_bit n l"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   474
  for k l :: int
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   475
  apply (induction n arbitrary: k l)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   476
   apply simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   477
  apply (subst or_int.rec)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   478
  apply (subst (2) or_int.rec)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   479
  apply simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   480
  done
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   481
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   482
lemma take_bit_xor [simp]:
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   483
  "take_bit n (k XOR l) = take_bit n k XOR take_bit n l"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   484
  for k l :: int
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   485
  apply (induction n arbitrary: k l)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   486
   apply simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   487
  apply (subst xor_int.rec)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   488
  apply (subst (2) xor_int.rec)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   489
  apply simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   490
  done
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   491
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   492
instantiation word :: (len) bit_operations
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   493
begin
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   494
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   495
lift_definition not_word :: "'a word \<Rightarrow> 'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   496
  is not
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   497
  by (simp add: take_bit_not_iff)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   498
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   499
lift_definition and_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   500
  is "and"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   501
  by simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   502
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   503
lift_definition or_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   504
  is or
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   505
  by simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   506
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   507
lift_definition xor_word ::  "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   508
  is xor
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   509
  by simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   510
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   511
lift_definition shift_left_word :: "'a word \<Rightarrow> nat \<Rightarrow> 'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   512
  is shift_left
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   513
proof -
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   514
  show "take_bit LENGTH('a) (k << n) = take_bit LENGTH('a) (l << n)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   515
    if "take_bit LENGTH('a) k = take_bit LENGTH('a) l" for k l :: int and n :: nat
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   516
  proof -
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   517
    from that
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   518
    have "take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   519
      = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   520
      by simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   521
    moreover have "min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   522
      by simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   523
    ultimately show ?thesis by (simp add: take_bit_push_bit)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   524
  qed
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   525
qed
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   526
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   527
lift_definition shift_right_word :: "'a word \<Rightarrow> nat \<Rightarrow> 'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   528
  is "\<lambda>k n. drop_bit n (take_bit LENGTH('a) k)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   529
  by simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   530
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   531
instance proof
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   532
  show "semilattice ((AND) :: 'a word \<Rightarrow> _)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   533
    by standard (transfer; simp add: ac_simps)+
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   534
  show "semilattice ((OR) :: 'a word \<Rightarrow> _)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   535
    by standard (transfer; simp add: ac_simps)+
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   536
  show "abel_semigroup ((XOR) :: 'a word \<Rightarrow> _)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   537
    by standard (transfer; simp add: ac_simps)+
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   538
  show "not = (of_bits \<circ> map Not \<circ> bits_of :: 'a word \<Rightarrow> 'a word)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   539
  proof
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   540
    fix a :: "'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   541
    have "NOT a = of_bits (map Not (bits_of a))"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   542
      by transfer (simp flip: unsigned_of_bits_take n_bits_of_not add: take_map)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   543
    then show "NOT a = (of_bits \<circ> map Not \<circ> bits_of) a"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   544
      by simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   545
  qed
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   546
  show "of_bits bs AND of_bits cs = (of_bits (map2 (\<and>) bs cs) :: 'a word)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   547
    if "length bs = length cs" for bs cs
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   548
    using that apply transfer
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   549
    apply (simp only: unsigned_of_bits_eq_of_bits)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   550
    apply (subst and_eq)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   551
    apply simp_all
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   552
    done
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   553
  show "of_bits bs OR of_bits cs = (of_bits (map2 (\<or>) bs cs) :: 'a word)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   554
    if "length bs = length cs" for bs cs
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   555
    using that apply transfer
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   556
    apply (simp only: unsigned_of_bits_eq_of_bits)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   557
    apply (subst or_eq)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   558
    apply simp_all
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   559
    done
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   560
  show "of_bits bs XOR of_bits cs = (of_bits (map2 (\<noteq>) bs cs) :: 'a word)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   561
    if "length bs = length cs" for bs cs
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   562
    using that apply transfer
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   563
    apply (simp only: unsigned_of_bits_eq_of_bits)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   564
    apply (subst xor_eq)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   565
    apply simp_all
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   566
    done
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   567
  show "a << n = of_bits (replicate n False @ bits_of a)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   568
    for a :: "'a word" and n :: nat
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   569
    by transfer (simp add: push_bit_take_bit)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   570
  show "a >> n = of_bits (drop n (bits_of a))"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   571
    if "n < length (bits_of a)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   572
    for a :: "'a word" and n :: nat
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   573
    using that by transfer simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   574
qed
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   575
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   576
end
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   577
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   578
global_interpretation bit_word: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: 'a::len word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   579
  rewrites "bit_word.xor = ((XOR) :: 'a word \<Rightarrow> _)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   580
proof -
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   581
  interpret bit_word: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: 'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   582
  proof
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   583
    show "a AND (b OR c) = a AND b OR a AND c"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   584
      for a b c :: "'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   585
      by transfer (simp add: bit_int.conj_disj_distrib)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   586
    show "a OR b AND c = (a OR b) AND (a OR c)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   587
      for a b c :: "'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   588
      by transfer (simp add: bit_int.disj_conj_distrib)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   589
    show "a AND NOT a = 0" for a :: "'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   590
      by transfer simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   591
    show "a OR NOT a = - 1" for a :: "'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   592
      by transfer simp
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   593
  qed (transfer; simp)+
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   594
  show "boolean_algebra (AND) (OR) NOT 0 (- 1 :: 'a word)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   595
    by (fact bit_word.boolean_algebra_axioms)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   596
  show "bit_word.xor = ((XOR) :: 'a word \<Rightarrow> _)"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   597
  proof (rule ext)+
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   598
    fix a b :: "'a word"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   599
    have "a XOR b = a AND NOT b OR NOT a AND b"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   600
      by transfer (simp add: bit_int.xor_def)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   601
    then show "bit_word.xor a b = a XOR b"
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   602
      by (simp add: bit_word.xor_def)
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   603
  qed
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   604
qed
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   605
525853e4ec80 bit operations for word type
haftmann
parents: 70903
diff changeset
   606
end