src/HOL/Codegenerator_Test/Generate_Target_Bit_Operations.thy
author wenzelm
Mon, 20 Oct 2025 21:07:53 +0200
changeset 83323 08722f90a439
parent 82580 cf506179fc4c
permissions -rw-r--r--
update to polyml-5.9.2; discontinue obsolete test from d4eaefc626ec: max_index is the largest possible word;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
81714
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
     2
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
     3
section \<open>Test of target-language bit operations\<close>
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
     4
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
     5
theory Generate_Target_Bit_Operations
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
     6
imports
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
     7
  "HOL-Library.Code_Test"
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
     8
begin
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
     9
81999
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents: 81818
diff changeset
    10
context
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents: 81818
diff changeset
    11
  includes bit_operations_syntax
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents: 81818
diff changeset
    12
begin
81714
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    13
81999
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents: 81818
diff changeset
    14
qualified definition computations where
81714
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    15
  \<open>computations = (
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    16
    [bit (473514 :: integer) 5, bit (- 805167 :: integer) 7],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    17
    [473514 AND 767063, - 805167 AND 767063, 473514 AND - 314527, - 805167 AND - 314527 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    18
    [473514 OR 767063, - 805167 OR 767063, 473514 OR - 314527, - 805167 OR - 314527 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    19
    [473514 XOR 767063, - 805167 XOR 767063, 473514 XOR - 314527, - 805167 XOR - 314527 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    20
    [NOT 473513, NOT (- 805167) :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    21
    mask 39 :: integer,
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    22
    [set_bit 15 473514, set_bit 11 (- 805167) :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    23
    [unset_bit 13 473514, unset_bit 12 (- 805167) :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    24
    [flip_bit 15 473514, flip_bit 12 (- 805167) :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    25
    [push_bit 12 473514, push_bit 12 (- 805167) :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    26
    [drop_bit 12 473514, drop_bit 12 (- 805167) :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    27
    [take_bit 12 473514, take_bit 12 (- 805167) :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    28
    [signed_take_bit 12 473514, signed_take_bit 12 (- 805167) :: integer]
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    29
  )\<close>
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    30
81999
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents: 81818
diff changeset
    31
qualified definition results where
81714
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    32
  \<open>results = (
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    33
    [True, True],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    34
    [208898, 242769, 209184, - 839103 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    35
    [1031679, - 280873, - 50197, - 280591 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    36
    [822781, - 523642, - 259381, 558512 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    37
    [- 473514, 805166 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    38
    549755813887 :: integer,
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    39
    [506282, - 803119 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    40
    [465322, - 809263 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    41
    [506282, - 809263 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    42
    [1939513344, - 3297964032 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    43
    [115, - 197 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    44
    [2474, 1745 :: integer],
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    45
    [- 1622, - 2351 :: integer]
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    46
  )\<close>
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    47
81999
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents: 81818
diff changeset
    48
qualified definition check where
81714
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    49
  \<open>check \<longleftrightarrow> computations = results\<close>
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    50
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    51
lemma check
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    52
  by code_simp
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    53
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    54
lemma check
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    55
  by normalization
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    56
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    57
lemma check
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    58
  by eval
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    59
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    60
test_code check in Scala
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    61
81999
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents: 81818
diff changeset
    62
end
81818
1085eb118dc7 restrict check to PolyML
haftmann
parents: 81814
diff changeset
    63
81714
5e3dd01a9eb2 separate theory for tests checking bit operations
haftmann
parents:
diff changeset
    64
end