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