| 
905
 | 
     1  | 
(*  Title:      ZF/ex/Bin.thy
  | 
| 
515
 | 
     2  | 
    ID:         $Id$
  | 
| 
905
 | 
     3  | 
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
  | 
| 
515
 | 
     4  | 
    Copyright   1994  University of Cambridge
  | 
| 
 | 
     5  | 
  | 
| 
 | 
     6  | 
Arithmetic on binary integers.
  | 
| 
905
 | 
     7  | 
  | 
| 
 | 
     8  | 
   The sign Plus stands for an infinite string of leading 0's.
  | 
| 
 | 
     9  | 
   The sign Minus stands for an infinite string of leading 1's.
  | 
| 
 | 
    10  | 
  | 
| 
 | 
    11  | 
A number can have multiple representations, namely leading 0's with sign
  | 
| 
 | 
    12  | 
Plus and leading 1's with sign Minus.  See twos-compl.ML/int_of_binary for
  | 
| 
 | 
    13  | 
the numerical interpretation.
  | 
| 
 | 
    14  | 
  | 
| 
 | 
    15  | 
The representation expects that (m mod 2) is 0 or 1, even if m is negative;
  | 
| 
 | 
    16  | 
For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1
  | 
| 
 | 
    17  | 
  | 
| 
 | 
    18  | 
Division is not defined yet!
  | 
| 
515
 | 
    19  | 
*)
  | 
| 
 | 
    20  | 
  | 
| 
935
 | 
    21  | 
Bin = Integ + Datatype + "twos_compl" +
  | 
| 
905
 | 
    22  | 
  | 
| 
515
 | 
    23  | 
consts
  | 
| 
1401
 | 
    24  | 
  bin_rec          :: [i, i, i, [i,i,i]=>i] => i
  | 
| 
 | 
    25  | 
  integ_of_bin     :: i=>i
  | 
| 
 | 
    26  | 
  norm_Bcons       :: [i,i]=>i
  | 
| 
 | 
    27  | 
  bin_succ         :: i=>i
  | 
| 
 | 
    28  | 
  bin_pred         :: i=>i
  | 
| 
 | 
    29  | 
  bin_minus        :: i=>i
  | 
| 
 | 
    30  | 
  bin_add,bin_mult :: [i,i]=>i
  | 
| 
515
 | 
    31  | 
  | 
| 
1401
 | 
    32  | 
  bin              :: i
  | 
| 
905
 | 
    33  | 
  | 
| 
 | 
    34  | 
syntax
  | 
| 
1401
 | 
    35  | 
  "_Int"           :: xnum => i        ("_")
 | 
| 
515
 | 
    36  | 
  | 
| 
 | 
    37  | 
datatype
  | 
| 
 | 
    38  | 
  "bin" = Plus
  | 
| 
 | 
    39  | 
        | Minus
  | 
| 
905
 | 
    40  | 
        | Bcons ("w: bin", "b: bool")
 | 
| 
515
 | 
    41  | 
  type_intrs "[bool_into_univ]"
  | 
| 
905
 | 
    42  | 
  | 
| 
 | 
    43  | 
  | 
| 
753
 | 
    44  | 
defs
  | 
| 
515
 | 
    45  | 
  | 
| 
 | 
    46  | 
  bin_rec_def
  | 
| 
1155
 | 
    47  | 
      "bin_rec(z,a,b,h) == 
  | 
| 
 | 
    48  | 
      Vrec(z, %z g. bin_case(a, b, %w x. h(w, x, g`w), z))"
  | 
| 
515
 | 
    49  | 
  | 
| 
905
 | 
    50  | 
  integ_of_bin_def
  | 
| 
515
 | 
    51  | 
      "integ_of_bin(w) == bin_rec(w, $#0, $~($#1), %w x r. $#x $+ r $+ r)"
  | 
| 
 | 
    52  | 
  | 
| 
905
 | 
    53  | 
    (** recall that cond(1,b,c)=b and cond(0,b,c)=0 **)
  | 
| 
 | 
    54  | 
  | 
| 
 | 
    55  | 
  (*norm_Bcons adds a bit, suppressing leading 0s and 1s*)
  | 
| 
 | 
    56  | 
  norm_Bcons_def
  | 
| 
1155
 | 
    57  | 
      "norm_Bcons(w,b) ==   
  | 
| 
 | 
    58  | 
       bin_case(cond(b,Bcons(w,b),w), cond(b,w,Bcons(w,b)),   
  | 
| 
 | 
    59  | 
                %w' x'. Bcons(w,b), w)"
  | 
| 
905
 | 
    60  | 
  | 
| 
515
 | 
    61  | 
  bin_succ_def
  | 
| 
1155
 | 
    62  | 
      "bin_succ(w0) ==   
  | 
| 
 | 
    63  | 
       bin_rec(w0, Bcons(Plus,1), Plus,   
  | 
| 
 | 
    64  | 
               %w x r. cond(x, Bcons(r,0), norm_Bcons(w,1)))"
  | 
| 
515
 | 
    65  | 
  | 
| 
 | 
    66  | 
  bin_pred_def
  | 
| 
1155
 | 
    67  | 
      "bin_pred(w0) == 
  | 
| 
 | 
    68  | 
       bin_rec(w0, Minus, Bcons(Minus,0),   
  | 
| 
 | 
    69  | 
               %w x r. cond(x, norm_Bcons(w,0), Bcons(r,1)))"
  | 
| 
515
 | 
    70  | 
  | 
| 
 | 
    71  | 
  bin_minus_def
  | 
| 
1155
 | 
    72  | 
      "bin_minus(w0) == 
  | 
| 
 | 
    73  | 
       bin_rec(w0, Plus, Bcons(Plus,1),   
  | 
| 
 | 
    74  | 
               %w x r. cond(x, bin_pred(Bcons(r,0)), Bcons(r,0)))"
  | 
| 
515
 | 
    75  | 
  | 
| 
 | 
    76  | 
  bin_add_def
  | 
| 
1155
 | 
    77  | 
      "bin_add(v0,w0) ==                        
  | 
| 
 | 
    78  | 
       bin_rec(v0,                             
  | 
| 
 | 
    79  | 
         lam w:bin. w,                 
  | 
| 
 | 
    80  | 
         lam w:bin. bin_pred(w),       
  | 
| 
 | 
    81  | 
         %v x r. lam w1:bin.           
  | 
| 
 | 
    82  | 
                  bin_rec(w1, Bcons(v,x), bin_pred(Bcons(v,x)),    
  | 
| 
 | 
    83  | 
                    %w y s. norm_Bcons(r`cond(x and y, bin_succ(w), w), 
  | 
| 
 | 
    84  | 
                                          x xor y)))    ` w0"
  | 
| 
515
 | 
    85  | 
  | 
| 
 | 
    86  | 
  bin_mult_def
  | 
| 
1155
 | 
    87  | 
      "bin_mult(v0,w) ==                        
  | 
| 
 | 
    88  | 
       bin_rec(v0, Plus, bin_minus(w),         
  | 
| 
 | 
    89  | 
         %v x r. cond(x, bin_add(norm_Bcons(r,0),w), norm_Bcons(r,0)))"
  | 
| 
515
 | 
    90  | 
end
  | 
| 
905
 | 
    91  | 
  | 
| 
 | 
    92  | 
  | 
| 
 | 
    93  | 
ML
  | 
| 
 | 
    94  | 
  | 
| 
 | 
    95  | 
(** Concrete syntax for integers **)
  | 
| 
 | 
    96  | 
  | 
| 
 | 
    97  | 
local
  | 
| 
 | 
    98  | 
  open Syntax;
  | 
| 
 | 
    99  | 
  | 
| 
 | 
   100  | 
  (* Bits *)
  | 
| 
 | 
   101  | 
  | 
| 
 | 
   102  | 
  fun mk_bit 0 = const "0"
  | 
| 
 | 
   103  | 
    | mk_bit 1 = const "succ" $ const "0"
  | 
| 
 | 
   104  | 
    | mk_bit _ = sys_error "mk_bit";
  | 
| 
 | 
   105  | 
  | 
| 
 | 
   106  | 
  fun dest_bit (Const ("0", _)) = 0
 | 
| 
 | 
   107  | 
    | dest_bit (Const ("succ", _) $ Const ("0", _)) = 1
 | 
| 
 | 
   108  | 
    | dest_bit _ = raise Match;
  | 
| 
 | 
   109  | 
  | 
| 
 | 
   110  | 
  | 
| 
 | 
   111  | 
  (* Bit strings *)   (*we try to handle superfluous leading digits nicely*)
  | 
| 
 | 
   112  | 
  | 
| 
 | 
   113  | 
  fun prefix_len _ [] = 0
  | 
| 
 | 
   114  | 
    | prefix_len pred (x :: xs) =
  | 
| 
 | 
   115  | 
        if pred x then 1 + prefix_len pred xs else 0;
  | 
| 
 | 
   116  | 
  | 
| 
 | 
   117  | 
  fun mk_bin str =
  | 
| 
 | 
   118  | 
    let
  | 
| 
 | 
   119  | 
      val (sign, digs) =
  | 
| 
 | 
   120  | 
        (case explode str of
  | 
| 
 | 
   121  | 
          "#" :: "~" :: cs => (~1, cs)
  | 
| 
 | 
   122  | 
        | "#" :: cs => (1, cs)
  | 
| 
 | 
   123  | 
        | _ => raise ERROR);
  | 
| 
 | 
   124  | 
  | 
| 
 | 
   125  | 
      val zs = prefix_len (equal "0") digs;
  | 
| 
 | 
   126  | 
  | 
| 
 | 
   127  | 
      fun bin_of 0 = replicate zs 0
  | 
| 
 | 
   128  | 
        | bin_of ~1 = replicate zs 1 @ [~1]
  | 
| 
 | 
   129  | 
        | bin_of n = (n mod 2) :: bin_of (n div 2);
  | 
| 
 | 
   130  | 
  | 
| 
 | 
   131  | 
      fun term_of [] = const "Plus"
  | 
| 
 | 
   132  | 
        | term_of [~1] = const "Minus"
  | 
| 
 | 
   133  | 
        | term_of (b :: bs) = const "Bcons" $ term_of bs $ mk_bit b;
  | 
| 
 | 
   134  | 
    in
  | 
| 
 | 
   135  | 
      term_of (bin_of (sign * (#1 (scan_int digs))))
  | 
| 
 | 
   136  | 
    end;
  | 
| 
 | 
   137  | 
  | 
| 
 | 
   138  | 
  fun dest_bin tm =
  | 
| 
 | 
   139  | 
    let
  | 
| 
 | 
   140  | 
      fun bin_of (Const ("Plus", _)) = []
 | 
| 
 | 
   141  | 
        | bin_of (Const ("Minus", _)) = [~1]
 | 
| 
 | 
   142  | 
        | bin_of (Const ("Bcons", _) $ bs $ b) = dest_bit b :: bin_of bs
 | 
| 
 | 
   143  | 
        | bin_of _ = raise Match;
  | 
| 
 | 
   144  | 
  | 
| 
 | 
   145  | 
      fun int_of [] = 0
  | 
| 
 | 
   146  | 
        | int_of (b :: bs) = b + 2 * int_of bs;
  | 
| 
 | 
   147  | 
  | 
| 
 | 
   148  | 
      val rev_digs = bin_of tm;
  | 
| 
 | 
   149  | 
      val (sign, zs) =
  | 
| 
 | 
   150  | 
        (case rev rev_digs of
  | 
| 
 | 
   151  | 
          ~1 :: bs => ("~", prefix_len (equal 1) bs)
 | 
| 
 | 
   152  | 
        | bs => ("", prefix_len (equal 0) bs));
 | 
| 
 | 
   153  | 
      val num = string_of_int (abs (int_of rev_digs));
  | 
| 
 | 
   154  | 
    in
  | 
| 
 | 
   155  | 
      "#" ^ sign ^ implode (replicate zs "0") ^ num
  | 
| 
 | 
   156  | 
    end;
  | 
| 
 | 
   157  | 
  | 
| 
 | 
   158  | 
  | 
| 
 | 
   159  | 
  (* translation of integer constant tokens to and from binary *)
  | 
| 
 | 
   160  | 
  | 
| 
 | 
   161  | 
  fun int_tr (*"_Int"*) [t as Free (str, _)] =
  | 
| 
 | 
   162  | 
        (const "integ_of_bin" $
  | 
| 
3795
 | 
   163  | 
          (mk_bin str handle ERROR => raise TERM ("int_tr", [t])))
 | 
| 
 | 
   164  | 
    | int_tr (*"_Int"*) ts = raise TERM ("int_tr", ts);
 | 
| 
905
 | 
   165  | 
  | 
| 
 | 
   166  | 
  fun int_tr' (*"integ_of"*) [t] = const "_Int" $ free (dest_bin t)
  | 
| 
 | 
   167  | 
    | int_tr' (*"integ_of"*) _ = raise Match;
  | 
| 
 | 
   168  | 
in
  | 
| 
 | 
   169  | 
  val parse_translation = [("_Int", int_tr)];
 | 
| 
 | 
   170  | 
  val print_translation = [("integ_of_bin", int_tr')];
 | 
| 
 | 
   171  | 
end;
  |