| author | wenzelm | 
| Wed, 14 Jan 1998 10:30:44 +0100 | |
| changeset 4571 | 6b02fc8a97f6 | 
| parent 3795 | e687069e7257 | 
| child 4709 | d24168720303 | 
| permissions | -rw-r--r-- | 
| 1632 | 1  | 
(* Title: HOL/Integ/Bin.thy  | 
2  | 
Authors: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
3  | 
David Spelt, University of Twente  | 
|
4  | 
Copyright 1994 University of Cambridge  | 
|
5  | 
Copyright 1996 University of Twente  | 
|
6  | 
||
7  | 
Arithmetic on binary integers.  | 
|
8  | 
||
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
9  | 
The sign PlusSign stands for an infinite string of leading F's.  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
10  | 
The sign MinusSign stands for an infinite string of leading T's.  | 
| 1632 | 11  | 
|
12  | 
A number can have multiple representations, namely leading F's with sign  | 
|
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
13  | 
PlusSign and leading T's with sign MinusSign. See twos-compl.ML/int_of_binary  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
14  | 
for the numerical interpretation.  | 
| 1632 | 15  | 
|
16  | 
The representation expects that (m mod 2) is 0 or 1, even if m is negative;  | 
|
17  | 
For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1  | 
|
18  | 
||
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
19  | 
Division is not defined yet! To do it efficiently requires computing the  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
20  | 
quotient and remainder using ML and checking the answer using multiplication  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
21  | 
by proof. Then uniqueness of the quotient and remainder yields theorems  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
22  | 
quoting the previously computed values. (Or code an oracle...)  | 
| 1632 | 23  | 
*)  | 
24  | 
||
25  | 
Bin = Integ +  | 
|
26  | 
||
27  | 
syntax  | 
|
28  | 
  "_Int"           :: xnum => int        ("_")
 | 
|
29  | 
||
30  | 
datatype  | 
|
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
31  | 
bin = PlusSign  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
32  | 
| MinusSign  | 
| 1632 | 33  | 
| Bcons bin bool  | 
34  | 
||
35  | 
consts  | 
|
36  | 
integ_of_bin :: bin=>int  | 
|
37  | 
norm_Bcons :: [bin,bool]=>bin  | 
|
38  | 
bin_succ :: bin=>bin  | 
|
39  | 
bin_pred :: bin=>bin  | 
|
40  | 
bin_minus :: bin=>bin  | 
|
41  | 
bin_add,bin_mult :: [bin,bin]=>bin  | 
|
42  | 
h_bin :: [bin,bool,bin]=>bin  | 
|
43  | 
||
44  | 
(*norm_Bcons adds a bit, suppressing leading 0s and 1s*)  | 
|
45  | 
||
46  | 
primrec norm_Bcons bin  | 
|
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
47  | 
norm_Plus "norm_Bcons PlusSign b = (if b then (Bcons PlusSign b) else PlusSign)"  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
48  | 
norm_Minus "norm_Bcons MinusSign b = (if b then MinusSign else (Bcons MinusSign b))"  | 
| 1632 | 49  | 
norm_Bcons "norm_Bcons (Bcons w' x') b = Bcons (Bcons w' x') b"  | 
50  | 
||
51  | 
primrec integ_of_bin bin  | 
|
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
52  | 
iob_Plus "integ_of_bin PlusSign = $#0"  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
53  | 
iob_Minus "integ_of_bin MinusSign = $~($#1)"  | 
| 1632 | 54  | 
iob_Bcons "integ_of_bin(Bcons w x) = (if x then $#1 else $#0) + (integ_of_bin w) + (integ_of_bin w)"  | 
55  | 
||
56  | 
primrec bin_succ bin  | 
|
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
57  | 
succ_Plus "bin_succ PlusSign = Bcons PlusSign True"  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
58  | 
succ_Minus "bin_succ MinusSign = PlusSign"  | 
| 1632 | 59  | 
succ_Bcons "bin_succ(Bcons w x) = (if x then (Bcons (bin_succ w) False) else (norm_Bcons w True))"  | 
60  | 
||
61  | 
primrec bin_pred bin  | 
|
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
62  | 
pred_Plus "bin_pred(PlusSign) = MinusSign"  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
63  | 
pred_Minus "bin_pred(MinusSign) = Bcons MinusSign False"  | 
| 1632 | 64  | 
pred_Bcons "bin_pred(Bcons w x) = (if x then (norm_Bcons w False) else (Bcons (bin_pred w) True))"  | 
65  | 
||
66  | 
primrec bin_minus bin  | 
|
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
67  | 
min_Plus "bin_minus PlusSign = PlusSign"  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
68  | 
min_Minus "bin_minus MinusSign = Bcons PlusSign True"  | 
| 1632 | 69  | 
min_Bcons "bin_minus(Bcons w x) = (if x then (bin_pred (Bcons (bin_minus w) False)) else (Bcons (bin_minus w) False))"  | 
70  | 
||
71  | 
primrec bin_add bin  | 
|
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
72  | 
add_Plus "bin_add PlusSign w = w"  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
73  | 
add_Minus "bin_add MinusSign w = bin_pred w"  | 
| 1632 | 74  | 
add_Bcons "bin_add (Bcons v x) w = h_bin v x w"  | 
75  | 
||
76  | 
primrec bin_mult bin  | 
|
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
77  | 
mult_Plus "bin_mult PlusSign w = PlusSign"  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
78  | 
mult_Minus "bin_mult MinusSign w = bin_minus w"  | 
| 1632 | 79  | 
mult_Bcons "bin_mult (Bcons v x) w = (if x then (bin_add (norm_Bcons (bin_mult v w) False) w) else (norm_Bcons (bin_mult v w) False))"  | 
80  | 
||
81  | 
primrec h_bin bin  | 
|
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
82  | 
h_Plus "h_bin v x PlusSign = Bcons v x"  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
83  | 
h_Minus "h_bin v x MinusSign = bin_pred (Bcons v x)"  | 
| 1632 | 84  | 
h_BCons "h_bin v x (Bcons w y) = norm_Bcons (bin_add v (if (x & y) then bin_succ w else w)) (x~=y)"  | 
85  | 
||
86  | 
end  | 
|
87  | 
||
88  | 
ML  | 
|
89  | 
||
90  | 
(** Concrete syntax for integers **)  | 
|
91  | 
||
92  | 
local  | 
|
93  | 
open Syntax;  | 
|
94  | 
||
95  | 
(* Bits *)  | 
|
96  | 
||
97  | 
fun mk_bit 0 = const "False"  | 
|
98  | 
| mk_bit 1 = const "True"  | 
|
99  | 
| mk_bit _ = sys_error "mk_bit";  | 
|
100  | 
||
101  | 
  fun dest_bit (Const ("False", _)) = 0
 | 
|
102  | 
    | dest_bit (Const ("True", _)) = 1
 | 
|
103  | 
| dest_bit _ = raise Match;  | 
|
104  | 
||
105  | 
||
106  | 
(* Bit strings *) (*we try to handle superfluous leading digits nicely*)  | 
|
107  | 
||
108  | 
fun prefix_len _ [] = 0  | 
|
109  | 
| prefix_len pred (x :: xs) =  | 
|
110  | 
if pred x then 1 + prefix_len pred xs else 0;  | 
|
111  | 
||
112  | 
fun mk_bin str =  | 
|
113  | 
let  | 
|
114  | 
val (sign, digs) =  | 
|
115  | 
(case explode str of  | 
|
116  | 
"#" :: "~" :: cs => (~1, cs)  | 
|
117  | 
| "#" :: cs => (1, cs)  | 
|
118  | 
| _ => raise ERROR);  | 
|
119  | 
||
120  | 
val zs = prefix_len (equal "0") digs;  | 
|
121  | 
||
122  | 
fun bin_of 0 = replicate zs 0  | 
|
123  | 
| bin_of ~1 = replicate zs 1 @ [~1]  | 
|
124  | 
| bin_of n = (n mod 2) :: bin_of (n div 2);  | 
|
125  | 
||
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
126  | 
fun term_of [] = const "PlusSign"  | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
127  | 
| term_of [~1] = const "MinusSign"  | 
| 1632 | 128  | 
| term_of (b :: bs) = const "Bcons" $ term_of bs $ mk_bit b;  | 
129  | 
in  | 
|
130  | 
term_of (bin_of (sign * (#1 (scan_int digs))))  | 
|
131  | 
end;  | 
|
132  | 
||
133  | 
fun dest_bin tm =  | 
|
134  | 
let  | 
|
| 
2988
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
135  | 
      fun bin_of (Const ("PlusSign", _)) = []
 | 
| 
 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 
paulson 
parents: 
1632 
diff
changeset
 | 
136  | 
        | bin_of (Const ("MinusSign", _)) = [~1]
 | 
| 1632 | 137  | 
        | bin_of (Const ("Bcons", _) $ bs $ b) = dest_bit b :: bin_of bs
 | 
138  | 
| bin_of _ = raise Match;  | 
|
139  | 
||
140  | 
fun int_of [] = 0  | 
|
141  | 
| int_of (b :: bs) = b + 2 * int_of bs;  | 
|
142  | 
||
143  | 
val rev_digs = bin_of tm;  | 
|
144  | 
val (sign, zs) =  | 
|
145  | 
(case rev rev_digs of  | 
|
146  | 
          ~1 :: bs => ("~", prefix_len (equal 1) bs)
 | 
|
147  | 
        | bs => ("", prefix_len (equal 0) bs));
 | 
|
148  | 
val num = string_of_int (abs (int_of rev_digs));  | 
|
149  | 
in  | 
|
150  | 
"#" ^ sign ^ implode (replicate zs "0") ^ num  | 
|
151  | 
end;  | 
|
152  | 
||
153  | 
||
154  | 
(* translation of integer constant tokens to and from binary *)  | 
|
155  | 
||
156  | 
fun int_tr (*"_Int"*) [t as Free (str, _)] =  | 
|
157  | 
(const "integ_of_bin" $  | 
|
| 3795 | 158  | 
          (mk_bin str handle ERROR => raise TERM ("int_tr", [t])))
 | 
159  | 
    | int_tr (*"_Int"*) ts = raise TERM ("int_tr", ts);
 | 
|
| 1632 | 160  | 
|
161  | 
fun int_tr' (*"integ_of"*) [t] = const "_Int" $ free (dest_bin t)  | 
|
162  | 
| int_tr' (*"integ_of"*) _ = raise Match;  | 
|
163  | 
in  | 
|
164  | 
  val parse_translation = [("_Int", int_tr)];
 | 
|
165  | 
  val print_translation = [("integ_of_bin", int_tr')]; 
 | 
|
166  | 
end;  |