src/HOL/Library/Code_Target_Bit_Shifts.thy
author haftmann
Mon, 27 Jan 2025 13:13:28 +0100
changeset 81986 3863850f4b0e
parent 81985 e23bd621eddb
child 81987 5dc2c98a6f16
permissions -rw-r--r--
clarified scopes
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/Code_Target_Bit_Shifts.thy
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
     3
*)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
     4
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
     5
section \<open>Implementation of bit-shifts on target-language integers by built-in operations\<close>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
     6
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
     7
theory Code_Target_Bit_Shifts
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
     8
imports Main
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
     9
begin
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    10
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    11
context
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    12
begin
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    13
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    14
qualified definition push_bit :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    15
  where \<open>push_bit i k = Bit_Operations.push_bit (nat_of_integer \<bar>i\<bar>) k\<close>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    16
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    17
qualified lemma push_bit_code [code]:
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    18
  \<open>push_bit i k = k * 2 ^ nat_of_integer \<bar>i\<bar>\<close>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    19
  by (simp add: push_bit_def push_bit_eq_mult)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    20
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    21
lemma push_bit_integer_code [code]:
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    22
  \<open>Bit_Operations.push_bit n k = push_bit (of_nat n) k\<close>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    23
  by (simp add: push_bit_def)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    24
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    25
qualified definition drop_bit :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    26
  where \<open>drop_bit i k = Bit_Operations.drop_bit (nat_of_integer \<bar>i\<bar>) k\<close>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    27
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    28
qualified lemma drop_bit_code [code]:
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    29
  \<open>drop_bit i k = k div 2 ^ nat_of_integer \<bar>i\<bar>\<close>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    30
  by (simp add: drop_bit_def drop_bit_eq_div)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    31
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    32
lemma drop_bit_integer_code [code]:
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    33
  \<open>Bit_Operations.drop_bit n k = drop_bit (of_nat n) k\<close>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    34
  by (simp add: drop_bit_def)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    35
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    36
end
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    37
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    38
code_printing code_module Bit_Shifts \<rightharpoonup>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    39
    (SML) \<open>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    40
structure Bit_Shifts : sig
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    41
  type int = IntInf.int
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    42
  val push : int -> int -> int
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    43
  val drop : int -> int -> int
81814
d4eaefc626ec explicit check for (experimentally determined) border value
haftmann
parents: 81762
diff changeset
    44
  val word_max_index : Word.word (*only for validation*)
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    45
end = struct
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    46
81986
3863850f4b0e clarified scopes
haftmann
parents: 81985
diff changeset
    47
open IntInf;
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    48
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    49
fun fold _ [] y = y
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    50
  | fold f (x :: xs) y = fold f xs (f x y);
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    51
81986
3863850f4b0e clarified scopes
haftmann
parents: 81985
diff changeset
    52
fun replicate n x = (if n <= 0 then [] else x :: replicate (n - 1) x);
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    53
81986
3863850f4b0e clarified scopes
haftmann
parents: 81985
diff changeset
    54
val max_index = pow (2, (Word.wordSize - 3)) - 1; (*experimentally determined*)
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    55
81986
3863850f4b0e clarified scopes
haftmann
parents: 81985
diff changeset
    56
val word_of_int = Word.fromLargeInt o toLarge;
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    57
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    58
val word_max_index = word_of_int max_index;
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    59
81986
3863850f4b0e clarified scopes
haftmann
parents: 81985
diff changeset
    60
fun words_of_int k = case divMod (k, max_index)
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    61
  of (b, s) => word_of_int s :: (replicate b word_max_index);
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    62
81986
3863850f4b0e clarified scopes
haftmann
parents: 81985
diff changeset
    63
fun push' i k = << (k, i);
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    64
81986
3863850f4b0e clarified scopes
haftmann
parents: 81985
diff changeset
    65
fun drop' i k = ~>> (k, i);
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    66
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    67
(* The implementations are formally total, though indices >~ max_index will produce heavy computation load *)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    68
81986
3863850f4b0e clarified scopes
haftmann
parents: 81985
diff changeset
    69
fun push i = fold push' (words_of_int (abs i));
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    70
81986
3863850f4b0e clarified scopes
haftmann
parents: 81985
diff changeset
    71
fun drop i = fold drop' (words_of_int (abs i));
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    72
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    73
end;\<close> for constant Code_Target_Bit_Shifts.push_bit Code_Target_Bit_Shifts.drop_bit
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    74
    and (OCaml) \<open>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    75
module Bit_Shifts : sig
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    76
  val push : Z.t -> Z.t -> Z.t
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    77
  val drop : Z.t -> Z.t -> Z.t
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    78
end = struct
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    79
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    80
let rec fold f xs y = match xs with
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    81
  [] -> y
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    82
  | (x :: xs) -> fold f xs (f x y);;
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    83
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    84
let rec replicate n x = (if Z.leq n Z.zero then [] else x :: replicate (Z.pred n) x);;
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    85
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    86
let max_index = Z.of_int max_int;;
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    87
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    88
let splitIndex i = let (b, s) = Z.div_rem i max_index
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    89
  in Z.to_int s :: (replicate b max_int);;
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    90
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    91
let push' i k = Z.shift_left k i;;
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    92
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    93
let drop' i k = Z.shift_right k i;;
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    94
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    95
(* The implementations are formally total, though indices >~ max_index will produce heavy computation load *)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    96
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    97
let push i = fold push' (splitIndex (Z.abs i));;
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    98
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
    99
let drop i = fold drop' (splitIndex (Z.abs i));;
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   100
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   101
end;;
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   102
\<close> for constant Code_Target_Bit_Shifts.push_bit Code_Target_Bit_Shifts.drop_bit
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   103
    and (Haskell) \<open>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   104
module Bit_Shifts where
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   105
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   106
import Prelude (Int, Integer, toInteger, fromInteger, maxBound, divMod, (-), (<=), abs, flip)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   107
import GHC.Bits (Bits)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   108
import Data.Bits (shiftL, shiftR)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   109
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   110
fold :: (a -> b -> b) -> [a] -> b -> b
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   111
fold _ [] y = y
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   112
fold f (x : xs) y = fold f xs (f x y)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   113
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   114
replicate :: Integer -> a -> [a]
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   115
replicate k x = if k <= 0 then [] else x : replicate (k - 1) x
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   116
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   117
maxIndex :: Integer
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   118
maxIndex = toInteger (maxBound :: Int)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   119
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   120
splitIndex :: Integer -> [Int]
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   121
splitIndex i = fromInteger s : replicate (fromInteger b) maxBound
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   122
  where (b, s) = i `divMod` maxIndex
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   123
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   124
{- The implementations are formally total, though indices >~ maxIndex will produce heavy computation load -}
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   125
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   126
push' :: Int -> Int -> Int
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   127
push' i = flip shiftL (abs i)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   128
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   129
push :: Integer -> Integer -> Integer
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   130
push i = fold (flip shiftL) (splitIndex (abs i))
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   131
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   132
drop' :: Int -> Int -> Int
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   133
drop' i = flip shiftR (abs i)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   134
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   135
drop :: Integer -> Integer -> Integer
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   136
drop i = fold (flip shiftR) (splitIndex (abs i))
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   137
\<close> for constant Code_Target_Bit_Shifts.push_bit Code_Target_Bit_Shifts.drop_bit
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   138
    and (Scala) \<open>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   139
object Bit_Shifts {
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   140
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   141
private val maxIndex : BigInt = BigInt(Int.MaxValue);
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   142
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   143
private def replicate[A](i : BigInt, x : A) : List[A] =
81762
8d790d757bfb compatibility with Scala 3
haftmann
parents: 81713
diff changeset
   144
  i <= 0 match {
8d790d757bfb compatibility with Scala 3
haftmann
parents: 81713
diff changeset
   145
    case true => Nil
8d790d757bfb compatibility with Scala 3
haftmann
parents: 81713
diff changeset
   146
    case false => x :: replicate[A](i - 1, x)
8d790d757bfb compatibility with Scala 3
haftmann
parents: 81713
diff changeset
   147
  }
81713
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   148
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   149
private def splitIndex(i : BigInt) : List[Int] = {
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   150
  val (b, s) = i /% maxIndex
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   151
  return s.intValue :: replicate(b, Int.MaxValue)
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   152
}
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   153
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   154
/* The implementations are formally total, though indices >~ maxIndex will produce heavy computation load */
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   155
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   156
def push(i: BigInt, k: BigInt) : BigInt =
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   157
  splitIndex(i).foldLeft(k) { (l, j) => l << j }
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   158
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   159
def drop(i: BigInt, k: BigInt) : BigInt =
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   160
  splitIndex(i).foldLeft(k) { (l, j) => l >> j }
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   161
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   162
}
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   163
\<close> for constant Code_Target_Bit_Shifts.push_bit Code_Target_Bit_Shifts.drop_bit
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   164
| constant Code_Target_Bit_Shifts.push_bit \<rightharpoonup>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   165
    (SML) "Bit'_Shifts.push"
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   166
    and (OCaml) "Bit'_Shifts.push"
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   167
    and (Haskell) "Bit'_Shifts.push"
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   168
    and (Haskell_Quickcheck) "Bit'_Shifts.push'"
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   169
    and (Scala) "Bit'_Shifts.push"
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   170
| constant Code_Target_Bit_Shifts.drop_bit \<rightharpoonup>
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   171
    (SML) "Bit'_Shifts.drop"
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   172
    and (OCaml) "Bit'_Shifts.drop"
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   173
    and (Haskell) "Bit'_Shifts.drop"
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   174
    and (Haskell_Quickcheck) "Bit'_Shifts.drop'"
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   175
    and (Scala) "Bit'_Shifts.drop"
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   176
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   177
code_reserved
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   178
  (SML) Bit_Shifts
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   179
  and (Haskell) Bit_Shifts
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   180
  and (Scala) Bit_Shifts
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   181
378b9d6c52b2 optionally use shift operations on target numerals for efficient execution
haftmann
parents:
diff changeset
   182
end