| author | nipkow | 
| Wed, 30 Jul 2008 16:07:00 +0200 | |
| changeset 27712 | 007a339b9e7d | 
| parent 27325 | 70e4eb732fa9 | 
| 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  | 
|
| 
23555
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
21  | 
val conj_intr: thm -> thm -> thm  | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
22  | 
val conj_elim: thm -> thm * thm  | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
23  | 
val conj_elims: thm -> thm list  | 
| 923 | 24  | 
val conj: term  | 
25  | 
val disj: term  | 
|
26  | 
val imp: term  | 
|
| 8429 | 27  | 
val Not: term  | 
| 7690 | 28  | 
val mk_conj: term * term -> term  | 
29  | 
val mk_disj: term * term -> term  | 
|
30  | 
val mk_imp: term * term -> term  | 
|
| 16835 | 31  | 
val mk_not: term -> term  | 
| 15151 | 32  | 
val dest_conj: term -> term list  | 
| 15945 | 33  | 
val dest_disj: term -> term list  | 
| 24958 | 34  | 
val disjuncts: term -> term list  | 
| 4571 | 35  | 
val dest_imp: term -> term * term  | 
| 15151 | 36  | 
val dest_not: term -> term  | 
| 923 | 37  | 
val eq_const: typ -> term  | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
38  | 
val mk_eq: term * term -> term  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
39  | 
val dest_eq: term -> term * term  | 
| 923 | 40  | 
val all_const: typ -> term  | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
41  | 
val mk_all: string * typ * term -> term  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
42  | 
val list_all: (string * typ) list * term -> term  | 
| 923 | 43  | 
val exists_const: typ -> term  | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
44  | 
val mk_exists: string * typ * term -> term  | 
| 15945 | 45  | 
val choice_const: typ -> term  | 
| 923 | 46  | 
val Collect_const: typ -> term  | 
47  | 
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
 | 
48  | 
val class_eq: string  | 
| 923 | 49  | 
val mk_mem: term * term -> term  | 
| 6380 | 50  | 
val dest_mem: term -> term * term  | 
| 11818 | 51  | 
val mk_UNIV: typ -> term  | 
| 2510 | 52  | 
val mk_binop: string -> term * term -> term  | 
53  | 
val mk_binrel: string -> term * term -> term  | 
|
54  | 
val dest_bin: string -> typ -> term -> term * term  | 
|
| 4571 | 55  | 
val unitT: typ  | 
| 9362 | 56  | 
val is_unitT: typ -> bool  | 
| 4571 | 57  | 
val unit: term  | 
58  | 
val is_unit: term -> bool  | 
|
59  | 
val mk_prodT: typ * typ -> typ  | 
|
60  | 
val dest_prodT: typ -> typ * typ  | 
|
| 14048 | 61  | 
val pair_const: typ -> typ -> term  | 
| 4571 | 62  | 
val mk_prod: term * term -> term  | 
63  | 
val dest_prod: term -> term * term  | 
|
64  | 
val mk_fst: term -> term  | 
|
65  | 
val mk_snd: term -> term  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
66  | 
val split_const: typ * typ * typ -> term  | 
| 18285 | 67  | 
val mk_split: term -> term  | 
| 
5096
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
68  | 
val prodT_factors: typ -> typ list  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
69  | 
val mk_tuple: typ -> term list -> term  | 
| 23745 | 70  | 
val dest_tuple: term -> term list  | 
71  | 
val ap_split: typ -> typ -> term -> term  | 
|
72  | 
val prod_factors: term -> int list list  | 
|
73  | 
val dest_tuple': int list list -> term -> term list  | 
|
74  | 
val prodT_factors': int list list -> typ -> typ list  | 
|
75  | 
val ap_split': int list list -> typ -> typ -> term -> term  | 
|
76  | 
val mk_tuple': int list list -> typ -> term list -> term  | 
|
77  | 
val mk_tupleT: int list list -> typ list -> typ  | 
|
78  | 
val strip_split: term -> term * typ list * int list list  | 
|
| 5207 | 79  | 
val natT: typ  | 
80  | 
val zero: term  | 
|
81  | 
val is_zero: term -> bool  | 
|
82  | 
val mk_Suc: term -> term  | 
|
83  | 
val dest_Suc: term -> term  | 
|
| 21621 | 84  | 
val Suc_zero: term  | 
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
23745 
diff
changeset
 | 
85  | 
val mk_nat: int -> term  | 
| 
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
23745 
diff
changeset
 | 
86  | 
val dest_nat: term -> int  | 
| 22994 | 87  | 
val class_size: string  | 
88  | 
val size_const: typ -> term  | 
|
| 26036 | 89  | 
val indexT: typ  | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
90  | 
val intT: typ  | 
| 7073 | 91  | 
val pls_const: term  | 
92  | 
val min_const: term  | 
|
| 
26086
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
93  | 
val bit0_const: term  | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
94  | 
val bit1_const: term  | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
95  | 
val mk_bit: int -> term  | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
96  | 
val dest_bit: term -> int  | 
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
23745 
diff
changeset
 | 
97  | 
val mk_numeral: int -> term  | 
| 
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
23745 
diff
changeset
 | 
98  | 
val dest_numeral: term -> int  | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
99  | 
val number_of_const: typ -> term  | 
| 23269 | 100  | 
val add_numerals: term -> (term * typ) list -> (term * typ) list  | 
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
23745 
diff
changeset
 | 
101  | 
val mk_number: typ -> int -> term  | 
| 
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
23745 
diff
changeset
 | 
102  | 
val dest_number: term -> typ * int  | 
| 
21778
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
103  | 
val realT: typ  | 
| 21755 | 104  | 
val nibbleT: typ  | 
105  | 
val mk_nibble: int -> term  | 
|
106  | 
val dest_nibble: term -> int  | 
|
107  | 
val charT: typ  | 
|
108  | 
val mk_char: int -> term  | 
|
109  | 
val dest_char: term -> int  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
110  | 
val listT: typ -> typ  | 
| 25887 | 111  | 
val nil_const: typ -> term  | 
112  | 
val cons_const: typ -> term  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
113  | 
val mk_list: typ -> term list -> term  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
114  | 
val dest_list: term -> term list  | 
| 
21778
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
115  | 
val stringT: typ  | 
| 
21820
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
116  | 
val mk_string: string -> term  | 
| 
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
117  | 
val dest_string: term -> string  | 
| 923 | 118  | 
end;  | 
119  | 
||
120  | 
structure HOLogic: HOLOGIC =  | 
|
121  | 
struct  | 
|
122  | 
||
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
123  | 
(* HOL syntax *)  | 
| 923 | 124  | 
|
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
125  | 
val typeS: sort = ["HOL.type"];  | 
| 
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
126  | 
val typeT = TypeInfer.anyT typeS;  | 
| 
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
11818 
diff
changeset
 | 
127  | 
|
| 923 | 128  | 
|
| 2510 | 129  | 
(* bool and set *)  | 
| 923 | 130  | 
|
| 8275 | 131  | 
val boolN = "bool";  | 
132  | 
val boolT = Type (boolN, []);  | 
|
| 923 | 133  | 
|
| 9856 | 134  | 
val true_const =  Const ("True", boolT);
 | 
135  | 
val false_const = Const ("False", boolT);
 | 
|
| 7073 | 136  | 
|
| 
26804
 
e2b1e6868c2f
Adapted functions mk_setT and dest_setT to encoding of sets as predicates.
 
berghofe 
parents: 
26086 
diff
changeset
 | 
137  | 
fun mk_setT T = T --> boolT;  | 
| 923 | 138  | 
|
| 
26804
 
e2b1e6868c2f
Adapted functions mk_setT and dest_setT to encoding of sets as predicates.
 
berghofe 
parents: 
26086 
diff
changeset
 | 
139  | 
fun dest_setT (Type ("fun", [T, Type ("bool", [])])) = T
 | 
| 3794 | 140  | 
  | dest_setT T = raise TYPE ("dest_setT: set type expected", [T], []);
 | 
| 923 | 141  | 
|
| 8275 | 142  | 
|
| 7073 | 143  | 
(* logic *)  | 
| 923 | 144  | 
|
145  | 
val Trueprop = Const ("Trueprop", boolT --> propT);
 | 
|
146  | 
||
147  | 
fun mk_Trueprop P = Trueprop $ P;  | 
|
148  | 
||
149  | 
fun dest_Trueprop (Const ("Trueprop", _) $ P) = P
 | 
|
| 3794 | 150  | 
  | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]);
 | 
| 923 | 151  | 
|
| 
23555
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
152  | 
fun conj_intr thP thQ =  | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
153  | 
let  | 
| 
23576
 
1ffe739e5ee0
removed dest_cTrueprop (cf. ObjectLogic.dest_judgment);
 
wenzelm 
parents: 
23566 
diff
changeset
 | 
154  | 
val (P, Q) = pairself (ObjectLogic.dest_judgment o Thm.cprop_of) (thP, thQ)  | 
| 
23555
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
155  | 
handle CTERM (msg, _) => raise THM (msg, 0, [thP, thQ]);  | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
156  | 
    val inst = Thm.instantiate ([], [(@{cpat "?P::bool"}, P), (@{cpat "?Q::bool"}, Q)]);
 | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
157  | 
  in Drule.implies_elim_list (inst @{thm conjI}) [thP, thQ] end;
 | 
| 923 | 158  | 
|
| 
23555
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
159  | 
fun conj_elim thPQ =  | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
160  | 
let  | 
| 
23576
 
1ffe739e5ee0
removed dest_cTrueprop (cf. ObjectLogic.dest_judgment);
 
wenzelm 
parents: 
23566 
diff
changeset
 | 
161  | 
val (P, Q) = Thm.dest_binop (ObjectLogic.dest_judgment (Thm.cprop_of thPQ))  | 
| 
23555
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
162  | 
handle CTERM (msg, _) => raise THM (msg, 0, [thPQ]);  | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
163  | 
    val inst = Thm.instantiate ([], [(@{cpat "?P::bool"}, P), (@{cpat "?Q::bool"}, Q)]);
 | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
164  | 
    val thP = Thm.implies_elim (inst @{thm conjunct1}) thPQ;
 | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
165  | 
    val thQ = Thm.implies_elim (inst @{thm conjunct2}) thPQ;
 | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
166  | 
in (thP, thQ) end;  | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
167  | 
|
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
168  | 
fun conj_elims th =  | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
169  | 
let val (th1, th2) = conj_elim th  | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
170  | 
in conj_elims th1 @ conj_elims th2 end handle THM _ => [th];  | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
171  | 
|
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
172  | 
val conj = @{term "op &"}
 | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
173  | 
and disj = @{term "op |"}
 | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
174  | 
and imp = @{term "op -->"}
 | 
| 
 
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
 
wenzelm 
parents: 
23297 
diff
changeset
 | 
175  | 
and Not = @{term "Not"};
 | 
| 923 | 176  | 
|
| 7690 | 177  | 
fun mk_conj (t1, t2) = conj $ t1 $ t2  | 
178  | 
and mk_disj (t1, t2) = disj $ t1 $ t2  | 
|
| 16835 | 179  | 
and mk_imp (t1, t2) = imp $ t1 $ t2  | 
180  | 
and mk_not t = Not $ t;  | 
|
| 7690 | 181  | 
|
| 15151 | 182  | 
fun dest_conj (Const ("op &", _) $ t $ t') = t :: dest_conj t'
 | 
183  | 
| dest_conj t = [t];  | 
|
184  | 
||
| 15945 | 185  | 
fun dest_disj (Const ("op |", _) $ t $ t') = t :: dest_disj t'
 | 
186  | 
| dest_disj t = [t];  | 
|
187  | 
||
| 24958 | 188  | 
(*Like dest_disj, but flattens disjunctions however nested*)  | 
189  | 
fun disjuncts_aux (Const ("op |", _) $ t $ t') disjs = disjuncts_aux t (disjuncts_aux t' disjs)
 | 
|
190  | 
| disjuncts_aux t disjs = t::disjs;  | 
|
191  | 
||
192  | 
fun disjuncts t = disjuncts_aux t [];  | 
|
193  | 
||
| 
4466
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4294 
diff
changeset
 | 
194  | 
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
 | 
195  | 
  | 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
 | 
196  | 
|
| 15151 | 197  | 
fun dest_not (Const ("Not", _) $ t) = t
 | 
198  | 
  | dest_not t = raise TERM ("dest_not", [t]);
 | 
|
| 8302 | 199  | 
|
| 923 | 200  | 
fun eq_const T = Const ("op =", [T, T] ---> boolT);
 | 
201  | 
fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u;  | 
|
202  | 
||
| 6031 | 203  | 
fun dest_eq (Const ("op =", _) $ lhs $ rhs) = (lhs, rhs)
 | 
204  | 
  | dest_eq t = raise TERM ("dest_eq", [t])
 | 
|
205  | 
||
| 923 | 206  | 
fun all_const T = Const ("All", [T --> boolT] ---> boolT);
 | 
207  | 
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
 | 
208  | 
fun list_all (xs, t) = fold_rev (fn (x, T) => fn P => all_const T $ Abs (x, T, P)) xs t;  | 
| 923 | 209  | 
|
210  | 
fun exists_const T = Const ("Ex", [T --> boolT] ---> boolT);
 | 
|
211  | 
fun mk_exists (x, T, P) = exists_const T $ absfree (x, T, P);  | 
|
212  | 
||
| 21755 | 213  | 
fun choice_const T = Const("Hilbert_Choice.Eps", (T --> boolT) --> T);
 | 
| 15945 | 214  | 
|
| 923 | 215  | 
fun Collect_const T = Const ("Collect", [T --> boolT] ---> mk_setT T);
 | 
216  | 
fun mk_Collect (a, T, t) = Collect_const T $ absfree (a, T, t);  | 
|
217  | 
||
| 23247 | 218  | 
val class_eq = "HOL.eq";  | 
| 21045 | 219  | 
|
| 923 | 220  | 
fun mk_mem (x, A) =  | 
221  | 
let val setT = fastype_of A in  | 
|
222  | 
    Const ("op :", [dest_setT setT, setT] ---> boolT) $ x $ A
 | 
|
223  | 
end;  | 
|
224  | 
||
| 6380 | 225  | 
fun dest_mem (Const ("op :", _) $ x $ A) = (x, A)
 | 
226  | 
  | dest_mem t = raise TERM ("dest_mem", [t]);
 | 
|
227  | 
||
| 11818 | 228  | 
fun mk_UNIV T = Const ("UNIV", mk_setT T);
 | 
229  | 
||
| 923 | 230  | 
|
| 
13743
 
f8f9393be64c
Fixed bug in simpdata.ML that prevented the use of congruence rules from a
 
ballarin 
parents: 
13640 
diff
changeset
 | 
231  | 
(* binary operations and relations *)  | 
| 2510 | 232  | 
|
233  | 
fun mk_binop c (t, u) =  | 
|
234  | 
let val T = fastype_of t in  | 
|
235  | 
Const (c, [T, T] ---> T) $ t $ u  | 
|
236  | 
end;  | 
|
237  | 
||
238  | 
fun mk_binrel c (t, u) =  | 
|
239  | 
let val T = fastype_of t in  | 
|
240  | 
Const (c, [T, T] ---> boolT) $ t $ u  | 
|
241  | 
end;  | 
|
242  | 
||
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14103 
diff
changeset
 | 
243  | 
(*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
 | 
244  | 
way of handling polymorphic operators.*)  | 
| 2510 | 245  | 
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
 | 
246  | 
if c = c' andalso (T=T' orelse T=dummyT) then (t, u)  | 
| 3794 | 247  | 
      else raise TERM ("dest_bin " ^ c, [tm])
 | 
248  | 
  | dest_bin c _ tm = raise TERM ("dest_bin " ^ c, [tm]);
 | 
|
| 2510 | 249  | 
|
250  | 
||
| 4571 | 251  | 
(* unit *)  | 
252  | 
||
| 11604 | 253  | 
val unitT = Type ("Product_Type.unit", []);
 | 
| 4571 | 254  | 
|
| 11604 | 255  | 
fun is_unitT (Type ("Product_Type.unit", [])) = true
 | 
| 9362 | 256  | 
| is_unitT _ = false;  | 
257  | 
||
| 11604 | 258  | 
val unit = Const ("Product_Type.Unity", unitT);
 | 
| 4571 | 259  | 
|
| 11604 | 260  | 
fun is_unit (Const ("Product_Type.Unity", _)) = true
 | 
| 4571 | 261  | 
| is_unit _ = false;  | 
262  | 
||
263  | 
||
264  | 
(* prod *)  | 
|
265  | 
||
266  | 
fun mk_prodT (T1, T2) = Type ("*", [T1, T2]);
 | 
|
267  | 
||
268  | 
fun dest_prodT (Type ("*", [T1, T2])) = (T1, T2)
 | 
|
269  | 
  | dest_prodT T = raise TYPE ("dest_prodT", [T], []);
 | 
|
270  | 
||
| 14048 | 271  | 
fun pair_const T1 T2 = Const ("Pair", [T1, T2] ---> mk_prodT (T1, T2));
 | 
272  | 
||
| 4571 | 273  | 
fun mk_prod (t1, t2) =  | 
274  | 
let val T1 = fastype_of t1 and T2 = fastype_of t2 in  | 
|
| 14048 | 275  | 
pair_const T1 T2 $ t1 $ t2  | 
| 4571 | 276  | 
end;  | 
277  | 
||
278  | 
fun dest_prod (Const ("Pair", _) $ t1 $ t2) = (t1, t2)
 | 
|
279  | 
  | dest_prod t = raise TERM ("dest_prod", [t]);
 | 
|
280  | 
||
281  | 
fun mk_fst p =  | 
|
282  | 
let val pT = fastype_of p in  | 
|
283  | 
    Const ("fst", pT --> fst (dest_prodT pT)) $ p
 | 
|
284  | 
end;  | 
|
285  | 
||
286  | 
fun mk_snd p =  | 
|
287  | 
let val pT = fastype_of p in  | 
|
288  | 
    Const ("snd", pT --> snd (dest_prodT pT)) $ p
 | 
|
289  | 
end;  | 
|
290  | 
||
| 18285 | 291  | 
fun split_const (A, B, C) =  | 
292  | 
  Const ("split", (A --> B --> C) --> mk_prodT (A, B) --> C);
 | 
|
293  | 
||
294  | 
fun mk_split t =  | 
|
295  | 
(case Term.fastype_of t of  | 
|
296  | 
    T as (Type ("fun", [A, Type ("fun", [B, C])])) =>
 | 
|
297  | 
      Const ("split", T --> mk_prodT (A, B) --> C) $ t
 | 
|
298  | 
  | _ => raise TERM ("mk_split: bad body type", [t]));
 | 
|
299  | 
||
| 
5096
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
300  | 
(*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
 | 
301  | 
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
 | 
302  | 
| prodT_factors T = [T];  | 
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
303  | 
|
| 
 
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
 
berghofe 
parents: 
4571 
diff
changeset
 | 
304  | 
(*Makes a nested tuple from a list, following the product type structure*)  | 
| 16971 | 305  | 
fun mk_tuple (Type ("*", [T1, T2])) tms =
 | 
306  | 
mk_prod (mk_tuple T1 tms,  | 
|
| 15570 | 307  | 
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
 | 
308  | 
| mk_tuple T (t::_) = t;  | 
| 4571 | 309  | 
|
| 23745 | 310  | 
fun dest_tuple (Const ("Pair", _) $ t $ u) = dest_tuple t @ dest_tuple u
 | 
311  | 
| dest_tuple t = [t];  | 
|
312  | 
||
313  | 
(*In ap_split S T u, term u expects separate arguments for the factors of S,  | 
|
314  | 
with result type T. The call creates a new term expecting one argument  | 
|
315  | 
of type S.*)  | 
|
316  | 
fun ap_split T T3 u =  | 
|
317  | 
let  | 
|
318  | 
fun ap (T :: Ts) =  | 
|
319  | 
(case T of  | 
|
320  | 
             Type ("*", [T1, T2]) =>
 | 
|
321  | 
split_const (T1, T2, Ts ---> T3) $ ap (T1 :: T2 :: Ts)  | 
|
322  | 
           | _ => Abs ("x", T, ap Ts))
 | 
|
323  | 
| ap [] =  | 
|
324  | 
let val k = length (prodT_factors T)  | 
|
325  | 
in list_comb (incr_boundvars k u, map Bound (k - 1 downto 0)) end  | 
|
326  | 
in ap [T] end;  | 
|
327  | 
||
328  | 
||
| 25172 | 329  | 
(* operations on tuples with specific arities *)  | 
330  | 
(*  | 
|
331  | 
an "arity" of a tuple is a list of lists of integers  | 
|
332  | 
  ("factors"), denoting paths to subterms that are pairs
 | 
|
333  | 
*)  | 
|
| 23745 | 334  | 
|
335  | 
fun prod_err s = raise TERM (s ^ ": inconsistent use of products", []);  | 
|
336  | 
||
337  | 
fun prod_factors t =  | 
|
338  | 
let  | 
|
339  | 
    fun factors p (Const ("Pair", _) $ t $ u) =
 | 
|
340  | 
p :: factors (1::p) t @ factors (2::p) u  | 
|
341  | 
| factors p _ = []  | 
|
342  | 
in factors [] t end;  | 
|
343  | 
||
344  | 
fun dest_tuple' ps =  | 
|
345  | 
let  | 
|
346  | 
fun dest p t = if p mem ps then (case t of  | 
|
347  | 
        Const ("Pair", _) $ t $ u =>
 | 
|
348  | 
dest (1::p) t @ dest (2::p) u  | 
|
349  | 
| _ => prod_err "dest_tuple'") else [t]  | 
|
350  | 
in dest [] end;  | 
|
351  | 
||
352  | 
fun prodT_factors' ps =  | 
|
353  | 
let  | 
|
354  | 
fun factors p T = if p mem ps then (case T of  | 
|
355  | 
        Type ("*", [T1, T2]) =>
 | 
|
356  | 
factors (1::p) T1 @ factors (2::p) T2  | 
|
357  | 
| _ => prod_err "prodT_factors'") else [T]  | 
|
358  | 
in factors [] end;  | 
|
359  | 
||
360  | 
(*In ap_split' ps S T u, term u expects separate arguments for the factors of S,  | 
|
361  | 
with result type T. The call creates a new term expecting one argument  | 
|
362  | 
of type S.*)  | 
|
363  | 
fun ap_split' ps T T3 u =  | 
|
364  | 
let  | 
|
365  | 
fun ap ((p, T) :: pTs) =  | 
|
366  | 
if p mem ps then (case T of  | 
|
367  | 
              Type ("*", [T1, T2]) =>
 | 
|
368  | 
split_const (T1, T2, map snd pTs ---> T3) $  | 
|
369  | 
ap ((1::p, T1) :: (2::p, T2) :: pTs)  | 
|
370  | 
| _ => prod_err "ap_split'")  | 
|
371  | 
          else Abs ("x", T, ap pTs)
 | 
|
372  | 
| ap [] =  | 
|
373  | 
let val k = length ps  | 
|
374  | 
in list_comb (incr_boundvars (k + 1) u, map Bound (k downto 0)) end  | 
|
375  | 
in ap [([], T)] end;  | 
|
376  | 
||
377  | 
fun mk_tuple' ps =  | 
|
378  | 
let  | 
|
379  | 
fun mk p T ts =  | 
|
380  | 
if p mem ps then (case T of  | 
|
381  | 
          Type ("*", [T1, T2]) =>
 | 
|
382  | 
let  | 
|
383  | 
val (t, ts') = mk (1::p) T1 ts;  | 
|
384  | 
val (u, ts'') = mk (2::p) T2 ts'  | 
|
385  | 
in (pair_const T1 T2 $ t $ u, ts'') end  | 
|
386  | 
| _ => prod_err "mk_tuple'")  | 
|
387  | 
else (hd ts, tl ts)  | 
|
388  | 
in fst oo mk [] end;  | 
|
389  | 
||
390  | 
fun mk_tupleT ps =  | 
|
391  | 
let  | 
|
392  | 
fun mk p Ts =  | 
|
393  | 
if p mem ps then  | 
|
394  | 
let  | 
|
395  | 
val (T, Ts') = mk (1::p) Ts;  | 
|
396  | 
val (U, Ts'') = mk (2::p) Ts'  | 
|
397  | 
in (mk_prodT (T, U), Ts'') end  | 
|
398  | 
else (hd Ts, tl Ts)  | 
|
399  | 
in fst o mk [] end;  | 
|
400  | 
||
401  | 
fun strip_split t =  | 
|
402  | 
let  | 
|
403  | 
fun strip [] qs Ts t = (t, Ts, qs)  | 
|
404  | 
      | strip (p :: ps) qs Ts (Const ("split", _) $ t) =
 | 
|
405  | 
strip ((1 :: p) :: (2 :: p) :: ps) (p :: qs) Ts t  | 
|
406  | 
| strip (p :: ps) qs Ts (Abs (s, T, t)) = strip ps qs (T :: Ts) t  | 
|
407  | 
| strip (p :: ps) qs Ts t = strip ps qs  | 
|
408  | 
(hd (binder_types (fastype_of1 (Ts, t))) :: Ts)  | 
|
409  | 
(incr_boundvars 1 t $ Bound 0)  | 
|
410  | 
in strip [[]] [] [] t end;  | 
|
411  | 
||
| 5207 | 412  | 
|
413  | 
(* nat *)  | 
|
414  | 
||
415  | 
val natT = Type ("nat", []);
 | 
|
416  | 
||
| 22994 | 417  | 
val zero = Const ("HOL.zero_class.zero", natT);
 | 
| 5207 | 418  | 
|
| 22994 | 419  | 
fun is_zero (Const ("HOL.zero_class.zero", _)) = true
 | 
| 5207 | 420  | 
| is_zero _ = false;  | 
421  | 
||
422  | 
fun mk_Suc t = Const ("Suc", natT --> natT) $ t;
 | 
|
423  | 
||
424  | 
fun dest_Suc (Const ("Suc", _) $ t) = t
 | 
|
425  | 
  | dest_Suc t = raise TERM ("dest_Suc", [t]);
 | 
|
426  | 
||
| 21621 | 427  | 
val Suc_zero = mk_Suc zero;  | 
428  | 
||
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
23745 
diff
changeset
 | 
429  | 
fun mk_nat n =  | 
| 22994 | 430  | 
let  | 
431  | 
fun mk 0 = zero  | 
|
| 23297 | 432  | 
| mk n = mk_Suc (mk (n - 1));  | 
| 
23576
 
1ffe739e5ee0
removed dest_cTrueprop (cf. ObjectLogic.dest_judgment);
 
wenzelm 
parents: 
23566 
diff
changeset
 | 
433  | 
  in if n < 0 then raise TERM ("mk_nat: negative number", []) else mk n end;
 | 
| 5207 | 434  | 
|
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
23745 
diff
changeset
 | 
435  | 
fun dest_nat (Const ("HOL.zero_class.zero", _)) = 0
 | 
| 23297 | 436  | 
  | dest_nat (Const ("Suc", _) $ t) = dest_nat t + 1
 | 
| 5207 | 437  | 
  | dest_nat t = raise TERM ("dest_nat", [t]);
 | 
438  | 
||
| 22994 | 439  | 
val class_size = "Nat.size";  | 
440  | 
||
441  | 
fun size_const T = Const ("Nat.size_class.size", T --> natT);
 | 
|
442  | 
||
| 5207 | 443  | 
|
| 26036 | 444  | 
(* index *)  | 
445  | 
||
446  | 
val indexT = Type ("Code_Index.index", []);
 | 
|
447  | 
||
448  | 
||
| 
21778
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
449  | 
(* 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
 | 
450  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25887 
diff
changeset
 | 
451  | 
val intT = Type ("Int.int", []);
 | 
| 
21778
 
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
 
wenzelm 
parents: 
21755 
diff
changeset
 | 
452  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25887 
diff
changeset
 | 
453  | 
val pls_const = Const ("Int.Pls", intT)
 | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25887 
diff
changeset
 | 
454  | 
and min_const = Const ("Int.Min", intT)
 | 
| 
26086
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
455  | 
and bit0_const = Const ("Int.Bit0", intT --> intT)
 | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
456  | 
and bit1_const = Const ("Int.Bit1", intT --> intT);
 | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
457  | 
|
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
458  | 
fun mk_bit 0 = bit0_const  | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
459  | 
| mk_bit 1 = bit1_const  | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
460  | 
  | mk_bit _ = raise TERM ("mk_bit", []);
 | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
461  | 
|
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
462  | 
fun dest_bit (Const ("Int.Bit0", _)) = 0
 | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
463  | 
  | dest_bit (Const ("Int.Bit1", _)) = 1
 | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
464  | 
  | dest_bit t = raise TERM ("dest_bit", [t]);
 | 
| 8768 | 465  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
466  | 
fun mk_numeral 0 = pls_const  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
467  | 
| mk_numeral ~1 = min_const  | 
| 
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
468  | 
| mk_numeral i =  | 
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
23745 
diff
changeset
 | 
469  | 
let val (q, r) = Integer.div_mod i 2;  | 
| 
26086
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
470  | 
in mk_bit r $ mk_numeral q end;  | 
| 13755 | 471  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25887 
diff
changeset
 | 
472  | 
fun dest_numeral (Const ("Int.Pls", _)) = 0
 | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25887 
diff
changeset
 | 
473  | 
  | dest_numeral (Const ("Int.Min", _)) = ~1
 | 
| 
26086
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
474  | 
  | dest_numeral (Const ("Int.Bit0", _) $ bs) = 2 * dest_numeral bs
 | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26036 
diff
changeset
 | 
475  | 
  | dest_numeral (Const ("Int.Bit1", _) $ bs) = 2 * dest_numeral bs + 1
 | 
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
476  | 
  | dest_numeral t = raise TERM ("dest_numeral", [t]);
 | 
| 13755 | 477  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25887 
diff
changeset
 | 
478  | 
fun number_of_const T = Const ("Int.number_class.number_of", intT --> T);
 | 
| 
21820
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
479  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25887 
diff
changeset
 | 
480  | 
fun add_numerals (Const ("Int.number_class.number_of", Type (_, [_, T])) $ t) = cons (t, T)
 | 
| 23269 | 481  | 
| add_numerals (t $ u) = add_numerals t #> add_numerals u  | 
482  | 
| add_numerals (Abs (_, _, t)) = add_numerals t  | 
|
483  | 
| add_numerals _ = I;  | 
|
| 22391 | 484  | 
|
| 22994 | 485  | 
fun mk_number T 0 = Const ("HOL.zero_class.zero", T)
 | 
486  | 
  | mk_number T 1 = Const ("HOL.one_class.one", T)
 | 
|
| 
21820
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
487  | 
| 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
 | 
488  | 
|
| 22994 | 489  | 
fun dest_number (Const ("HOL.zero_class.zero", T)) = (T, 0)
 | 
490  | 
  | dest_number (Const ("HOL.one_class.one", T)) = (T, 1)
 | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25887 
diff
changeset
 | 
491  | 
  | dest_number (Const ("Int.number_class.number_of", Type ("fun", [_, T])) $ t) =
 | 
| 22994 | 492  | 
(T, dest_numeral t)  | 
| 
21820
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
493  | 
  | dest_number t = raise TERM ("dest_number", [t]);
 | 
| 13755 | 494  | 
|
| 
21829
 
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
 
wenzelm 
parents: 
21820 
diff
changeset
 | 
495  | 
|
| 13755 | 496  | 
(* real *)  | 
497  | 
||
| 16971 | 498  | 
val realT = Type ("RealDef.real", []);
 | 
| 13755 | 499  | 
|
500  | 
||
| 21755 | 501  | 
(* nibble *)  | 
502  | 
||
503  | 
val nibbleT = Type ("List.nibble", []);
 | 
|
| 13755 | 504  | 
|
| 21755 | 505  | 
fun mk_nibble n =  | 
506  | 
let val s =  | 
|
507  | 
if 0 <= n andalso n <= 9 then chr (n + ord "0")  | 
|
508  | 
else if 10 <= n andalso n <= 15 then chr (n + ord "A" - 10)  | 
|
509  | 
    else raise TERM ("mk_nibble", [])
 | 
|
510  | 
  in Const ("List.nibble.Nibble" ^ s, nibbleT) end;
 | 
|
| 17083 | 511  | 
|
| 21755 | 512  | 
fun dest_nibble t =  | 
513  | 
  let fun err () = raise TERM ("dest_nibble", [t]) in
 | 
|
514  | 
(case try (unprefix "List.nibble.Nibble" o fst o Term.dest_Const) t of  | 
|
515  | 
NONE => err ()  | 
|
516  | 
| SOME c =>  | 
|
517  | 
if size c <> 1 then err ()  | 
|
518  | 
else if "0" <= c andalso c <= "9" then ord c - ord "0"  | 
|
519  | 
else if "A" <= c andalso c <= "F" then ord c - ord "A" + 10  | 
|
520  | 
else err ())  | 
|
521  | 
end;  | 
|
| 15062 | 522  | 
|
| 21455 | 523  | 
|
524  | 
(* char *)  | 
|
525  | 
||
526  | 
val charT = Type ("List.char", []);
 | 
|
527  | 
||
| 21755 | 528  | 
fun mk_char n =  | 
529  | 
if 0 <= n andalso n <= 255 then  | 
|
530  | 
    Const ("List.char.Char", nibbleT --> nibbleT --> charT) $
 | 
|
531  | 
mk_nibble (n div 16) $ mk_nibble (n mod 16)  | 
|
532  | 
  else raise TERM ("mk_char", []);
 | 
|
| 21455 | 533  | 
|
| 21755 | 534  | 
fun dest_char (Const ("List.char.Char", _) $ t $ u) =
 | 
535  | 
dest_nibble t * 16 + dest_nibble u  | 
|
536  | 
  | dest_char t = raise TERM ("dest_char", [t]);
 | 
|
| 21455 | 537  | 
|
| 21755 | 538  | 
|
539  | 
(* list *)  | 
|
| 21455 | 540  | 
|
| 21755 | 541  | 
fun listT T = Type ("List.list", [T]);
 | 
| 21455 | 542  | 
|
| 25887 | 543  | 
fun nil_const T = Const ("List.list.Nil", listT T);
 | 
544  | 
||
545  | 
fun cons_const T =  | 
|
546  | 
let val lT = listT T  | 
|
547  | 
  in Const ("List.list.Cons", T --> lT --> lT) end;
 | 
|
548  | 
||
| 21755 | 549  | 
fun mk_list T ts =  | 
| 21455 | 550  | 
let  | 
| 21755 | 551  | 
val lT = listT T;  | 
552  | 
    val Nil = Const ("List.list.Nil", lT);
 | 
|
553  | 
    fun Cons t u = Const ("List.list.Cons", T --> lT --> lT) $ t $ u;
 | 
|
554  | 
in fold_rev Cons ts Nil end;  | 
|
555  | 
||
556  | 
fun dest_list (Const ("List.list.Nil", _)) = []
 | 
|
557  | 
  | dest_list (Const ("List.list.Cons", _) $ t $ u) = t :: dest_list u
 | 
|
558  | 
  | dest_list t = raise TERM ("dest_list", [t]);
 | 
|
| 21455 | 559  | 
|
560  | 
||
561  | 
(* string *)  | 
|
562  | 
||
| 
21820
 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 
haftmann 
parents: 
21788 
diff
changeset
 | 
563  | 
val stringT = Type ("List.string", []);
 | 
| 21455 | 564  | 
|
| 21755 | 565  | 
val mk_string = mk_list charT o map (mk_char o ord) o explode;  | 
566  | 
val dest_string = implode o map (chr o dest_char) o dest_list;  | 
|
| 21455 | 567  | 
|
| 923 | 568  | 
end;  |