| author | wenzelm | 
| Mon, 26 Jun 2000 00:00:40 +0200 | |
| changeset 9143 | 6180c29d2db6 | 
| parent 8768 | 9f18aba48519 | 
| child 9362 | b78d4246a320 | 
| 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  | 
|
10  | 
val termC: class  | 
|
11  | 
val termS: sort  | 
|
| 8100 | 12  | 
val termT: typ  | 
| 8275 | 13  | 
val boolN: string  | 
| 923 | 14  | 
val boolT: typ  | 
| 7073 | 15  | 
val false_const: term  | 
16  | 
val true_const: term  | 
|
| 923 | 17  | 
val mk_setT: typ -> typ  | 
18  | 
val dest_setT: typ -> typ  | 
|
19  | 
val mk_Trueprop: term -> term  | 
|
20  | 
val dest_Trueprop: term -> term  | 
|
21  | 
val conj: term  | 
|
22  | 
val disj: term  | 
|
23  | 
val imp: term  | 
|
| 8429 | 24  | 
val Not: term  | 
| 7690 | 25  | 
val mk_conj: term * term -> term  | 
26  | 
val mk_disj: term * term -> term  | 
|
27  | 
val mk_imp: term * term -> term  | 
|
| 4571 | 28  | 
val dest_imp: term -> term * term  | 
| 8302 | 29  | 
val dest_conj: term -> term list  | 
| 923 | 30  | 
val eq_const: typ -> term  | 
31  | 
val all_const: typ -> term  | 
|
32  | 
val exists_const: typ -> term  | 
|
33  | 
val Collect_const: typ -> term  | 
|
34  | 
val mk_eq: term * term -> term  | 
|
| 6031 | 35  | 
val dest_eq: term -> term * term  | 
| 923 | 36  | 
val mk_all: string * typ * term -> term  | 
37  | 
val mk_exists: string * typ * term -> term  | 
|
38  | 
val mk_Collect: string * typ * term -> term  | 
|
39  | 
val mk_mem: term * term -> term  | 
|
| 6380 | 40  | 
val dest_mem: term -> term * term  | 
| 2510 | 41  | 
val mk_binop: string -> term * term -> term  | 
42  | 
val mk_binrel: string -> term * term -> term  | 
|
43  | 
val dest_bin: string -> typ -> term -> term * term  | 
|
| 4571 | 44  | 
val unitT: typ  | 
45  | 
val unit: term  | 
|
46  | 
val is_unit: term -> bool  | 
|
47  | 
val mk_prodT: typ * typ -> typ  | 
|
48  | 
val dest_prodT: typ -> typ * typ  | 
|
49  | 
val mk_prod: term * term -> term  | 
|
50  | 
val dest_prod: term -> term * term  | 
|
51  | 
val mk_fst: term -> term  | 
|
52  | 
val mk_snd: term -> term  | 
|
| 
5096
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
53  | 
val prodT_factors: typ -> typ list  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
54  | 
val split_const: typ * typ * typ -> term  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
55  | 
val mk_tuple: typ -> term list -> term  | 
| 5207 | 56  | 
val natT: typ  | 
57  | 
val zero: term  | 
|
58  | 
val is_zero: term -> bool  | 
|
59  | 
val mk_Suc: term -> term  | 
|
60  | 
val dest_Suc: term -> term  | 
|
61  | 
val mk_nat: int -> term  | 
|
62  | 
val dest_nat: term -> int  | 
|
| 7073 | 63  | 
val intT: typ  | 
| 7163 | 64  | 
val realT: typ  | 
| 7073 | 65  | 
val binT: typ  | 
66  | 
val pls_const: term  | 
|
67  | 
val min_const: term  | 
|
68  | 
val bit_const: term  | 
|
| 8768 | 69  | 
val number_of_const: typ -> term  | 
| 
7548
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
70  | 
val int_of: int list -> int  | 
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
71  | 
val dest_binum: term -> int  | 
| 923 | 72  | 
end;  | 
73  | 
||
| 2510 | 74  | 
|
| 923 | 75  | 
structure HOLogic: HOLOGIC =  | 
76  | 
struct  | 
|
77  | 
||
| 2510 | 78  | 
(* basics *)  | 
| 923 | 79  | 
|
80  | 
val termC: class = "term";  | 
|
81  | 
val termS: sort = [termC];  | 
|
| 8100 | 82  | 
val termT = TypeInfer.anyT termS;  | 
| 923 | 83  | 
|
84  | 
||
| 2510 | 85  | 
(* bool and set *)  | 
| 923 | 86  | 
|
| 8275 | 87  | 
val boolN = "bool";  | 
88  | 
val boolT = Type (boolN, []);  | 
|
| 923 | 89  | 
|
| 7073 | 90  | 
val true_const =  Const ("True", boolT)
 | 
91  | 
and false_const = Const ("False", boolT);
 | 
|
92  | 
||
| 923 | 93  | 
fun mk_setT T = Type ("set", [T]);
 | 
94  | 
||
95  | 
fun dest_setT (Type ("set", [T])) = T
 | 
|
| 3794 | 96  | 
  | dest_setT T = raise TYPE ("dest_setT: set type expected", [T], []);
 | 
| 923 | 97  | 
|
| 8275 | 98  | 
|
| 7073 | 99  | 
(* logic *)  | 
| 923 | 100  | 
|
101  | 
val Trueprop = Const ("Trueprop", boolT --> propT);
 | 
|
102  | 
||
103  | 
fun mk_Trueprop P = Trueprop $ P;  | 
|
104  | 
||
105  | 
fun dest_Trueprop (Const ("Trueprop", _) $ P) = P
 | 
|
| 3794 | 106  | 
  | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]);
 | 
| 923 | 107  | 
|
108  | 
||
109  | 
val conj = Const ("op &", [boolT, boolT] ---> boolT)
 | 
|
110  | 
and disj = Const ("op |", [boolT, boolT] ---> boolT)
 | 
|
| 8429 | 111  | 
and imp = Const ("op -->", [boolT, boolT] ---> boolT)
 | 
112  | 
and Not = Const ("Not", boolT --> boolT);
 | 
|
| 923 | 113  | 
|
| 7690 | 114  | 
fun mk_conj (t1, t2) = conj $ t1 $ t2  | 
115  | 
and mk_disj (t1, t2) = disj $ t1 $ t2  | 
|
116  | 
and mk_imp (t1, t2) = imp $ t1 $ t2;  | 
|
117  | 
||
| 
4466
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4294 
diff
changeset
 | 
118  | 
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
 | 
119  | 
  | 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
 | 
120  | 
|
| 8302 | 121  | 
fun dest_conj (Const ("op &", _) $ t $ t') = t :: dest_conj t'
 | 
122  | 
| dest_conj t = [t];  | 
|
123  | 
||
| 923 | 124  | 
fun eq_const T = Const ("op =", [T, T] ---> boolT);
 | 
125  | 
fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u;  | 
|
126  | 
||
| 6031 | 127  | 
fun dest_eq (Const ("op =", _) $ lhs $ rhs) = (lhs, rhs)
 | 
128  | 
  | dest_eq t = raise TERM ("dest_eq", [t])
 | 
|
129  | 
||
| 923 | 130  | 
fun all_const T = Const ("All", [T --> boolT] ---> boolT);
 | 
131  | 
fun mk_all (x, T, P) = all_const T $ absfree (x, T, P);  | 
|
132  | 
||
133  | 
fun exists_const T = Const ("Ex", [T --> boolT] ---> boolT);
 | 
|
134  | 
fun mk_exists (x, T, P) = exists_const T $ absfree (x, T, P);  | 
|
135  | 
||
136  | 
fun Collect_const T = Const ("Collect", [T --> boolT] ---> mk_setT T);
 | 
|
137  | 
fun mk_Collect (a, T, t) = Collect_const T $ absfree (a, T, t);  | 
|
138  | 
||
139  | 
fun mk_mem (x, A) =  | 
|
140  | 
let val setT = fastype_of A in  | 
|
141  | 
    Const ("op :", [dest_setT setT, setT] ---> boolT) $ x $ A
 | 
|
142  | 
end;  | 
|
143  | 
||
| 6380 | 144  | 
fun dest_mem (Const ("op :", _) $ x $ A) = (x, A)
 | 
145  | 
  | dest_mem t = raise TERM ("dest_mem", [t]);
 | 
|
146  | 
||
| 923 | 147  | 
|
| 2510 | 148  | 
(* binary oprations and relations *)  | 
149  | 
||
150  | 
fun mk_binop c (t, u) =  | 
|
151  | 
let val T = fastype_of t in  | 
|
152  | 
Const (c, [T, T] ---> T) $ t $ u  | 
|
153  | 
end;  | 
|
154  | 
||
155  | 
fun mk_binrel c (t, u) =  | 
|
156  | 
let val T = fastype_of t in  | 
|
157  | 
Const (c, [T, T] ---> boolT) $ t $ u  | 
|
158  | 
end;  | 
|
159  | 
||
160  | 
fun dest_bin c T (tm as Const (c', Type ("fun", [T', _])) $ t $ u) =
 | 
|
161  | 
if c = c' andalso T = T' then (t, u)  | 
|
| 3794 | 162  | 
      else raise TERM ("dest_bin " ^ c, [tm])
 | 
163  | 
  | dest_bin c _ tm = raise TERM ("dest_bin " ^ c, [tm]);
 | 
|
| 2510 | 164  | 
|
165  | 
||
| 4571 | 166  | 
(* unit *)  | 
167  | 
||
168  | 
val unitT = Type ("unit", []);
 | 
|
169  | 
||
170  | 
val unit = Const ("()", unitT);
 | 
|
171  | 
||
172  | 
fun is_unit (Const ("()", _)) = true
 | 
|
173  | 
| is_unit _ = false;  | 
|
174  | 
||
175  | 
||
176  | 
(* prod *)  | 
|
177  | 
||
178  | 
fun mk_prodT (T1, T2) = Type ("*", [T1, T2]);
 | 
|
179  | 
||
180  | 
fun dest_prodT (Type ("*", [T1, T2])) = (T1, T2)
 | 
|
181  | 
  | dest_prodT T = raise TYPE ("dest_prodT", [T], []);
 | 
|
182  | 
||
183  | 
fun mk_prod (t1, t2) =  | 
|
184  | 
let val T1 = fastype_of t1 and T2 = fastype_of t2 in  | 
|
185  | 
    Const ("Pair", [T1, T2] ---> mk_prodT (T1, T2)) $ t1 $ t2
 | 
|
186  | 
end;  | 
|
187  | 
||
188  | 
fun dest_prod (Const ("Pair", _) $ t1 $ t2) = (t1, t2)
 | 
|
189  | 
  | dest_prod t = raise TERM ("dest_prod", [t]);
 | 
|
190  | 
||
191  | 
fun mk_fst p =  | 
|
192  | 
let val pT = fastype_of p in  | 
|
193  | 
    Const ("fst", pT --> fst (dest_prodT pT)) $ p
 | 
|
194  | 
end;  | 
|
195  | 
||
196  | 
fun mk_snd p =  | 
|
197  | 
let val pT = fastype_of p in  | 
|
198  | 
    Const ("snd", pT --> snd (dest_prodT pT)) $ p
 | 
|
199  | 
end;  | 
|
200  | 
||
| 
5096
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
201  | 
(*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
 | 
202  | 
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
 | 
203  | 
| prodT_factors T = [T];  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
204  | 
|
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
205  | 
fun split_const (Ta, Tb, Tc) =  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
206  | 
    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
 | 
207  | 
|
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
208  | 
(*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
 | 
209  | 
fun mk_tuple (Type ("*", [T1, T2])) tms = 
 | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
210  | 
mk_prod (mk_tuple T1 tms,  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
211  | 
mk_tuple T2 (drop (length (prodT_factors T1), tms)))  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
212  | 
| mk_tuple T (t::_) = t;  | 
| 4571 | 213  | 
|
| 5207 | 214  | 
|
215  | 
||
216  | 
(* nat *)  | 
|
217  | 
||
218  | 
val natT = Type ("nat", []);
 | 
|
219  | 
||
220  | 
val zero = Const ("0", natT);
 | 
|
221  | 
||
222  | 
fun is_zero (Const ("0", _)) = true
 | 
|
223  | 
| is_zero _ = false;  | 
|
224  | 
||
225  | 
fun mk_Suc t = Const ("Suc", natT --> natT) $ t;
 | 
|
226  | 
||
227  | 
fun dest_Suc (Const ("Suc", _) $ t) = t
 | 
|
228  | 
  | dest_Suc t = raise TERM ("dest_Suc", [t]);
 | 
|
229  | 
||
230  | 
fun mk_nat 0 = zero  | 
|
231  | 
| mk_nat n = mk_Suc (mk_nat (n - 1));  | 
|
232  | 
||
233  | 
fun dest_nat (Const ("0", _)) = 0
 | 
|
234  | 
  | dest_nat (Const ("Suc", _) $ t) = dest_nat t + 1
 | 
|
235  | 
  | dest_nat t = raise TERM ("dest_nat", [t]);
 | 
|
236  | 
||
237  | 
||
| 7073 | 238  | 
val intT = Type ("IntDef.int", []);
 | 
239  | 
||
| 7163 | 240  | 
val realT = Type("RealDef.real",[]);
 | 
241  | 
||
| 7073 | 242  | 
|
243  | 
(* binary numerals *)  | 
|
244  | 
||
245  | 
val binT = Type ("Numeral.bin", []);
 | 
|
246  | 
||
247  | 
val pls_const =  Const ("Numeral.bin.Pls", binT)
 | 
|
248  | 
and min_const = Const ("Numeral.bin.Min", binT)
 | 
|
| 8768 | 249  | 
and bit_const = Const ("Numeral.bin.Bit", [binT, boolT] ---> binT);
 | 
250  | 
||
251  | 
fun number_of_const T = Const ("Numeral.number_of", binT --> T);
 | 
|
| 8739 | 252  | 
|
| 7073 | 253  | 
|
| 
7548
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
254  | 
fun int_of [] = 0  | 
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
255  | 
| int_of (b :: bs) = b + 2 * int_of bs;  | 
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
256  | 
|
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
257  | 
fun dest_bit (Const ("False", _)) = 0
 | 
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
258  | 
  | dest_bit (Const ("True", _)) = 1
 | 
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
259  | 
  | 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
 | 
260  | 
|
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
261  | 
fun bin_of (Const ("Numeral.bin.Pls", _)) = []
 | 
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
262  | 
  | bin_of (Const ("Numeral.bin.Min", _)) = [~1]
 | 
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
263  | 
  | bin_of (Const ("Numeral.bin.Bit", _) $ bs $ b) = dest_bit b :: bin_of bs
 | 
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
264  | 
  | 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
 | 
265  | 
|
| 
 
9e29a3af64ab
ROOT: Integ/bin_simprocs.ML now loaded in Integ/Bin.ML
 
nipkow 
parents: 
7163 
diff
changeset
 | 
266  | 
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
 | 
267  | 
|
| 923 | 268  | 
end;  |