| author | quigley | 
| Mon, 20 Jun 2005 18:39:24 +0200 | |
| changeset 16478 | d0a1f6231e2f | 
| parent 15965 | f422f8283491 | 
| child 16835 | 2e7d7ec7a268 | 
| permissions | -rw-r--r-- | 
| 923 | 1  | 
(* Title: HOL/hologic.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson and Markus Wenzel  | 
|
4  | 
||
5  | 
Abstract syntax operations for HOL.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature HOLOGIC =  | 
|
9  | 
sig  | 
|
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
10  | 
val typeS: sort  | 
| 
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
11  | 
val typeT: typ  | 
| 
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
12  | 
val read_cterm: Sign.sg -> string -> cterm  | 
| 8275 | 13  | 
val boolN: string  | 
| 923 | 14  | 
val boolT: typ  | 
| 7073 | 15  | 
val false_const: term  | 
16  | 
val true_const: term  | 
|
| 9856 | 17  | 
val not_const: term  | 
| 923 | 18  | 
val mk_setT: typ -> typ  | 
19  | 
val dest_setT: typ -> typ  | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
20  | 
val Trueprop: term  | 
| 923 | 21  | 
val mk_Trueprop: term -> term  | 
22  | 
val dest_Trueprop: term -> term  | 
|
23  | 
val conj: term  | 
|
24  | 
val disj: term  | 
|
25  | 
val imp: term  | 
|
| 8429 | 26  | 
val Not: term  | 
| 7690 | 27  | 
val mk_conj: term * term -> term  | 
28  | 
val mk_disj: term * term -> term  | 
|
29  | 
val mk_imp: term * term -> term  | 
|
| 15151 | 30  | 
val dest_conj: term -> term list  | 
| 15945 | 31  | 
val dest_disj: term -> term list  | 
| 4571 | 32  | 
val dest_imp: term -> term * term  | 
| 15151 | 33  | 
val dest_not: term -> term  | 
| 11683 | 34  | 
val dest_concls: term -> term list  | 
| 923 | 35  | 
val eq_const: typ -> term  | 
36  | 
val all_const: typ -> term  | 
|
37  | 
val exists_const: typ -> term  | 
|
| 15945 | 38  | 
val choice_const: typ -> term  | 
| 923 | 39  | 
val Collect_const: typ -> term  | 
40  | 
val mk_eq: term * term -> term  | 
|
| 6031 | 41  | 
val dest_eq: term -> term * term  | 
| 923 | 42  | 
val mk_all: string * typ * term -> term  | 
| 13640 | 43  | 
val list_all: (string * typ) list * term -> term  | 
| 923 | 44  | 
val mk_exists: string * typ * term -> term  | 
45  | 
val mk_Collect: string * typ * term -> term  | 
|
46  | 
val mk_mem: term * term -> term  | 
|
| 6380 | 47  | 
val dest_mem: term -> term * term  | 
| 11818 | 48  | 
val mk_UNIV: typ -> term  | 
| 2510 | 49  | 
val mk_binop: string -> term * term -> term  | 
50  | 
val mk_binrel: string -> term * term -> term  | 
|
51  | 
val dest_bin: string -> typ -> term -> term * term  | 
|
| 4571 | 52  | 
val unitT: typ  | 
| 9362 | 53  | 
val is_unitT: typ -> bool  | 
| 4571 | 54  | 
val unit: term  | 
55  | 
val is_unit: term -> bool  | 
|
56  | 
val mk_prodT: typ * typ -> typ  | 
|
57  | 
val dest_prodT: typ -> typ * typ  | 
|
| 14048 | 58  | 
val pair_const: typ -> typ -> term  | 
| 4571 | 59  | 
val mk_prod: term * term -> term  | 
60  | 
val dest_prod: term -> term * term  | 
|
61  | 
val mk_fst: term -> term  | 
|
62  | 
val mk_snd: term -> term  | 
|
| 
5096
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
63  | 
val prodT_factors: typ -> typ list  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
64  | 
val split_const: typ * typ * typ -> term  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
65  | 
val mk_tuple: typ -> term list -> term  | 
| 5207 | 66  | 
val natT: typ  | 
67  | 
val zero: term  | 
|
68  | 
val is_zero: term -> bool  | 
|
69  | 
val mk_Suc: term -> term  | 
|
70  | 
val dest_Suc: term -> term  | 
|
71  | 
val mk_nat: int -> term  | 
|
72  | 
val dest_nat: term -> int  | 
|
| 7073 | 73  | 
val intT: typ  | 
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
74  | 
val mk_int: IntInf.int -> term  | 
| 7163 | 75  | 
val realT: typ  | 
| 
15620
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
76  | 
val bitT: typ  | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
77  | 
val B0_const: term  | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
78  | 
val B1_const: term  | 
| 7073 | 79  | 
val binT: typ  | 
80  | 
val pls_const: term  | 
|
81  | 
val min_const: term  | 
|
82  | 
val bit_const: term  | 
|
| 8768 | 83  | 
val number_of_const: typ -> term  | 
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
84  | 
val int_of: int list -> IntInf.int  | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
85  | 
val dest_binum: term -> IntInf.int  | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
86  | 
val mk_bin: IntInf.int -> term  | 
| 
15620
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
87  | 
val bin_of : term -> int list  | 
| 13755 | 88  | 
  val mk_list: ('a -> term) -> typ -> 'a list -> term
 | 
| 15062 | 89  | 
val dest_list: term -> term list  | 
| 923 | 90  | 
end;  | 
91  | 
||
| 2510 | 92  | 
|
| 923 | 93  | 
structure HOLogic: HOLOGIC =  | 
94  | 
struct  | 
|
95  | 
||
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
96  | 
(* HOL syntax *)  | 
| 923 | 97  | 
|
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
98  | 
val typeS: sort = ["HOL.type"];  | 
| 
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
99  | 
val typeT = TypeInfer.anyT typeS;  | 
| 
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
100  | 
|
| 
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
101  | 
fun read_cterm sg s = Thm.read_cterm sg (s, typeT);  | 
| 923 | 102  | 
|
103  | 
||
| 2510 | 104  | 
(* bool and set *)  | 
| 923 | 105  | 
|
| 8275 | 106  | 
val boolN = "bool";  | 
107  | 
val boolT = Type (boolN, []);  | 
|
| 923 | 108  | 
|
| 9856 | 109  | 
val true_const =  Const ("True", boolT);
 | 
110  | 
val false_const = Const ("False", boolT);
 | 
|
111  | 
val not_const = Const ("Not", boolT --> boolT);
 | 
|
| 7073 | 112  | 
|
| 923 | 113  | 
fun mk_setT T = Type ("set", [T]);
 | 
114  | 
||
115  | 
fun dest_setT (Type ("set", [T])) = T
 | 
|
| 3794 | 116  | 
  | dest_setT T = raise TYPE ("dest_setT: set type expected", [T], []);
 | 
| 923 | 117  | 
|
| 8275 | 118  | 
|
| 7073 | 119  | 
(* logic *)  | 
| 923 | 120  | 
|
121  | 
val Trueprop = Const ("Trueprop", boolT --> propT);
 | 
|
122  | 
||
123  | 
fun mk_Trueprop P = Trueprop $ P;  | 
|
124  | 
||
125  | 
fun dest_Trueprop (Const ("Trueprop", _) $ P) = P
 | 
|
| 3794 | 126  | 
  | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]);
 | 
| 923 | 127  | 
|
128  | 
||
129  | 
val conj = Const ("op &", [boolT, boolT] ---> boolT)
 | 
|
130  | 
and disj = Const ("op |", [boolT, boolT] ---> boolT)
 | 
|
| 8429 | 131  | 
and imp = Const ("op -->", [boolT, boolT] ---> boolT)
 | 
132  | 
and Not = Const ("Not", boolT --> boolT);
 | 
|
| 923 | 133  | 
|
| 7690 | 134  | 
fun mk_conj (t1, t2) = conj $ t1 $ t2  | 
135  | 
and mk_disj (t1, t2) = disj $ t1 $ t2  | 
|
136  | 
and mk_imp (t1, t2) = imp $ t1 $ t2;  | 
|
137  | 
||
| 15151 | 138  | 
fun dest_conj (Const ("op &", _) $ t $ t') = t :: dest_conj t'
 | 
139  | 
| dest_conj t = [t];  | 
|
140  | 
||
| 15945 | 141  | 
fun dest_disj (Const ("op |", _) $ t $ t') = t :: dest_disj t'
 | 
142  | 
| dest_disj t = [t];  | 
|
143  | 
||
| 
4466
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4294 
diff
changeset
 | 
144  | 
fun dest_imp (Const("op -->",_) $ A $ B) = (A, B)
 | 
| 
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4294 
diff
changeset
 | 
145  | 
  | dest_imp  t = raise TERM ("dest_imp", [t]);
 | 
| 
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4294 
diff
changeset
 | 
146  | 
|
| 15151 | 147  | 
fun dest_not (Const ("Not", _) $ t) = t
 | 
148  | 
  | dest_not t = raise TERM ("dest_not", [t]);
 | 
|
| 8302 | 149  | 
|
| 11683 | 150  | 
fun imp_concl_of t = imp_concl_of (#2 (dest_imp t)) handle TERM _ => t;  | 
151  | 
val dest_concls = map imp_concl_of o dest_conj o dest_Trueprop;  | 
|
152  | 
||
| 923 | 153  | 
fun eq_const T = Const ("op =", [T, T] ---> boolT);
 | 
154  | 
fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u;  | 
|
155  | 
||
| 6031 | 156  | 
fun dest_eq (Const ("op =", _) $ lhs $ rhs) = (lhs, rhs)
 | 
157  | 
  | dest_eq t = raise TERM ("dest_eq", [t])
 | 
|
158  | 
||
| 923 | 159  | 
fun all_const T = Const ("All", [T --> boolT] ---> boolT);
 | 
160  | 
fun mk_all (x, T, P) = all_const T $ absfree (x, T, P);  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
161  | 
fun list_all (vs,x) = foldr (fn ((x, T), P) => all_const T $ Abs (x, T, P)) x vs;  | 
| 923 | 162  | 
|
163  | 
fun exists_const T = Const ("Ex", [T --> boolT] ---> boolT);
 | 
|
164  | 
fun mk_exists (x, T, P) = exists_const T $ absfree (x, T, P);  | 
|
165  | 
||
| 15945 | 166  | 
fun choice_const T = Const("Hilbert_Choice.Eps", (T --> boolT) --> T)
 | 
167  | 
||
| 923 | 168  | 
fun Collect_const T = Const ("Collect", [T --> boolT] ---> mk_setT T);
 | 
169  | 
fun mk_Collect (a, T, t) = Collect_const T $ absfree (a, T, t);  | 
|
170  | 
||
171  | 
fun mk_mem (x, A) =  | 
|
172  | 
let val setT = fastype_of A in  | 
|
173  | 
    Const ("op :", [dest_setT setT, setT] ---> boolT) $ x $ A
 | 
|
174  | 
end;  | 
|
175  | 
||
| 6380 | 176  | 
fun dest_mem (Const ("op :", _) $ x $ A) = (x, A)
 | 
177  | 
  | dest_mem t = raise TERM ("dest_mem", [t]);
 | 
|
178  | 
||
| 11818 | 179  | 
fun mk_UNIV T = Const ("UNIV", mk_setT T);
 | 
180  | 
||
| 923 | 181  | 
|
| 
13743
 
f8f9393be64c
Fixed bug in simpdata.ML that prevented the use of congruence rules from a
 
ballarin 
parents: 
13640 
diff
changeset
 | 
182  | 
(* binary operations and relations *)  | 
| 2510 | 183  | 
|
184  | 
fun mk_binop c (t, u) =  | 
|
185  | 
let val T = fastype_of t in  | 
|
186  | 
Const (c, [T, T] ---> T) $ t $ u  | 
|
187  | 
end;  | 
|
188  | 
||
189  | 
fun mk_binrel c (t, u) =  | 
|
190  | 
let val T = fastype_of t in  | 
|
191  | 
Const (c, [T, T] ---> boolT) $ t $ u  | 
|
192  | 
end;  | 
|
193  | 
||
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14103 
diff
changeset
 | 
194  | 
(*destruct the application of a binary operator. The dummyT case is a crude  | 
| 
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14103 
diff
changeset
 | 
195  | 
way of handling polymorphic operators.*)  | 
| 2510 | 196  | 
fun dest_bin c T (tm as Const (c', Type ("fun", [T', _])) $ t $ u) =
 | 
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14103 
diff
changeset
 | 
197  | 
if c = c' andalso (T=T' orelse T=dummyT) then (t, u)  | 
| 3794 | 198  | 
      else raise TERM ("dest_bin " ^ c, [tm])
 | 
199  | 
  | dest_bin c _ tm = raise TERM ("dest_bin " ^ c, [tm]);
 | 
|
| 2510 | 200  | 
|
201  | 
||
| 4571 | 202  | 
(* unit *)  | 
203  | 
||
| 11604 | 204  | 
val unitT = Type ("Product_Type.unit", []);
 | 
| 4571 | 205  | 
|
| 11604 | 206  | 
fun is_unitT (Type ("Product_Type.unit", [])) = true
 | 
| 9362 | 207  | 
| is_unitT _ = false;  | 
208  | 
||
| 11604 | 209  | 
val unit = Const ("Product_Type.Unity", unitT);
 | 
| 4571 | 210  | 
|
| 11604 | 211  | 
fun is_unit (Const ("Product_Type.Unity", _)) = true
 | 
| 4571 | 212  | 
| is_unit _ = false;  | 
213  | 
||
214  | 
||
215  | 
(* prod *)  | 
|
216  | 
||
217  | 
fun mk_prodT (T1, T2) = Type ("*", [T1, T2]);
 | 
|
218  | 
||
219  | 
fun dest_prodT (Type ("*", [T1, T2])) = (T1, T2)
 | 
|
220  | 
  | dest_prodT T = raise TYPE ("dest_prodT", [T], []);
 | 
|
221  | 
||
| 14048 | 222  | 
fun pair_const T1 T2 = Const ("Pair", [T1, T2] ---> mk_prodT (T1, T2));
 | 
223  | 
||
| 4571 | 224  | 
fun mk_prod (t1, t2) =  | 
225  | 
let val T1 = fastype_of t1 and T2 = fastype_of t2 in  | 
|
| 14048 | 226  | 
pair_const T1 T2 $ t1 $ t2  | 
| 4571 | 227  | 
end;  | 
228  | 
||
229  | 
fun dest_prod (Const ("Pair", _) $ t1 $ t2) = (t1, t2)
 | 
|
230  | 
  | dest_prod t = raise TERM ("dest_prod", [t]);
 | 
|
231  | 
||
232  | 
fun mk_fst p =  | 
|
233  | 
let val pT = fastype_of p in  | 
|
234  | 
    Const ("fst", pT --> fst (dest_prodT pT)) $ p
 | 
|
235  | 
end;  | 
|
236  | 
||
237  | 
fun mk_snd p =  | 
|
238  | 
let val pT = fastype_of p in  | 
|
239  | 
    Const ("snd", pT --> snd (dest_prodT pT)) $ p
 | 
|
240  | 
end;  | 
|
241  | 
||
| 
5096
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
242  | 
(*Maps the type T1 * ... * Tn to [T1, ..., Tn], however nested*)  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
243  | 
fun prodT_factors (Type ("*", [T1, T2])) = prodT_factors T1 @ prodT_factors T2
 | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
244  | 
| prodT_factors T = [T];  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
245  | 
|
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
246  | 
fun split_const (Ta, Tb, Tc) =  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
247  | 
    Const ("split", [[Ta, Tb] ---> Tc, mk_prodT (Ta, Tb)] ---> Tc);
 | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
248  | 
|
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
249  | 
(*Makes a nested tuple from a list, following the product type structure*)  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
250  | 
fun mk_tuple (Type ("*", [T1, T2])) tms = 
 | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
251  | 
mk_prod (mk_tuple T1 tms,  | 
| 15570 | 252  | 
mk_tuple T2 (Library.drop (length (prodT_factors T1), tms)))  | 
| 
5096
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
253  | 
| mk_tuple T (t::_) = t;  | 
| 4571 | 254  | 
|
| 5207 | 255  | 
|
256  | 
||
| 9362 | 257  | 
(* proper tuples *)  | 
258  | 
||
259  | 
local (*currently unused*)  | 
|
260  | 
||
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
261  | 
fun mk_tupleT Ts = foldr mk_prodT unitT Ts;  | 
| 9362 | 262  | 
|
| 11604 | 263  | 
fun dest_tupleT (Type ("Product_Type.unit", [])) = []
 | 
| 9362 | 264  | 
  | dest_tupleT (Type ("*", [T, U])) = T :: dest_tupleT U
 | 
265  | 
  | dest_tupleT T = raise TYPE ("dest_tupleT", [T], []);
 | 
|
266  | 
||
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
267  | 
fun mk_tuple ts = foldr mk_prod unit ts;  | 
| 9362 | 268  | 
|
| 11604 | 269  | 
fun dest_tuple (Const ("Product_Type.Unity", _)) = []
 | 
| 9362 | 270  | 
  | dest_tuple (Const ("Pair", _) $ t $ u) = t :: dest_tuple u
 | 
271  | 
  | dest_tuple t = raise TERM ("dest_tuple", [t]);
 | 
|
272  | 
||
273  | 
in val _ = unit end;  | 
|
274  | 
||
275  | 
||
| 5207 | 276  | 
(* nat *)  | 
277  | 
||
278  | 
val natT = Type ("nat", []);
 | 
|
279  | 
||
280  | 
val zero = Const ("0", natT);
 | 
|
281  | 
||
282  | 
fun is_zero (Const ("0", _)) = true
 | 
|
283  | 
| is_zero _ = false;  | 
|
284  | 
||
285  | 
fun mk_Suc t = Const ("Suc", natT --> natT) $ t;
 | 
|
286  | 
||
287  | 
fun dest_Suc (Const ("Suc", _) $ t) = t
 | 
|
288  | 
  | dest_Suc t = raise TERM ("dest_Suc", [t]);
 | 
|
289  | 
||
290  | 
fun mk_nat 0 = zero  | 
|
291  | 
| mk_nat n = mk_Suc (mk_nat (n - 1));  | 
|
292  | 
||
293  | 
fun dest_nat (Const ("0", _)) = 0
 | 
|
294  | 
  | dest_nat (Const ("Suc", _) $ t) = dest_nat t + 1
 | 
|
295  | 
  | dest_nat t = raise TERM ("dest_nat", [t]);
 | 
|
296  | 
||
297  | 
||
| 7073 | 298  | 
(* binary numerals *)  | 
299  | 
||
300  | 
val binT = Type ("Numeral.bin", []);
 | 
|
301  | 
||
| 
15620
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
302  | 
val bitT = Type ("Numeral.bit", []);
 | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
303  | 
|
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
304  | 
val B0_const = Const ("Numeral.bit.B0", bitT);
 | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
305  | 
val B1_const =  Const ("Numeral.bit.B1", bitT);
 | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
306  | 
|
| 15013 | 307  | 
val pls_const = Const ("Numeral.Pls", binT)
 | 
308  | 
and min_const = Const ("Numeral.Min", binT)
 | 
|
| 
15620
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
309  | 
and bit_const = Const ("Numeral.Bit", [binT, bitT] ---> binT);
 | 
| 8768 | 310  | 
|
311  | 
fun number_of_const T = Const ("Numeral.number_of", binT --> T);
 | 
|
| 8739 | 312  | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
313  | 
fun int_of [] = 0  | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
314  | 
| int_of (b :: bs) = IntInf.fromInt b + (2 * int_of bs);  | 
| 15595 | 315  | 
|
| 
15620
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
316  | 
(*When called from a print translation, the Numeral qualifier will probably have  | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
317  | 
been removed from Bit, bin.B0, etc.*)  | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
318  | 
fun dest_bit (Const ("Numeral.bit.B0", _)) = 0
 | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
319  | 
  | dest_bit (Const ("Numeral.bit.B1", _)) = 1
 | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
320  | 
  | dest_bit (Const ("bit.B0", _)) = 0
 | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
321  | 
  | dest_bit (Const ("bit.B1", _)) = 1
 | 
| 
7548
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
322  | 
  | dest_bit t = raise TERM("dest_bit", [t]);
 | 
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
323  | 
|
| 15013 | 324  | 
fun bin_of (Const ("Numeral.Pls", _)) = []
 | 
325  | 
  | bin_of (Const ("Numeral.Min", _)) = [~1]
 | 
|
326  | 
  | bin_of (Const ("Numeral.Bit", _) $ bs $ b) = dest_bit b :: bin_of bs
 | 
|
| 
15620
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
327  | 
  | bin_of (Const ("Bit", _) $ bs $ b) = dest_bit b :: bin_of bs
 | 
| 
7548
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
328  | 
  | bin_of t = raise TERM("bin_of", [t]);
 | 
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
329  | 
|
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
330  | 
val dest_binum = int_of o bin_of;  | 
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
331  | 
|
| 
15620
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
332  | 
fun mk_bit 0 = B0_const  | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
333  | 
| mk_bit 1 = B1_const  | 
| 10693 | 334  | 
| mk_bit _ = sys_error "mk_bit";  | 
335  | 
||
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
336  | 
fun mk_bin n =  | 
| 15595 | 337  | 
let  | 
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
338  | 
fun mk_bit n = if n = 0 then B0_const else B1_const  | 
| 15595 | 339  | 
|
340  | 
fun bin_of n =  | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
341  | 
if n = 0 then pls_const  | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
342  | 
else if n = ~1 then min_const  | 
| 15595 | 343  | 
else  | 
344  | 
let  | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
345  | 
(*val (q,r) = IntInf.divMod (n, 2): doesn't work in SML 10.0.7, but in newer versions! FIXME: put this back after new SML released!*)  | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
346  | 
val q = IntInf.div (n, 2)  | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
347  | 
val r = IntInf.mod (n, 2)  | 
| 15595 | 348  | 
in  | 
349  | 
bit_const $ bin_of q $ mk_bit r  | 
|
350  | 
end  | 
|
351  | 
in  | 
|
352  | 
bin_of n  | 
|
353  | 
end  | 
|
| 13755 | 354  | 
|
| 
15620
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
355  | 
|
| 13755 | 356  | 
(* int *)  | 
357  | 
||
358  | 
val intT = Type ("IntDef.int", []);
 | 
|
359  | 
||
| 
14103
 
afd168fdcd3a
mk_int now produces specific constants for 0 and 1.
 
berghofe 
parents: 
14048 
diff
changeset
 | 
360  | 
fun mk_int 0 = Const ("0", intT)
 | 
| 
 
afd168fdcd3a
mk_int now produces specific constants for 0 and 1.
 
berghofe 
parents: 
14048 
diff
changeset
 | 
361  | 
  | mk_int 1 = Const ("1", intT)
 | 
| 
 
afd168fdcd3a
mk_int now produces specific constants for 0 and 1.
 
berghofe 
parents: 
14048 
diff
changeset
 | 
362  | 
| mk_int i = number_of_const intT $ mk_bin i;  | 
| 13755 | 363  | 
|
364  | 
||
365  | 
(* real *)  | 
|
366  | 
||
367  | 
val realT = Type("RealDef.real", []);
 | 
|
368  | 
||
369  | 
||
370  | 
(* list *)  | 
|
371  | 
||
372  | 
fun mk_list f T [] = Const ("List.list.Nil", Type ("List.list", [T]))
 | 
|
373  | 
  | mk_list f T (x :: xs) = Const ("List.list.Cons",
 | 
|
374  | 
      T --> Type ("List.list", [T]) --> Type ("List.list", [T])) $ f x $
 | 
|
375  | 
mk_list f T xs;  | 
|
376  | 
||
| 15062 | 377  | 
fun dest_list (Const ("List.list.Nil", _)) = []
 | 
378  | 
  | dest_list (Const ("List.list.Cons", _) $ x $ xs) = x :: dest_list xs
 | 
|
379  | 
  | dest_list t = raise TERM ("dest_list", [t]);
 | 
|
380  | 
||
| 923 | 381  | 
end;  |