| author | haftmann | 
| Thu, 04 Jan 2007 20:01:00 +0100 | |
| changeset 22005 | 0faa5afd5d69 | 
| parent 21829 | 016eff9c5699 | 
| child 22391 | 56861fe9c22c | 
| 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  | 
| 8275 | 12  | 
val boolN: string  | 
| 923 | 13  | 
val boolT: typ  | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
14  | 
val true_const: term  | 
| 7073 | 15  | 
val false_const: term  | 
| 923 | 16  | 
val mk_setT: typ -> typ  | 
17  | 
val dest_setT: typ -> typ  | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15945 
diff
changeset
 | 
18  | 
val Trueprop: term  | 
| 923 | 19  | 
val mk_Trueprop: term -> term  | 
20  | 
val dest_Trueprop: term -> term  | 
|
| 21550 | 21  | 
val Trueprop_conv: (cterm -> thm) -> cterm -> thm  | 
| 923 | 22  | 
val conj: term  | 
23  | 
val disj: term  | 
|
24  | 
val imp: term  | 
|
| 8429 | 25  | 
val Not: term  | 
| 7690 | 26  | 
val mk_conj: term * term -> term  | 
27  | 
val mk_disj: term * term -> term  | 
|
28  | 
val mk_imp: term * term -> term  | 
|
| 16835 | 29  | 
val mk_not: 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  | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
36  | 
val mk_eq: term * term -> term  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
37  | 
val dest_eq: term -> term * term  | 
| 923 | 38  | 
val all_const: typ -> term  | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
39  | 
val mk_all: string * typ * term -> term  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
40  | 
val list_all: (string * typ) list * term -> term  | 
| 923 | 41  | 
val exists_const: typ -> term  | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
42  | 
val mk_exists: string * typ * term -> term  | 
| 15945 | 43  | 
val choice_const: typ -> term  | 
| 923 | 44  | 
val Collect_const: typ -> term  | 
45  | 
val mk_Collect: string * typ * term -> term  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
46  | 
val class_eq: string  | 
| 923 | 47  | 
val mk_mem: term * term -> term  | 
| 6380 | 48  | 
val dest_mem: term -> term * term  | 
| 11818 | 49  | 
val mk_UNIV: typ -> term  | 
| 2510 | 50  | 
val mk_binop: string -> term * term -> term  | 
51  | 
val mk_binrel: string -> term * term -> term  | 
|
52  | 
val dest_bin: string -> typ -> term -> term * term  | 
|
| 4571 | 53  | 
val unitT: typ  | 
| 9362 | 54  | 
val is_unitT: typ -> bool  | 
| 4571 | 55  | 
val unit: term  | 
56  | 
val is_unit: term -> bool  | 
|
57  | 
val mk_prodT: typ * typ -> typ  | 
|
58  | 
val dest_prodT: typ -> typ * typ  | 
|
| 14048 | 59  | 
val pair_const: typ -> typ -> term  | 
| 4571 | 60  | 
val mk_prod: term * term -> term  | 
61  | 
val dest_prod: term -> term * term  | 
|
62  | 
val mk_fst: term -> term  | 
|
63  | 
val mk_snd: term -> term  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
64  | 
val split_const: typ * typ * typ -> term  | 
| 18285 | 65  | 
val mk_split: term -> term  | 
| 
5096
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
66  | 
val prodT_factors: typ -> typ list  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
67  | 
val mk_tuple: typ -> term list -> term  | 
| 5207 | 68  | 
val natT: typ  | 
69  | 
val zero: term  | 
|
70  | 
val is_zero: term -> bool  | 
|
71  | 
val mk_Suc: term -> term  | 
|
72  | 
val dest_Suc: term -> term  | 
|
| 21621 | 73  | 
val Suc_zero: term  | 
74  | 
val mk_nat: IntInf.int -> term  | 
|
75  | 
val dest_nat: term -> IntInf.int  | 
|
| 
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  | 
| 
21778
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
79  | 
val mk_bit: int -> term  | 
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
80  | 
val dest_bit: term -> int  | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
81  | 
val intT: typ  | 
| 7073 | 82  | 
val pls_const: term  | 
83  | 
val min_const: term  | 
|
84  | 
val bit_const: term  | 
|
| 
21820
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
85  | 
val mk_numeral: IntInf.int -> term  | 
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
86  | 
val dest_numeral: term -> IntInf.int  | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
87  | 
val number_of_const: typ -> term  | 
| 
21820
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
88  | 
val mk_number: typ -> IntInf.int -> term  | 
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
89  | 
val dest_number: term -> typ * IntInf.int  | 
| 
21778
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
90  | 
val realT: typ  | 
| 21755 | 91  | 
val nibbleT: typ  | 
92  | 
val mk_nibble: int -> term  | 
|
93  | 
val dest_nibble: term -> int  | 
|
94  | 
val charT: typ  | 
|
95  | 
val mk_char: int -> term  | 
|
96  | 
val dest_char: term -> int  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
97  | 
val listT: typ -> typ  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
98  | 
val mk_list: typ -> term list -> term  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
99  | 
val dest_list: term -> term list  | 
| 
21778
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
100  | 
val stringT: typ  | 
| 
21820
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
101  | 
val mk_string: string -> term  | 
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
102  | 
val dest_string: term -> string  | 
| 923 | 103  | 
end;  | 
104  | 
||
105  | 
structure HOLogic: HOLOGIC =  | 
|
106  | 
struct  | 
|
107  | 
||
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
108  | 
(* HOL syntax *)  | 
| 923 | 109  | 
|
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
110  | 
val typeS: sort = ["HOL.type"];  | 
| 
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
111  | 
val typeT = TypeInfer.anyT typeS;  | 
| 
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
112  | 
|
| 923 | 113  | 
|
| 2510 | 114  | 
(* bool and set *)  | 
| 923 | 115  | 
|
| 8275 | 116  | 
val boolN = "bool";  | 
117  | 
val boolT = Type (boolN, []);  | 
|
| 923 | 118  | 
|
| 9856 | 119  | 
val true_const =  Const ("True", boolT);
 | 
120  | 
val false_const = Const ("False", boolT);
 | 
|
| 7073 | 121  | 
|
| 923 | 122  | 
fun mk_setT T = Type ("set", [T]);
 | 
123  | 
||
124  | 
fun dest_setT (Type ("set", [T])) = T
 | 
|
| 3794 | 125  | 
  | dest_setT T = raise TYPE ("dest_setT: set type expected", [T], []);
 | 
| 923 | 126  | 
|
| 8275 | 127  | 
|
| 7073 | 128  | 
(* logic *)  | 
| 923 | 129  | 
|
130  | 
val Trueprop = Const ("Trueprop", boolT --> propT);
 | 
|
131  | 
||
132  | 
fun mk_Trueprop P = Trueprop $ P;  | 
|
133  | 
||
134  | 
fun dest_Trueprop (Const ("Trueprop", _) $ P) = P
 | 
|
| 3794 | 135  | 
  | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]);
 | 
| 923 | 136  | 
|
| 21550 | 137  | 
fun Trueprop_conv conv ct = (case term_of ct of  | 
138  | 
    Const ("Trueprop", _) $ _ =>
 | 
|
139  | 
let val (ct1, ct2) = Thm.dest_comb ct  | 
|
140  | 
in Thm.combination (Thm.reflexive ct1) (conv ct2) end  | 
|
141  | 
  | _ => raise TERM ("Trueprop_conv", []));
 | 
|
| 923 | 142  | 
|
143  | 
val conj = Const ("op &", [boolT, boolT] ---> boolT)
 | 
|
144  | 
and disj = Const ("op |", [boolT, boolT] ---> boolT)
 | 
|
| 8429 | 145  | 
and imp = Const ("op -->", [boolT, boolT] ---> boolT)
 | 
146  | 
and Not = Const ("Not", boolT --> boolT);
 | 
|
| 923 | 147  | 
|
| 7690 | 148  | 
fun mk_conj (t1, t2) = conj $ t1 $ t2  | 
149  | 
and mk_disj (t1, t2) = disj $ t1 $ t2  | 
|
| 16835 | 150  | 
and mk_imp (t1, t2) = imp $ t1 $ t2  | 
151  | 
and mk_not t = Not $ t;  | 
|
| 7690 | 152  | 
|
| 15151 | 153  | 
fun dest_conj (Const ("op &", _) $ t $ t') = t :: dest_conj t'
 | 
154  | 
| dest_conj t = [t];  | 
|
155  | 
||
| 15945 | 156  | 
fun dest_disj (Const ("op |", _) $ t $ t') = t :: dest_disj t'
 | 
157  | 
| dest_disj t = [t];  | 
|
158  | 
||
| 
4466
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4294 
diff
changeset
 | 
159  | 
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
 | 
160  | 
  | 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
 | 
161  | 
|
| 15151 | 162  | 
fun dest_not (Const ("Not", _) $ t) = t
 | 
163  | 
  | dest_not t = raise TERM ("dest_not", [t]);
 | 
|
| 8302 | 164  | 
|
| 11683 | 165  | 
fun imp_concl_of t = imp_concl_of (#2 (dest_imp t)) handle TERM _ => t;  | 
166  | 
val dest_concls = map imp_concl_of o dest_conj o dest_Trueprop;  | 
|
167  | 
||
| 923 | 168  | 
fun eq_const T = Const ("op =", [T, T] ---> boolT);
 | 
169  | 
fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u;  | 
|
170  | 
||
| 6031 | 171  | 
fun dest_eq (Const ("op =", _) $ lhs $ rhs) = (lhs, rhs)
 | 
172  | 
  | dest_eq t = raise TERM ("dest_eq", [t])
 | 
|
173  | 
||
| 923 | 174  | 
fun all_const T = Const ("All", [T --> boolT] ---> boolT);
 | 
175  | 
fun mk_all (x, T, P) = all_const T $ absfree (x, T, P);  | 
|
| 
21173
 
663a7b39894c
removed is_Trueprop (use can dest_Trueprop'' instead);
 
wenzelm 
parents: 
21078 
diff
changeset
 | 
176  | 
fun list_all (xs, t) = fold_rev (fn (x, T) => fn P => all_const T $ Abs (x, T, P)) xs t;  | 
| 923 | 177  | 
|
178  | 
fun exists_const T = Const ("Ex", [T --> boolT] ---> boolT);
 | 
|
179  | 
fun mk_exists (x, T, P) = exists_const T $ absfree (x, T, P);  | 
|
180  | 
||
| 21755 | 181  | 
fun choice_const T = Const("Hilbert_Choice.Eps", (T --> boolT) --> T);
 | 
| 15945 | 182  | 
|
| 923 | 183  | 
fun Collect_const T = Const ("Collect", [T --> boolT] ---> mk_setT T);
 | 
184  | 
fun mk_Collect (a, T, t) = Collect_const T $ absfree (a, T, t);  | 
|
185  | 
||
| 21045 | 186  | 
val class_eq = "Code_Generator.eq";  | 
187  | 
||
| 923 | 188  | 
fun mk_mem (x, A) =  | 
189  | 
let val setT = fastype_of A in  | 
|
190  | 
    Const ("op :", [dest_setT setT, setT] ---> boolT) $ x $ A
 | 
|
191  | 
end;  | 
|
192  | 
||
| 6380 | 193  | 
fun dest_mem (Const ("op :", _) $ x $ A) = (x, A)
 | 
194  | 
  | dest_mem t = raise TERM ("dest_mem", [t]);
 | 
|
195  | 
||
| 11818 | 196  | 
fun mk_UNIV T = Const ("UNIV", mk_setT T);
 | 
197  | 
||
| 923 | 198  | 
|
| 
13743
 
f8f9393be64c
Fixed bug in simpdata.ML that prevented the use of congruence rules from a
 
ballarin 
parents: 
13640 
diff
changeset
 | 
199  | 
(* binary operations and relations *)  | 
| 2510 | 200  | 
|
201  | 
fun mk_binop c (t, u) =  | 
|
202  | 
let val T = fastype_of t in  | 
|
203  | 
Const (c, [T, T] ---> T) $ t $ u  | 
|
204  | 
end;  | 
|
205  | 
||
206  | 
fun mk_binrel c (t, u) =  | 
|
207  | 
let val T = fastype_of t in  | 
|
208  | 
Const (c, [T, T] ---> boolT) $ t $ u  | 
|
209  | 
end;  | 
|
210  | 
||
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14103 
diff
changeset
 | 
211  | 
(*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
 | 
212  | 
way of handling polymorphic operators.*)  | 
| 2510 | 213  | 
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
 | 
214  | 
if c = c' andalso (T=T' orelse T=dummyT) then (t, u)  | 
| 3794 | 215  | 
      else raise TERM ("dest_bin " ^ c, [tm])
 | 
216  | 
  | dest_bin c _ tm = raise TERM ("dest_bin " ^ c, [tm]);
 | 
|
| 2510 | 217  | 
|
218  | 
||
| 4571 | 219  | 
(* unit *)  | 
220  | 
||
| 11604 | 221  | 
val unitT = Type ("Product_Type.unit", []);
 | 
| 4571 | 222  | 
|
| 11604 | 223  | 
fun is_unitT (Type ("Product_Type.unit", [])) = true
 | 
| 9362 | 224  | 
| is_unitT _ = false;  | 
225  | 
||
| 11604 | 226  | 
val unit = Const ("Product_Type.Unity", unitT);
 | 
| 4571 | 227  | 
|
| 11604 | 228  | 
fun is_unit (Const ("Product_Type.Unity", _)) = true
 | 
| 4571 | 229  | 
| is_unit _ = false;  | 
230  | 
||
231  | 
||
232  | 
(* prod *)  | 
|
233  | 
||
234  | 
fun mk_prodT (T1, T2) = Type ("*", [T1, T2]);
 | 
|
235  | 
||
236  | 
fun dest_prodT (Type ("*", [T1, T2])) = (T1, T2)
 | 
|
237  | 
  | dest_prodT T = raise TYPE ("dest_prodT", [T], []);
 | 
|
238  | 
||
| 14048 | 239  | 
fun pair_const T1 T2 = Const ("Pair", [T1, T2] ---> mk_prodT (T1, T2));
 | 
240  | 
||
| 4571 | 241  | 
fun mk_prod (t1, t2) =  | 
242  | 
let val T1 = fastype_of t1 and T2 = fastype_of t2 in  | 
|
| 14048 | 243  | 
pair_const T1 T2 $ t1 $ t2  | 
| 4571 | 244  | 
end;  | 
245  | 
||
246  | 
fun dest_prod (Const ("Pair", _) $ t1 $ t2) = (t1, t2)
 | 
|
247  | 
  | dest_prod t = raise TERM ("dest_prod", [t]);
 | 
|
248  | 
||
249  | 
fun mk_fst p =  | 
|
250  | 
let val pT = fastype_of p in  | 
|
251  | 
    Const ("fst", pT --> fst (dest_prodT pT)) $ p
 | 
|
252  | 
end;  | 
|
253  | 
||
254  | 
fun mk_snd p =  | 
|
255  | 
let val pT = fastype_of p in  | 
|
256  | 
    Const ("snd", pT --> snd (dest_prodT pT)) $ p
 | 
|
257  | 
end;  | 
|
258  | 
||
| 18285 | 259  | 
fun split_const (A, B, C) =  | 
260  | 
  Const ("split", (A --> B --> C) --> mk_prodT (A, B) --> C);
 | 
|
261  | 
||
262  | 
fun mk_split t =  | 
|
263  | 
(case Term.fastype_of t of  | 
|
264  | 
    T as (Type ("fun", [A, Type ("fun", [B, C])])) =>
 | 
|
265  | 
      Const ("split", T --> mk_prodT (A, B) --> C) $ t
 | 
|
266  | 
  | _ => raise TERM ("mk_split: bad body type", [t]));
 | 
|
267  | 
||
| 
5096
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
268  | 
(*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
 | 
269  | 
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
 | 
270  | 
| prodT_factors T = [T];  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
271  | 
|
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
272  | 
(*Makes a nested tuple from a list, following the product type structure*)  | 
| 16971 | 273  | 
fun mk_tuple (Type ("*", [T1, T2])) tms =
 | 
274  | 
mk_prod (mk_tuple T1 tms,  | 
|
| 15570 | 275  | 
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
 | 
276  | 
| mk_tuple T (t::_) = t;  | 
| 4571 | 277  | 
|
| 5207 | 278  | 
|
279  | 
(* nat *)  | 
|
280  | 
||
281  | 
val natT = Type ("nat", []);
 | 
|
282  | 
||
| 
20713
 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 
haftmann 
parents: 
20485 
diff
changeset
 | 
283  | 
val zero = Const ("HOL.zero", natT);
 | 
| 5207 | 284  | 
|
| 
20713
 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 
haftmann 
parents: 
20485 
diff
changeset
 | 
285  | 
fun is_zero (Const ("HOL.zero", _)) = true
 | 
| 5207 | 286  | 
| is_zero _ = false;  | 
287  | 
||
288  | 
fun mk_Suc t = Const ("Suc", natT --> natT) $ t;
 | 
|
289  | 
||
290  | 
fun dest_Suc (Const ("Suc", _) $ t) = t
 | 
|
291  | 
  | dest_Suc t = raise TERM ("dest_Suc", [t]);
 | 
|
292  | 
||
| 21621 | 293  | 
val Suc_zero = mk_Suc zero;  | 
294  | 
||
| 5207 | 295  | 
fun mk_nat 0 = zero  | 
| 21621 | 296  | 
| mk_nat n = mk_Suc (mk_nat (IntInf.- (n, 1)));  | 
| 5207 | 297  | 
|
| 
20713
 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 
haftmann 
parents: 
20485 
diff
changeset
 | 
298  | 
fun dest_nat (Const ("HOL.zero", _)) = 0
 | 
| 21621 | 299  | 
  | dest_nat (Const ("Suc", _) $ t) = IntInf.+ (dest_nat t, 1)
 | 
| 5207 | 300  | 
  | dest_nat t = raise TERM ("dest_nat", [t]);
 | 
301  | 
||
302  | 
||
| 
21778
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
303  | 
(* bit *)  | 
| 7073 | 304  | 
|
| 
15620
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
305  | 
val bitT = Type ("Numeral.bit", []);
 | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
306  | 
|
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
307  | 
val B0_const = Const ("Numeral.bit.B0", bitT);
 | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
308  | 
val B1_const =  Const ("Numeral.bit.B1", bitT);
 | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15595 
diff
changeset
 | 
309  | 
|
| 
21778
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
310  | 
fun mk_bit 0 = B0_const  | 
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
311  | 
| mk_bit 1 = B1_const  | 
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
312  | 
  | mk_bit _ = raise TERM ("mk_bit", []);
 | 
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
313  | 
|
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
314  | 
fun dest_bit (Const ("Numeral.bit.B0", _)) = 0
 | 
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
315  | 
  | dest_bit (Const ("Numeral.bit.B1", _)) = 1
 | 
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
316  | 
  | dest_bit t = raise TERM ("dest_bit", [t]);
 | 
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
317  | 
|
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
318  | 
|
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
319  | 
(* binary numerals and int -- non-unique representation due to leading zeros/ones! *)  | 
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
320  | 
|
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
321  | 
val intT = Type ("IntDef.int", []);
 | 
| 
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
322  | 
|
| 20485 | 323  | 
val pls_const = Const ("Numeral.Pls", intT)
 | 
324  | 
and min_const = Const ("Numeral.Min", intT)
 | 
|
325  | 
and bit_const = Const ("Numeral.Bit", [intT, bitT] ---> intT);
 | 
|
| 8768 | 326  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
327  | 
fun mk_numeral 0 = pls_const  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
328  | 
| mk_numeral ~1 = min_const  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
329  | 
| mk_numeral i =  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
330  | 
let val (q, r) = IntInf.divMod (i, 2)  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
331  | 
in bit_const $ mk_numeral q $ mk_bit (IntInf.toInt r) end;  | 
| 13755 | 332  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
333  | 
fun dest_numeral (Const ("Numeral.Pls", _)) = 0
 | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
334  | 
  | dest_numeral (Const ("Numeral.Min", _)) = ~1
 | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
335  | 
  | dest_numeral (Const ("Numeral.Bit", _) $ bs $ b) =
 | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
336  | 
2 * dest_numeral bs + IntInf.fromInt (dest_bit b)  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
337  | 
  | dest_numeral t = raise TERM ("dest_numeral", [t]);
 | 
| 13755 | 338  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
339  | 
fun number_of_const T = Const ("Numeral.number_of", intT --> T);
 | 
| 
21820
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
340  | 
|
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
341  | 
fun mk_number T 0 = Const ("HOL.zero", T)
 | 
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
342  | 
  | mk_number T 1 = Const ("HOL.one", T)
 | 
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
343  | 
| mk_number T i = number_of_const T $ mk_numeral i;  | 
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
344  | 
|
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
345  | 
fun dest_number (Const ("HOL.zero", T)) = (T, 0)
 | 
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
346  | 
  | dest_number (Const ("HOL.one", T)) = (T, 1)
 | 
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
347  | 
  | dest_number (Const ("Numeral.number_of", Type ("fun", [_, T])) $ t) = (T, dest_numeral t)
 | 
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
348  | 
  | dest_number t = raise TERM ("dest_number", [t]);
 | 
| 13755 | 349  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
350  | 
|
| 13755 | 351  | 
(* real *)  | 
352  | 
||
| 16971 | 353  | 
val realT = Type ("RealDef.real", []);
 | 
| 13755 | 354  | 
|
355  | 
||
| 21755 | 356  | 
(* nibble *)  | 
357  | 
||
358  | 
val nibbleT = Type ("List.nibble", []);
 | 
|
| 13755 | 359  | 
|
| 21755 | 360  | 
fun mk_nibble n =  | 
361  | 
let val s =  | 
|
362  | 
if 0 <= n andalso n <= 9 then chr (n + ord "0")  | 
|
363  | 
else if 10 <= n andalso n <= 15 then chr (n + ord "A" - 10)  | 
|
364  | 
    else raise TERM ("mk_nibble", [])
 | 
|
365  | 
  in Const ("List.nibble.Nibble" ^ s, nibbleT) end;
 | 
|
| 17083 | 366  | 
|
| 21755 | 367  | 
fun dest_nibble t =  | 
368  | 
  let fun err () = raise TERM ("dest_nibble", [t]) in
 | 
|
369  | 
(case try (unprefix "List.nibble.Nibble" o fst o Term.dest_Const) t of  | 
|
370  | 
NONE => err ()  | 
|
371  | 
| SOME c =>  | 
|
372  | 
if size c <> 1 then err ()  | 
|
373  | 
else if "0" <= c andalso c <= "9" then ord c - ord "0"  | 
|
374  | 
else if "A" <= c andalso c <= "F" then ord c - ord "A" + 10  | 
|
375  | 
else err ())  | 
|
376  | 
end;  | 
|
| 15062 | 377  | 
|
| 21455 | 378  | 
|
379  | 
(* char *)  | 
|
380  | 
||
381  | 
val charT = Type ("List.char", []);
 | 
|
382  | 
||
| 21755 | 383  | 
fun mk_char n =  | 
384  | 
if 0 <= n andalso n <= 255 then  | 
|
385  | 
    Const ("List.char.Char", nibbleT --> nibbleT --> charT) $
 | 
|
386  | 
mk_nibble (n div 16) $ mk_nibble (n mod 16)  | 
|
387  | 
  else raise TERM ("mk_char", []);
 | 
|
| 21455 | 388  | 
|
| 21755 | 389  | 
fun dest_char (Const ("List.char.Char", _) $ t $ u) =
 | 
390  | 
dest_nibble t * 16 + dest_nibble u  | 
|
391  | 
  | dest_char t = raise TERM ("dest_char", [t]);
 | 
|
| 21455 | 392  | 
|
| 21755 | 393  | 
|
394  | 
(* list *)  | 
|
| 21455 | 395  | 
|
| 21755 | 396  | 
fun listT T = Type ("List.list", [T]);
 | 
| 21455 | 397  | 
|
| 21755 | 398  | 
fun mk_list T ts =  | 
| 21455 | 399  | 
let  | 
| 21755 | 400  | 
val lT = listT T;  | 
401  | 
    val Nil = Const ("List.list.Nil", lT);
 | 
|
402  | 
    fun Cons t u = Const ("List.list.Cons", T --> lT --> lT) $ t $ u;
 | 
|
403  | 
in fold_rev Cons ts Nil end;  | 
|
404  | 
||
405  | 
fun dest_list (Const ("List.list.Nil", _)) = []
 | 
|
406  | 
  | dest_list (Const ("List.list.Cons", _) $ t $ u) = t :: dest_list u
 | 
|
407  | 
  | dest_list t = raise TERM ("dest_list", [t]);
 | 
|
| 21455 | 408  | 
|
409  | 
||
410  | 
(* string *)  | 
|
411  | 
||
| 
21820
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
412  | 
val stringT = Type ("List.string", []);
 | 
| 21455 | 413  | 
|
| 21755 | 414  | 
val mk_string = mk_list charT o map (mk_char o ord) o explode;  | 
415  | 
val dest_string = implode o map (chr o dest_char) o dest_list;  | 
|
| 21455 | 416  | 
|
| 923 | 417  | 
end;  |