src/HOL/ex/Word_Type.thy
author haftmann
Tue, 16 Apr 2019 19:50:07 +0000
changeset 70171 3173d7878274
parent 67961 9c31678d2139
child 70348 bde161c740ca
permissions -rw-r--r--
tuned
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
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 64593
diff changeset
     9
    "HOL-Library.Type_Length"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    10
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    11
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    12
lemma take_bit_uminus:
70171
haftmann
parents: 67961
diff changeset
    13
  "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
    14
  by (simp add: take_bit_eq_mod mod_minus_eq)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    15
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    16
lemma take_bit_minus:
70171
haftmann
parents: 67961
diff changeset
    17
  "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
    18
  by (simp add: take_bit_eq_mod mod_diff_eq)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    19
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    20
lemma take_bit_nonnegative [simp]:
70171
haftmann
parents: 67961
diff changeset
    21
  "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
    22
  by (simp add: take_bit_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    23
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    24
definition signed_take_bit :: "nat \<Rightarrow> int \<Rightarrow> int"
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    25
  where signed_take_bit_eq_take_bit:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    26
    "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
    27
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    28
lemma signed_take_bit_eq_take_bit':
70171
haftmann
parents: 67961
diff changeset
    29
  "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
    30
  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
    31
  
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    32
lemma signed_take_bit_0 [simp]:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    33
  "signed_take_bit 0 k = - (k mod 2)"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    34
proof (cases "even k")
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    35
  case True
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    36
  then have "odd (k + 1)"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    37
    by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    38
  then have "(k + 1) mod 2 = 1"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    39
    by (simp add: even_iff_mod_2_eq_zero)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    40
  with True show ?thesis
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    41
    by (simp add: signed_take_bit_eq_take_bit)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    42
next
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    43
  case False
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    44
  then show ?thesis
67907
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 odd_iff_mod_2_eq_one)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    46
qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    47
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    48
lemma signed_take_bit_Suc [simp]:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    49
  "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
    50
  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
    51
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    52
lemma signed_take_bit_of_0 [simp]:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    53
  "signed_take_bit n 0 = 0"
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    54
  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
    55
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    56
lemma signed_take_bit_of_minus_1 [simp]:
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    57
  "signed_take_bit n (- 1) = - 1"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    58
  by (induct n) simp_all
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    59
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    60
lemma signed_take_bit_eq_iff_take_bit_eq:
70171
haftmann
parents: 67961
diff changeset
    61
  "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
    62
  if "n > 0"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    63
proof -
70171
haftmann
parents: 67961
diff changeset
    64
  from that obtain m where m: "n = Suc m"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    65
    by (cases n) auto
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    66
  show ?thesis
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    67
  proof 
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    68
    assume ?Q
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    69
    have "take_bit (Suc m) (k + 2 ^ m) =
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    70
      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
    71
      by (simp only: take_bit_add)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    72
    also have "\<dots> =
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    73
      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
    74
      by (simp only: \<open>?Q\<close> m [symmetric])
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    75
    also have "\<dots> = take_bit (Suc m) (l + 2 ^ m)"
67961
9c31678d2139 even more on bit operations
haftmann
parents: 67907
diff changeset
    76
      by (simp only: take_bit_add)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    77
    finally show ?P
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    78
      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
    79
  next
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    80
    assume ?P
70171
haftmann
parents: 67961
diff changeset
    81
    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
    82
      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
    83
    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
    84
      by (metis mod_add_eq)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    85
    then have "k mod 2 ^ n = l mod 2 ^ n"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    86
      by (metis add_diff_cancel_right' uminus_add_conv_diff)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    87
    then show ?Q
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    88
      by (simp add: take_bit_eq_mod)
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    89
  qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    90
qed 
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    91
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    92
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    93
subsection \<open>Bit strings as quotient type\<close>
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
subsubsection \<open>Basic properties\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    96
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
    97
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
    98
  by (auto intro!: equivpI reflpI sympI transpI)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
    99
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   100
instantiation word :: (len0) "{semiring_numeral, comm_semiring_0, comm_ring}"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   101
begin
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
lift_definition zero_word :: "'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   104
  is 0
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
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   107
lift_definition one_word :: "'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   108
  is 1
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
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   111
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
   112
  is plus
67961
9c31678d2139 even more on bit operations
haftmann
parents: 67907
diff changeset
   113
  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
   114
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   115
lift_definition uminus_word :: "'a word \<Rightarrow> 'a word"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   116
  is uminus
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   117
  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
   118
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   119
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
   120
  is minus
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   121
  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
   122
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   123
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
   124
  is times
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   125
  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
   126
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   127
instance
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   128
  by standard (transfer; simp add: algebra_simps)+
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
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   131
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   132
instance word :: (len) comm_ring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   133
  by standard (transfer; simp)+
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
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   136
subsubsection \<open>Conversions\<close>
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   137
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   138
lemma [transfer_rule]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   139
  "rel_fun HOL.eq pcr_word int of_nat"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   140
proof -
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   141
  note transfer_rule_of_nat [transfer_rule]
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   142
  show ?thesis by transfer_prover
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   143
qed
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
lemma [transfer_rule]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   146
  "rel_fun HOL.eq pcr_word (\<lambda>k. k) of_int"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   147
proof -
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   148
  note transfer_rule_of_int [transfer_rule]
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   149
  have "rel_fun HOL.eq pcr_word (of_int :: int \<Rightarrow> int) (of_int :: int \<Rightarrow> 'a word)"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   150
    by transfer_prover
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   151
  then show ?thesis by (simp add: id_def)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   152
qed
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   153
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   154
context semiring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   155
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   156
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   157
lift_definition unsigned :: "'b::len0 word \<Rightarrow> 'a"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   158
  is "of_nat \<circ> nat \<circ> take_bit LENGTH('b)"
64015
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   159
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   160
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   161
lemma unsigned_0 [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   162
  "unsigned 0 = 0"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   163
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   164
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   165
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   166
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   167
context semiring_char_0
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   168
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   169
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   170
lemma word_eq_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   171
  "a = b \<longleftrightarrow> unsigned a = unsigned b"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   172
  by safe (transfer; simp add: eq_nat_nat_iff)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   173
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   174
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   175
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   176
context ring_1
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   177
begin
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
lift_definition signed :: "'b::len word \<Rightarrow> 'a"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   180
  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
   181
  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
   182
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   183
lemma signed_0 [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   184
  "signed 0 = 0"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   185
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   186
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   187
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   188
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   189
lemma unsigned_of_nat [simp]:
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   190
  "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
   191
  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
   192
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   193
lemma of_nat_unsigned [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   194
  "of_nat (unsigned a) = a"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   195
  by transfer simp
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 of_int_unsigned [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   198
  "of_int (unsigned a) = a"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   199
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   200
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   201
context ring_char_0
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   202
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   203
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   204
lemma word_eq_iff_signed:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   205
  "a = b \<longleftrightarrow> signed a = signed b"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   206
  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
   207
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   208
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   209
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   210
lemma signed_of_int [simp]:
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   211
  "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
   212
  by transfer simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   213
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   214
lemma of_int_signed [simp]:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   215
  "of_int (signed a) = a"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   216
  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
   217
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
subsubsection \<open>Properties\<close>
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
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   222
subsubsection \<open>Division\<close>
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
instantiation word :: (len0) modulo
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   225
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   226
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   227
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
   228
  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
   229
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   230
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   231
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
   232
  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
   233
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   234
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   235
instance ..
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   236
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   237
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   238
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
subsubsection \<open>Orderings\<close>
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
instantiation word :: (len0) linorder
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   243
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   244
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   245
lift_definition 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
   246
  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
   247
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   248
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   249
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
   250
  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
   251
  by simp
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   252
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   253
instance
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   254
  by standard (transfer; auto)+
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
end
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   257
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   258
context linordered_semidom
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   259
begin
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   260
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   261
lemma word_less_eq_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   262
  "a \<le> b \<longleftrightarrow> unsigned a \<le> unsigned b"
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   263
  by (transfer fixing: less_eq) (simp add: nat_le_eq_zle)
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   264
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   265
lemma word_less_iff_unsigned:
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   266
  "a < b \<longleftrightarrow> unsigned a < unsigned b"
67907
02a14c1cb917 prefer convention to place operation name before type name
haftmann
parents: 67816
diff changeset
   267
  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
   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
c9f3a94cb825 proof of concept for algebraically founded word types
haftmann
parents:
diff changeset
   271
end