src/HOL/ex/Word_Lsb_Msb.thy
author wenzelm
Mon, 11 Sep 2023 19:30:48 +0200
changeset 78659 b5f3d1051b13
parent 75713 40af1efeadee
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
     1
theory Word_Lsb_Msb
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
     2
  imports "HOL-Library.Word"          
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
     3
begin
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
     4
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
     5
class word = ring_bit_operations +
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
     6
  fixes word_length :: \<open>'a itself \<Rightarrow> nat\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
     7
  assumes word_length_positive [simp]: \<open>0 < word_length TYPE('a)\<close>
75712
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
     8
    and possible_bit_msb: \<open>possible_bit TYPE('a) (word_length TYPE('a) - Suc 0)\<close>
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
     9
    and not_possible_bit_length: \<open>\<not> possible_bit TYPE('a) (word_length TYPE('a))\<close>
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    10
begin
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    11
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    12
lemma word_length_not_0 [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    13
  \<open>word_length TYPE('a) \<noteq> 0\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    14
  using word_length_positive
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    15
  by simp 
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    16
75712
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    17
lemma possible_bit_iff_less_word_length:
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    18
  \<open>possible_bit TYPE('a) n \<longleftrightarrow> n < word_length TYPE('a)\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    19
proof
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    20
  assume \<open>?P\<close>
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    21
  show ?Q
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    22
  proof (rule ccontr)
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    23
    assume \<open>\<not> n < word_length TYPE('a)\<close>
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    24
    then have \<open>word_length TYPE('a) \<le> n\<close>
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    25
      by simp
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    26
    with \<open>?P\<close> have \<open>possible_bit TYPE('a) (word_length TYPE('a))\<close>
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    27
      by (rule possible_bit_less_imp)
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    28
    with not_possible_bit_length show False ..
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    29
  qed
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    30
next
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    31
  assume \<open>?Q\<close>
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    32
  then have \<open>n \<le> word_length TYPE('a) - Suc 0\<close>
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    33
    by simp
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    34
  with possible_bit_msb show ?P
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    35
    by (rule possible_bit_less_imp)
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    36
qed
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    37
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    38
end
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    39
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    40
instantiation word :: (len) word
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    41
begin
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    42
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    43
definition word_length_word :: \<open>'a word itself \<Rightarrow> nat\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    44
  where [simp, code_unfold]: \<open>word_length_word _ = LENGTH('a)\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    45
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    46
instance
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    47
  by standard simp_all
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    48
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    49
end
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    50
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    51
context word
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    52
begin
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    53
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    54
context
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    55
  includes bit_operations_syntax
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    56
begin
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    57
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    58
abbreviation lsb :: \<open>'a \<Rightarrow> bool\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    59
  where \<open>lsb \<equiv> odd\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    60
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    61
definition msb :: \<open>'a \<Rightarrow> bool\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    62
  where \<open>msb w = bit w (word_length TYPE('a) - Suc 0)\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    63
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    64
lemma not_msb_0 [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    65
  \<open>\<not> msb 0\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    66
  by (simp add: msb_def)
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    67
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    68
lemma msb_minus_1 [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    69
  \<open>msb (- 1)\<close>
75712
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    70
  by (simp add: msb_def possible_bit_iff_less_word_length)
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    71
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    72
lemma msb_1_iff [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    73
  \<open>msb 1 \<longleftrightarrow> word_length TYPE('a) = 1\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    74
  by (auto simp add: msb_def bit_simps le_less)
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    75
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    76
lemma msb_minus_iff [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    77
  \<open>msb (- w) \<longleftrightarrow> \<not> msb (w - 1)\<close>
75712
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    78
  by (simp add: msb_def bit_simps possible_bit_iff_less_word_length)
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    79
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    80
lemma msb_not_iff [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    81
  \<open>msb (NOT w) \<longleftrightarrow> \<not> msb w\<close>
75712
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    82
  by (simp add: msb_def bit_simps possible_bit_iff_less_word_length)
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    83
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    84
lemma msb_and_iff [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    85
  \<open>msb (v AND w) \<longleftrightarrow> msb v \<and> msb w\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    86
  by (simp add: msb_def bit_simps)
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    87
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    88
lemma msb_or_iff [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    89
  \<open>msb (v OR w) \<longleftrightarrow> msb v \<or> msb w\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    90
  by (simp add: msb_def bit_simps)
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    91
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    92
lemma msb_xor_iff [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    93
  \<open>msb (v XOR w) \<longleftrightarrow> \<not> (msb v \<longleftrightarrow> msb w)\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    94
  by (simp add: msb_def bit_simps)
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    95
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    96
lemma msb_exp_iff [simp]:                                             
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    97
  \<open>msb (2 ^ n) \<longleftrightarrow> n = word_length TYPE('a) - Suc 0\<close>
75712
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
    98
  by (simp add: msb_def bit_simps possible_bit_iff_less_word_length)
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
    99
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   100
lemma msb_mask_iff [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   101
  \<open>msb (mask n) \<longleftrightarrow> word_length TYPE('a) \<le> n\<close>
75712
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
   102
  by (simp add: msb_def bit_simps less_diff_conv2 Suc_le_eq less_Suc_eq_le possible_bit_iff_less_word_length)
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   103
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   104
lemma msb_set_bit_iff [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   105
  \<open>msb (set_bit n w) \<longleftrightarrow> n = word_length TYPE('a) - Suc 0 \<or> msb w\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   106
  by (simp add: set_bit_eq_or ac_simps)
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   107
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   108
lemma msb_unset_bit_iff [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   109
  \<open>msb (unset_bit n w) \<longleftrightarrow> n \<noteq> word_length TYPE('a) - Suc 0 \<and> msb w\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   110
  by (simp add: unset_bit_eq_and_not ac_simps)
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   111
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   112
lemma msb_flip_bit_iff [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   113
  \<open>msb (flip_bit n w) \<longleftrightarrow> (n \<noteq> word_length TYPE('a) - Suc 0 \<longleftrightarrow> msb w)\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   114
  by (auto simp add: flip_bit_eq_xor)
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   115
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   116
lemma msb_push_bit_iff:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   117
  \<open>msb (push_bit n w) \<longleftrightarrow> n < word_length TYPE('a) \<and> bit w (word_length TYPE('a) - Suc n)\<close>
75712
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
   118
  by (simp add: msb_def bit_simps le_diff_conv2 Suc_le_eq possible_bit_iff_less_word_length)
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   119
75712
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
   120
lemma msb_drop_bit_iff [simp]:
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   121
  \<open>msb (drop_bit n w) \<longleftrightarrow> n = 0 \<and> msb w\<close>
75712
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
   122
  by (cases n)
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
   123
    (auto simp add: msb_def bit_simps possible_bit_iff_less_word_length intro!: impossible_bit)
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   124
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   125
lemma msb_take_bit_iff [simp]:
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   126
  \<open>msb (take_bit n w) \<longleftrightarrow> word_length TYPE('a) \<le> n \<and> msb w\<close>
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   127
  by (simp add: take_bit_eq_mask ac_simps)
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   128
75712
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
   129
lemma msb_signed_take_bit_iff:
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
   130
  \<open>msb (signed_take_bit n w) \<longleftrightarrow> bit w (min n (word_length TYPE('a) - Suc 0))\<close>
ff0aceed8923 Some more proofs.
haftmann
parents: 75655
diff changeset
   131
  by (simp add: msb_def bit_simps possible_bit_iff_less_word_length)
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   132
75713
40af1efeadee More lemmas.
haftmann
parents: 75712
diff changeset
   133
definition signed_drop_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
40af1efeadee More lemmas.
haftmann
parents: 75712
diff changeset
   134
  where \<open>signed_drop_bit n w = drop_bit n w
40af1efeadee More lemmas.
haftmann
parents: 75712
diff changeset
   135
    OR (of_bool (bit w (word_length TYPE('a) - Suc 0)) * NOT (mask (word_length TYPE('a) - Suc n)))\<close>
40af1efeadee More lemmas.
haftmann
parents: 75712
diff changeset
   136
40af1efeadee More lemmas.
haftmann
parents: 75712
diff changeset
   137
lemma msb_signed_drop_bit_iff [simp]:
40af1efeadee More lemmas.
haftmann
parents: 75712
diff changeset
   138
  \<open>msb (signed_drop_bit n w) \<longleftrightarrow> msb w\<close>
40af1efeadee More lemmas.
haftmann
parents: 75712
diff changeset
   139
  by (simp add: signed_drop_bit_def bit_simps not_le not_less)
40af1efeadee More lemmas.
haftmann
parents: 75712
diff changeset
   140
    (simp add: msb_def)
40af1efeadee More lemmas.
haftmann
parents: 75712
diff changeset
   141
75655
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   142
end
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   143
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   144
end
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   145
be0865060346 sketch for word-specific lsb and msb
haftmann
parents:
diff changeset
   146
end