src/HOL/ex/Bit_Operation_Computations.thy
author nipkow
Tue, 17 Jun 2025 14:11:40 +0200
changeset 82733 8b537e1af2ec
parent 82012 194de6d02827
permissions -rw-r--r--
reinstated intersection of lists as inter_list_set
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
     2
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
     3
section \<open>Tests for simplifying bit operations on ground terms\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
     4
82012
194de6d02827 clarifed terminology
haftmann
parents: 81876
diff changeset
     5
theory Bit_Operation_Computations
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
     6
  imports
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
     7
    "HOL.Bit_Operations"
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
     8
    "HOL-Library.Word"
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
     9
begin
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    10
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    11
unbundle bit_operations_syntax
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    12
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    13
subsection \<open>Generic unsigned computations\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    14
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    15
locale unsigned_computations =
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    16
  fixes type :: \<open>'a::semiring_bit_operations itself\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    17
begin
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    18
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    19
definition computations where
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    20
  \<open>computations = (
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    21
    [bit (473514 :: 'a) 5],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    22
    [473514 AND 767063 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    23
    [473514 OR 767063 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    24
    [473514 XOR 767063 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    25
    mask 11 :: 'a,
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    26
    [set_bit 15 473514 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    27
    [unset_bit 13 473514 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    28
    [flip_bit 15 473514 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    29
    [push_bit 12 473514 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    30
    [drop_bit 12 65344 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    31
    [take_bit 12 473514 :: 'a]
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    32
  )\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    33
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    34
definition results where
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    35
  \<open>results = (
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    36
    [True],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    37
    [208898 :: 'a::semiring_bit_operations],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    38
    [1031679 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    39
    [822781 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    40
    2047 :: 'a,
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    41
    [506282 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    42
    [465322 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    43
    [506282 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    44
    [1939513344 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    45
    [15 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    46
    [2474 :: 'a]
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    47
  )\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    48
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    49
lemmas evaluation_simps = computations_def results_def mask_numeral
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    50
   \<comment> \<open>Expressions like \<term>\<open>mask 42\<close> are deliberately not simplified by default\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    51
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    52
end
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    53
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    54
global_interpretation unsigned_computations_nat: unsigned_computations \<open>TYPE(nat)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    55
  defines unsigned_computations_nat = unsigned_computations_nat.computations
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    56
    and unsigned_results_nat = unsigned_computations_nat.results .
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    57
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    58
lemma \<open>unsigned_computations_nat.computations = unsigned_computations_nat.results\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    59
  by (simp add: unsigned_computations_nat.evaluation_simps)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    60
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    61
global_interpretation unsigned_computations_int: unsigned_computations \<open>TYPE(int)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    62
  defines unsigned_computations_int = unsigned_computations_int.computations
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    63
    and unsigned_results_int = unsigned_computations_int.results .
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    64
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    65
lemma \<open>unsigned_computations_int.computations = unsigned_computations_int.results\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    66
  by (simp add: unsigned_computations_int.evaluation_simps)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    67
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    68
global_interpretation unsigned_computations_word16: unsigned_computations \<open>TYPE(16 word)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    69
  defines unsigned_computations_word16 = unsigned_computations_word16.computations
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    70
    and unsigned_results_word16 = unsigned_computations_word16.results .
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    71
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    72
lemma \<open>unsigned_computations_word16 = unsigned_results_word16\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    73
  by (simp add: unsigned_computations_word16.evaluation_simps)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    74
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    75
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    76
subsection \<open>Generic unsigned computations\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    77
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    78
locale signed_computations =
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    79
  fixes type :: \<open>'a::ring_bit_operations itself\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    80
begin
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    81
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    82
definition computations where
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    83
  \<open>computations = (
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    84
    [bit (- 805167 :: 'a) 7],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    85
    [- 805167 AND 767063, 473514 AND - 314527, - 805167 AND - 314527 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    86
    [- 805167 OR 767063, 473514 OR - 314527, - 805167 OR - 314527 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    87
    [- 805167 XOR 767063, 473514 XOR - 314527, - 805167 XOR - 314527 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    88
    [NOT 473513, NOT (- 805167) :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    89
    [set_bit 11 (- 805167) :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    90
    [unset_bit 12 (- 805167) :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    91
    [flip_bit 12 (- 805167) :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    92
    [push_bit 12 (- 805167) :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    93
    [take_bit 12 (- 805167) :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    94
    [signed_take_bit 12 473514, signed_take_bit 12 (- 805167) :: 'a]
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    95
  )\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    96
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    97
definition results where
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    98
  \<open>results = (
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
    99
    [True],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   100
    [242769, 209184, - 839103 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   101
    [- 280873, - 50197, - 280591 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   102
    [- 523642, - 259381, 558512 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   103
    [- 473514, 805166 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   104
    [- 803119 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   105
    [- 809263 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   106
    [- 809263 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   107
    [- 3297964032 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   108
    [1745 :: 'a],
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   109
    [- 1622, - 2351 :: 'a]
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   110
  )\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   111
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   112
lemmas evaluation_simps = computations_def results_def
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   113
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   114
end
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   115
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   116
global_interpretation signed_computations_int: signed_computations \<open>TYPE(int)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   117
  defines signed_computations_int = signed_computations_int.computations
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   118
    and signed_results_int = signed_computations_int.results .
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   119
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   120
lemma \<open>signed_computations_int.computations = signed_computations_int.results\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   121
  by (simp add: signed_computations_int.evaluation_simps)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   122
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   123
global_interpretation signed_computations_word16: signed_computations \<open>TYPE(16 word)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   124
  defines signed_computations_word16 = signed_computations_word16.computations
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   125
    and signed_results_word16 = signed_computations_word16.results .
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   126
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   127
lemma \<open>signed_computations_word16 = signed_results_word16\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   128
  by (simp add: signed_computations_word16.evaluation_simps)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   129
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   130
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   131
subsection \<open>Special cases on particular type instances\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   132
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   133
lemma
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   134
  \<open>drop_bit 12 (- 17405 :: int) = - 5\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   135
  by simp
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   136
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   137
lemma
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   138
  \<open>signed_drop_bit 12 (- 17405 :: 16 word) = - 5\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   139
  by simp
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   140
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   141
lemma
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   142
  \<open>drop_bit 12 (- 17405 :: 16 word) = 11\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   143
  by simp
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   144
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents:
diff changeset
   145
end