| author | huffman | 
| Sat, 27 Nov 2010 16:08:10 -0800 | |
| changeset 40774 | 0437dbc127b3 | 
| parent 40327 | src/HOLCF/Tools/holcf_library.ML@1dfdbd66093a | 
| child 40832 | 4352ca878c41 | 
| permissions | -rw-r--r-- | 
| 35475 | 1  | 
(* Title: HOLCF/Tools/holcf_library.ML  | 
2  | 
Author: Brian Huffman  | 
|
3  | 
||
4  | 
Functions for constructing HOLCF types and terms.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
structure HOLCF_Library =  | 
|
8  | 
struct  | 
|
9  | 
||
| 35476 | 10  | 
infixr 6 ->>;  | 
| 40217 | 11  | 
infixr -->>;  | 
| 35555 | 12  | 
infix 9 `;  | 
| 35476 | 13  | 
|
| 35475 | 14  | 
(*** Operations from Isabelle/HOL ***)  | 
15  | 
||
16  | 
val boolT = HOLogic.boolT;  | 
|
| 35489 | 17  | 
val natT = HOLogic.natT;  | 
| 35475 | 18  | 
|
19  | 
val mk_equals = Logic.mk_equals;  | 
|
20  | 
val mk_eq = HOLogic.mk_eq;  | 
|
21  | 
val mk_trp = HOLogic.mk_Trueprop;  | 
|
22  | 
val mk_fst = HOLogic.mk_fst;  | 
|
23  | 
val mk_snd = HOLogic.mk_snd;  | 
|
24  | 
val mk_not = HOLogic.mk_not;  | 
|
25  | 
val mk_conj = HOLogic.mk_conj;  | 
|
26  | 
val mk_disj = HOLogic.mk_disj;  | 
|
| 40024 | 27  | 
val mk_imp = HOLogic.mk_imp;  | 
| 35475 | 28  | 
|
29  | 
fun mk_ex (x, t) = HOLogic.exists_const (fastype_of x) $ Term.lambda x t;  | 
|
| 40024 | 30  | 
fun mk_all (x, t) = HOLogic.all_const (fastype_of x) $ Term.lambda x t;  | 
| 35475 | 31  | 
|
32  | 
||
33  | 
(*** Basic HOLCF concepts ***)  | 
|
34  | 
||
35  | 
fun mk_bottom T = Const (@{const_name UU}, T);
 | 
|
36  | 
||
37  | 
fun below_const T = Const (@{const_name below}, [T, T] ---> boolT);
 | 
|
38  | 
fun mk_below (t, u) = below_const (fastype_of t) $ t $ u;  | 
|
39  | 
||
40  | 
fun mk_undef t = mk_eq (t, mk_bottom (fastype_of t));  | 
|
41  | 
||
42  | 
fun mk_defined t = mk_not (mk_undef t);  | 
|
43  | 
||
| 40021 | 44  | 
fun mk_adm t =  | 
45  | 
  Const (@{const_name adm}, fastype_of t --> boolT) $ t;
 | 
|
46  | 
||
| 35475 | 47  | 
fun mk_compact t =  | 
48  | 
  Const (@{const_name compact}, fastype_of t --> boolT) $ t;
 | 
|
49  | 
||
50  | 
fun mk_cont t =  | 
|
51  | 
  Const (@{const_name cont}, fastype_of t --> boolT) $ t;
 | 
|
52  | 
||
| 35489 | 53  | 
fun mk_chain t =  | 
54  | 
  Const (@{const_name chain}, Term.fastype_of t --> boolT) $ t;
 | 
|
55  | 
||
| 35555 | 56  | 
fun mk_lub t =  | 
57  | 
let  | 
|
58  | 
val T = Term.range_type (Term.fastype_of t);  | 
|
59  | 
    val lub_const = Const (@{const_name lub}, (T --> boolT) --> T);
 | 
|
60  | 
    val UNIV_const = @{term "UNIV :: nat set"};
 | 
|
61  | 
val image_type = (natT --> T) --> (natT --> boolT) --> T --> boolT;  | 
|
62  | 
    val image_const = Const (@{const_name image}, image_type);
 | 
|
63  | 
in  | 
|
64  | 
lub_const $ (image_const $ t $ UNIV_const)  | 
|
65  | 
end;  | 
|
66  | 
||
| 35475 | 67  | 
|
68  | 
(*** Continuous function space ***)  | 
|
69  | 
||
| 35525 | 70  | 
fun mk_cfunT (T, U) = Type(@{type_name cfun}, [T, U]);
 | 
| 35475 | 71  | 
|
| 35476 | 72  | 
val (op ->>) = mk_cfunT;  | 
73  | 
val (op -->>) = Library.foldr mk_cfunT;  | 
|
| 35475 | 74  | 
|
| 35525 | 75  | 
fun dest_cfunT (Type(@{type_name cfun}, [T, U])) = (T, U)
 | 
| 35475 | 76  | 
  | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []);
 | 
77  | 
||
78  | 
fun capply_const (S, T) =  | 
|
| 40327 | 79  | 
  Const(@{const_name Rep_cfun}, (S ->> T) --> (S --> T));
 | 
| 35475 | 80  | 
|
81  | 
fun cabs_const (S, T) =  | 
|
| 40327 | 82  | 
  Const(@{const_name Abs_cfun}, (S --> T) --> (S ->> T));
 | 
| 35475 | 83  | 
|
84  | 
fun mk_cabs t =  | 
|
85  | 
let val T = fastype_of t  | 
|
86  | 
in cabs_const (Term.domain_type T, Term.range_type T) $ t end  | 
|
87  | 
||
88  | 
(* builds the expression (% v1 v2 .. vn. rhs) *)  | 
|
89  | 
fun lambdas [] rhs = rhs  | 
|
90  | 
| lambdas (v::vs) rhs = Term.lambda v (lambdas vs rhs);  | 
|
91  | 
||
92  | 
(* builds the expression (LAM v. rhs) *)  | 
|
93  | 
fun big_lambda v rhs =  | 
|
94  | 
cabs_const (fastype_of v, fastype_of rhs) $ Term.lambda v rhs;  | 
|
95  | 
||
96  | 
(* builds the expression (LAM v1 v2 .. vn. rhs) *)  | 
|
97  | 
fun big_lambdas [] rhs = rhs  | 
|
98  | 
| big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs);  | 
|
99  | 
||
100  | 
fun mk_capply (t, u) =  | 
|
101  | 
let val (S, T) =  | 
|
102  | 
case fastype_of t of  | 
|
| 35525 | 103  | 
        Type(@{type_name cfun}, [S, T]) => (S, T)
 | 
| 35475 | 104  | 
      | _ => raise TERM ("mk_capply " ^ ML_Syntax.print_list ML_Syntax.print_term [t, u], [t, u]);
 | 
105  | 
in capply_const (S, T) $ t $ u end;  | 
|
106  | 
||
| 35555 | 107  | 
val (op `) = mk_capply;  | 
| 35475 | 108  | 
|
109  | 
val list_ccomb : term * term list -> term = Library.foldl mk_capply;  | 
|
110  | 
||
111  | 
fun mk_ID T = Const (@{const_name ID}, T ->> T);
 | 
|
112  | 
||
113  | 
fun cfcomp_const (T, U, V) =  | 
|
114  | 
  Const (@{const_name cfcomp}, (U ->> V) ->> (T ->> U) ->> (T ->> V));
 | 
|
115  | 
||
116  | 
fun mk_cfcomp (f, g) =  | 
|
117  | 
let  | 
|
118  | 
val (U, V) = dest_cfunT (fastype_of f);  | 
|
119  | 
val (T, U') = dest_cfunT (fastype_of g);  | 
|
120  | 
in  | 
|
121  | 
if U = U'  | 
|
122  | 
then mk_capply (mk_capply (cfcomp_const (T, U, V), f), g)  | 
|
123  | 
    else raise TYPE ("mk_cfcomp", [U, U'], [f, g])
 | 
|
124  | 
end;  | 
|
125  | 
||
| 35785 | 126  | 
fun strictify_const T = Const (@{const_name strictify}, T ->> T);
 | 
127  | 
fun mk_strictify t = strictify_const (fastype_of t) ` t;  | 
|
128  | 
||
| 35475 | 129  | 
fun mk_strict t =  | 
130  | 
let val (T, U) = dest_cfunT (fastype_of t);  | 
|
131  | 
in mk_eq (t ` mk_bottom T, mk_bottom U) end;  | 
|
132  | 
||
133  | 
||
134  | 
(*** Product type ***)  | 
|
135  | 
||
136  | 
val mk_prodT = HOLogic.mk_prodT  | 
|
137  | 
||
138  | 
fun mk_tupleT [] = HOLogic.unitT  | 
|
139  | 
| mk_tupleT [T] = T  | 
|
140  | 
| mk_tupleT (T :: Ts) = mk_prodT (T, mk_tupleT Ts);  | 
|
141  | 
||
142  | 
(* builds the expression (v1,v2,..,vn) *)  | 
|
143  | 
fun mk_tuple [] = HOLogic.unit  | 
|
144  | 
| mk_tuple (t::[]) = t  | 
|
145  | 
| mk_tuple (t::ts) = HOLogic.mk_prod (t, mk_tuple ts);  | 
|
146  | 
||
147  | 
(* builds the expression (%(v1,v2,..,vn). rhs) *)  | 
|
148  | 
fun lambda_tuple [] rhs = Term.lambda (Free("unit", HOLogic.unitT)) rhs
 | 
|
149  | 
| lambda_tuple (v::[]) rhs = Term.lambda v rhs  | 
|
150  | 
| lambda_tuple (v::vs) rhs =  | 
|
151  | 
HOLogic.mk_split (Term.lambda v (lambda_tuple vs rhs));  | 
|
152  | 
||
153  | 
||
154  | 
(*** Lifted cpo type ***)  | 
|
155  | 
||
156  | 
fun mk_upT T = Type(@{type_name "u"}, [T]);
 | 
|
157  | 
||
158  | 
fun dest_upT (Type(@{type_name "u"}, [T])) = T
 | 
|
159  | 
  | dest_upT T = raise TYPE ("dest_upT", [T], []);
 | 
|
160  | 
||
161  | 
fun up_const T = Const(@{const_name up}, T ->> mk_upT T);
 | 
|
162  | 
||
163  | 
fun mk_up t = up_const (fastype_of t) ` t;  | 
|
164  | 
||
165  | 
fun fup_const (T, U) =  | 
|
166  | 
  Const(@{const_name fup}, (T ->> U) ->> mk_upT T ->> U);
 | 
|
167  | 
||
| 35785 | 168  | 
fun mk_fup t = fup_const (dest_cfunT (fastype_of t)) ` t;  | 
169  | 
||
| 35475 | 170  | 
fun from_up T = fup_const (T, T) ` mk_ID T;  | 
171  | 
||
172  | 
||
| 35785 | 173  | 
(*** Lifted unit type ***)  | 
| 35475 | 174  | 
|
175  | 
val oneT = @{typ "one"};
 | 
|
176  | 
||
| 40212 | 177  | 
fun one_case_const T = Const (@{const_name one_case}, T ->> oneT ->> T);
 | 
178  | 
fun mk_one_case t = one_case_const (fastype_of t) ` t;  | 
|
| 35785 | 179  | 
|
180  | 
||
181  | 
(*** Strict product type ***)  | 
|
182  | 
||
| 35525 | 183  | 
fun mk_sprodT (T, U) = Type(@{type_name sprod}, [T, U]);
 | 
| 35475 | 184  | 
|
| 35525 | 185  | 
fun dest_sprodT (Type(@{type_name sprod}, [T, U])) = (T, U)
 | 
| 35475 | 186  | 
  | dest_sprodT T = raise TYPE ("dest_sprodT", [T], []);
 | 
187  | 
||
188  | 
fun spair_const (T, U) =  | 
|
189  | 
  Const(@{const_name spair}, T ->> U ->> mk_sprodT (T, U));
 | 
|
190  | 
||
191  | 
(* builds the expression (:t, u:) *)  | 
|
192  | 
fun mk_spair (t, u) =  | 
|
193  | 
spair_const (fastype_of t, fastype_of u) ` t ` u;  | 
|
194  | 
||
195  | 
(* builds the expression (:t1,t2,..,tn:) *)  | 
|
196  | 
fun mk_stuple [] = @{term "ONE"}
 | 
|
197  | 
| mk_stuple (t::[]) = t  | 
|
198  | 
| mk_stuple (t::ts) = mk_spair (t, mk_stuple ts);  | 
|
199  | 
||
200  | 
fun sfst_const (T, U) =  | 
|
201  | 
  Const(@{const_name sfst}, mk_sprodT (T, U) ->> T);
 | 
|
202  | 
||
203  | 
fun ssnd_const (T, U) =  | 
|
204  | 
  Const(@{const_name ssnd}, mk_sprodT (T, U) ->> U);
 | 
|
205  | 
||
| 35785 | 206  | 
fun ssplit_const (T, U, V) =  | 
207  | 
  Const (@{const_name ssplit}, (T ->> U ->> V) ->> mk_sprodT (T, U) ->> V);
 | 
|
208  | 
||
209  | 
fun mk_ssplit t =  | 
|
210  | 
let val (T, (U, V)) = apsnd dest_cfunT (dest_cfunT (fastype_of t));  | 
|
211  | 
in ssplit_const (T, U, V) ` t end;  | 
|
212  | 
||
| 35475 | 213  | 
|
214  | 
(*** Strict sum type ***)  | 
|
215  | 
||
| 35525 | 216  | 
fun mk_ssumT (T, U) = Type(@{type_name ssum}, [T, U]);
 | 
| 35475 | 217  | 
|
| 35525 | 218  | 
fun dest_ssumT (Type(@{type_name ssum}, [T, U])) = (T, U)
 | 
| 35475 | 219  | 
  | dest_ssumT T = raise TYPE ("dest_ssumT", [T], []);
 | 
220  | 
||
221  | 
fun sinl_const (T, U) = Const(@{const_name sinl}, T ->> mk_ssumT (T, U));
 | 
|
222  | 
fun sinr_const (T, U) = Const(@{const_name sinr}, U ->> mk_ssumT (T, U));
 | 
|
223  | 
||
224  | 
(* builds the list [sinl(t1), sinl(sinr(t2)), ... sinr(...sinr(tn))] *)  | 
|
225  | 
fun mk_sinjects ts =  | 
|
226  | 
let  | 
|
227  | 
val Ts = map fastype_of ts;  | 
|
228  | 
fun combine (t, T) (us, U) =  | 
|
229  | 
let  | 
|
230  | 
val v = sinl_const (T, U) ` t;  | 
|
231  | 
val vs = map (fn u => sinr_const (T, U) ` u) us;  | 
|
232  | 
in  | 
|
233  | 
(v::vs, mk_ssumT (T, U))  | 
|
234  | 
end  | 
|
| 35912 | 235  | 
fun inj [] = raise Fail "mk_sinjects: empty list"  | 
| 35475 | 236  | 
| inj ((t, T)::[]) = ([t], T)  | 
237  | 
| inj ((t, T)::ts) = combine (t, T) (inj ts);  | 
|
238  | 
in  | 
|
239  | 
fst (inj (ts ~~ Ts))  | 
|
240  | 
end;  | 
|
241  | 
||
242  | 
fun sscase_const (T, U, V) =  | 
|
243  | 
  Const(@{const_name sscase},
 | 
|
244  | 
(T ->> V) ->> (U ->> V) ->> mk_ssumT (T, U) ->> V);  | 
|
245  | 
||
| 35785 | 246  | 
fun mk_sscase (t, u) =  | 
247  | 
let val (T, V) = dest_cfunT (fastype_of t);  | 
|
248  | 
val (U, V) = dest_cfunT (fastype_of u);  | 
|
249  | 
in sscase_const (T, U, V) ` t ` u end;  | 
|
250  | 
||
| 35475 | 251  | 
fun from_sinl (T, U) =  | 
252  | 
sscase_const (T, U, T) ` mk_ID T ` mk_bottom (U ->> T);  | 
|
253  | 
||
254  | 
fun from_sinr (T, U) =  | 
|
255  | 
sscase_const (T, U, U) ` mk_bottom (T ->> U) ` mk_ID U;  | 
|
256  | 
||
257  | 
||
258  | 
(*** pattern match monad type ***)  | 
|
259  | 
||
| 
37108
 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 
huffman 
parents: 
35912 
diff
changeset
 | 
260  | 
fun mk_matchT T = Type (@{type_name "match"}, [T]);
 | 
| 35475 | 261  | 
|
| 
37108
 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 
huffman 
parents: 
35912 
diff
changeset
 | 
262  | 
fun dest_matchT (Type(@{type_name "match"}, [T])) = T
 | 
| 35475 | 263  | 
  | dest_matchT T = raise TYPE ("dest_matchT", [T], []);
 | 
264  | 
||
265  | 
fun mk_fail T = Const (@{const_name "Fixrec.fail"}, mk_matchT T);
 | 
|
266  | 
||
| 
37108
 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 
huffman 
parents: 
35912 
diff
changeset
 | 
267  | 
fun succeed_const T = Const (@{const_name "Fixrec.succeed"}, T ->> mk_matchT T);
 | 
| 
 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 
huffman 
parents: 
35912 
diff
changeset
 | 
268  | 
fun mk_succeed t = succeed_const (fastype_of t) ` t;  | 
| 35475 | 269  | 
|
270  | 
||
271  | 
(*** lifted boolean type ***)  | 
|
272  | 
||
273  | 
val trT = @{typ "tr"};
 | 
|
274  | 
||
275  | 
||
276  | 
(*** theory of fixed points ***)  | 
|
277  | 
||
278  | 
fun mk_fix t =  | 
|
279  | 
let val (T, _) = dest_cfunT (fastype_of t)  | 
|
280  | 
  in mk_capply (Const(@{const_name fix}, (T ->> T) ->> T), t) end;
 | 
|
281  | 
||
| 35489 | 282  | 
fun iterate_const T =  | 
283  | 
  Const (@{const_name iterate}, natT --> (T ->> T) ->> (T ->> T));
 | 
|
284  | 
||
285  | 
fun mk_iterate (n, f) =  | 
|
286  | 
let val (T, _) = dest_cfunT (Term.fastype_of f);  | 
|
287  | 
in (iterate_const T $ n) ` f ` mk_bottom T end;  | 
|
| 35475 | 288  | 
|
289  | 
end;  |