| author | paulson | 
| Thu, 15 May 1997 12:54:02 +0200 | |
| changeset 3196 | c522bc46aea7 | 
| parent 1461 | 6bcb44e4d6e5 | 
| child 5512 | 4327eec06849 | 
| permissions | -rw-r--r-- | 
| 1461 | 1  | 
(* Title: ZF/ex/twos-compl.ML  | 
| 0 | 2  | 
ID: $Id$  | 
| 1461 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 0 | 4  | 
Copyright 1993 University of Cambridge  | 
5  | 
||
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
6  | 
ML code for Arithmetic on binary integers; the model for theory Bin  | 
| 0 | 7  | 
|
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
8  | 
The sign Plus stands for an infinite string of leading 0s.  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
9  | 
The sign Minus stands for an infinite string of leading 1s.  | 
| 0 | 10  | 
|
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
11  | 
See int_of_binary for the numerical interpretation. A number can have  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
12  | 
multiple representations, namely leading 0s with sign Plus and leading 1s with  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
13  | 
sign Minus. A number is in NORMAL FORM if it has no such extra bits.  | 
| 0 | 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  | 
Still needs division!  | 
|
19  | 
||
20  | 
print_depth 40;  | 
|
21  | 
System.Control.Print.printDepth := 350;  | 
|
22  | 
*)  | 
|
23  | 
||
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
24  | 
infix 5 $$ $$$  | 
| 0 | 25  | 
|
26  | 
(*Recursive datatype of binary integers*)  | 
|
27  | 
datatype bin = Plus | Minus | op $$ of bin * int;  | 
|
28  | 
||
29  | 
(** Conversions between bin and int **)  | 
|
30  | 
||
31  | 
fun int_of_binary Plus = 0  | 
|
32  | 
| int_of_binary Minus = ~1  | 
|
33  | 
| int_of_binary (w$$b) = 2 * int_of_binary w + b;  | 
|
34  | 
||
35  | 
fun binary_of_int 0 = Plus  | 
|
36  | 
| binary_of_int ~1 = Minus  | 
|
37  | 
| binary_of_int n = binary_of_int (n div 2) $$ (n mod 2);  | 
|
38  | 
||
39  | 
(*** Addition ***)  | 
|
40  | 
||
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
41  | 
(*Attach a bit while preserving the normal form. Cases left as default  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
42  | 
are Plus$$$1 and Minus$$$0. *)  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
43  | 
fun Plus $$$ 0 = Plus  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
44  | 
| Minus $$$ 1 = Minus  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
45  | 
| v $$$ x = v$$x;  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
46  | 
|
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
47  | 
(*Successor of an integer, assumed to be in normal form.  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
48  | 
If w$$1 is normal then w is not Minus, so bin_succ(w) $$ 0 is normal.  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
49  | 
But Minus$$0 is normal while Minus$$1 is not.*)  | 
| 0 | 50  | 
fun bin_succ Plus = Plus$$1  | 
51  | 
| bin_succ Minus = Plus  | 
|
52  | 
| bin_succ (w$$1) = bin_succ(w) $$ 0  | 
|
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
53  | 
| bin_succ (w$$0) = w $$$ 1;  | 
| 0 | 54  | 
|
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
55  | 
(*Predecessor of an integer, assumed to be in normal form.  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
56  | 
If w$$0 is normal then w is not Plus, so bin_pred(w) $$ 1 is normal.  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
57  | 
But Plus$$1 is normal while Plus$$0 is not.*)  | 
| 0 | 58  | 
fun bin_pred Plus = Minus  | 
59  | 
| bin_pred Minus = Minus$$0  | 
|
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
60  | 
| bin_pred (w$$1) = w $$$ 0  | 
| 0 | 61  | 
| bin_pred (w$$0) = bin_pred(w) $$ 1;  | 
62  | 
||
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
63  | 
(*Sum of two binary integers in normal form.  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
64  | 
Ensure last $$ preserves normal form! *)  | 
| 0 | 65  | 
fun bin_add (Plus, w) = w  | 
66  | 
| bin_add (Minus, w) = bin_pred w  | 
|
67  | 
| bin_add (v$$x, Plus) = v$$x  | 
|
68  | 
| bin_add (v$$x, Minus) = bin_pred (v$$x)  | 
|
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
69  | 
| bin_add (v$$x, w$$y) =  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
70  | 
bin_add(v, if x+y=2 then bin_succ w else w) $$$ ((x+y) mod 2);  | 
| 0 | 71  | 
|
72  | 
(*** Subtraction ***)  | 
|
73  | 
||
74  | 
(*Unary minus*)  | 
|
75  | 
fun bin_minus Plus = Plus  | 
|
76  | 
| bin_minus Minus = Plus$$1  | 
|
77  | 
| bin_minus (w$$1) = bin_pred (bin_minus(w) $$ 0)  | 
|
78  | 
| bin_minus (w$$0) = bin_minus(w) $$ 0;  | 
|
79  | 
||
80  | 
(*** Multiplication ***)  | 
|
81  | 
||
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
82  | 
(*product of two bins; a factor of 0 might cause leading 0s in result*)  | 
| 0 | 83  | 
fun bin_mult (Plus, _) = Plus  | 
84  | 
| bin_mult (Minus, v) = bin_minus v  | 
|
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
85  | 
| bin_mult (w$$1, v) = bin_add(bin_mult(w,v) $$$ 0, v)  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
86  | 
| bin_mult (w$$0, v) = bin_mult(w,v) $$$ 0;  | 
| 0 | 87  | 
|
88  | 
(*** Testing ***)  | 
|
89  | 
||
90  | 
(*tests addition*)  | 
|
91  | 
fun checksum m n =  | 
|
92  | 
let val wm = binary_of_int m  | 
|
93  | 
and wn = binary_of_int n  | 
|
94  | 
val wsum = bin_add(wm,wn)  | 
|
95  | 
in if m+n = int_of_binary wsum then (wm, wn, wsum, m+n)  | 
|
96  | 
else raise Match  | 
|
97  | 
end;  | 
|
98  | 
||
99  | 
fun bfact n = if n=0 then Plus$$1  | 
|
100  | 
else bin_mult(binary_of_int n, bfact(n-1));  | 
|
101  | 
||
102  | 
(*Examples...  | 
|
103  | 
bfact 5;  | 
|
104  | 
int_of_binary it;  | 
|
105  | 
bfact 69;  | 
|
106  | 
int_of_binary it;  | 
|
107  | 
||
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
108  | 
(*leading zeros??*)  | 
| 0 | 109  | 
bin_add(binary_of_int 1234, binary_of_int ~1234);  | 
110  | 
bin_mult(binary_of_int 1234, Plus);  | 
|
111  | 
||
| 
912
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
112  | 
(*leading ones??*)  | 
| 
 
ed9e0c70a5da
Redefined functions to suppress redundant leading 0s and 1s
 
lcp 
parents: 
0 
diff
changeset
 | 
113  | 
bin_add(binary_of_int 1, binary_of_int ~2);  | 
| 0 | 114  | 
bin_add(binary_of_int 1234, binary_of_int ~1235);  | 
115  | 
*)  |