author | wenzelm |
Sat, 13 May 2006 02:51:49 +0200 | |
changeset 19636 | 50515882049e |
parent 19092 | e32cf29f01fc |
child 22675 | acf10be7dcca |
permissions | -rw-r--r-- |
2276 | 1 |
(* Title: HOLCF/domain/library.ML |
2445
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents:
2276
diff
changeset
|
2 |
ID: $Id$ |
12030 | 3 |
Author: David von Oheimb |
2445
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents:
2276
diff
changeset
|
4 |
|
12030 | 5 |
Library for domain section. |
1274 | 6 |
*) |
7 |
||
2446
c2a9bf6c0948
The previous log message was wrong. The correct one is:
oheimb
parents:
2445
diff
changeset
|
8 |
|
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
9 |
(* ----- general support ---------------------------------------------------- *) |
1274 | 10 |
|
11 |
fun mapn f n [] = [] |
|
12 |
| mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs; |
|
13 |
||
15570 | 14 |
fun foldr'' f (l,f2) = let fun itr [] = raise Fail "foldr''" |
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
15 |
| itr [a] = f2 a |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
16 |
| itr (a::l) = f(a, itr l) |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
17 |
in itr l end; |
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
18 |
fun map_cumulr f start xs = foldr (fn (x,(ys,res))=>case f(x,res) of (y,res2) => |
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
19 |
(y::ys,res2)) ([],start) xs; |
1274 | 20 |
|
21 |
||
22 |
fun first (x,_,_) = x; fun second (_,x,_) = x; fun third (_,_,x) = x; |
|
23 |
fun upd_first f (x,y,z) = (f x, y, z); |
|
24 |
fun upd_second f (x,y,z) = ( x, f y, z); |
|
25 |
fun upd_third f (x,y,z) = ( x, y, f z); |
|
26 |
||
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
27 |
fun atomize thm = let val r_inst = read_instantiate; |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
28 |
fun at thm = case concl_of thm of |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
29 |
_$(Const("op &",_)$_$_) => at(thm RS conjunct1)@at(thm RS conjunct2) |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
30 |
| _$(Const("All" ,_)$Abs(s,_,_))=> at(thm RS (r_inst [("x","?"^s)] spec)) |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
31 |
| _ => [thm]; |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
32 |
in map zero_var_indexes (at thm) end; |
1274 | 33 |
|
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
34 |
(* ----- specific support for domain ---------------------------------------- *) |
1274 | 35 |
|
36 |
structure Domain_Library = struct |
|
37 |
||
4122 | 38 |
open HOLCFLogic; |
39 |
||
1274 | 40 |
exception Impossible of string; |
41 |
fun Imposs msg = raise Impossible ("Domain:"^msg); |
|
42 |
||
43 |
(* ----- name handling ----- *) |
|
44 |
||
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
45 |
val strip_esc = let fun strip ("'" :: c :: cs) = c :: strip cs |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
46 |
| strip ["'"] = [] |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
47 |
| strip (c :: cs) = c :: strip cs |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
48 |
| strip [] = []; |
4709 | 49 |
in implode o strip o Symbol.explode end; |
1274 | 50 |
|
4709 | 51 |
fun extern_name con = case Symbol.explode con of |
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
52 |
("o"::"p"::" "::rest) => implode rest |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
53 |
| _ => con; |
1274 | 54 |
fun dis_name con = "is_"^ (extern_name con); |
55 |
fun dis_name_ con = "is_"^ (strip_esc con); |
|
16224
57094b83774e
Domain package generates match functions for new datatypes, for use with the fixrec package
huffman
parents:
15638
diff
changeset
|
56 |
fun mat_name con = "match_"^ (extern_name con); |
57094b83774e
Domain package generates match functions for new datatypes, for use with the fixrec package
huffman
parents:
15638
diff
changeset
|
57 |
fun mat_name_ con = "match_"^ (strip_esc con); |
18113 | 58 |
fun pat_name con = (extern_name con) ^ "_pat"; |
59 |
fun pat_name_ con = (strip_esc con) ^ "_pat"; |
|
1274 | 60 |
|
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
61 |
(* make distinct names out of the type list, |
5292
92ea5ff34c79
minor correction: n must not be used as free variable
oheimb
parents:
5291
diff
changeset
|
62 |
forbidding "o","n..","x..","f..","P.." as names *) |
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
63 |
(* a number string is added if necessary *) |
4122 | 64 |
fun mk_var_names ids : string list = let |
5292
92ea5ff34c79
minor correction: n must not be used as free variable
oheimb
parents:
5291
diff
changeset
|
65 |
fun nonreserved s = if s mem ["n","x","f","P"] then s^"'" else s; |
1819
245721624c8d
minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents:
1637
diff
changeset
|
66 |
fun index_vnames(vn::vns,occupied) = |
17325 | 67 |
(case AList.lookup (op =) occupied vn of |
15531 | 68 |
NONE => if vn mem vns |
1819
245721624c8d
minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents:
1637
diff
changeset
|
69 |
then (vn^"1") :: index_vnames(vns,(vn,1) ::occupied) |
245721624c8d
minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents:
1637
diff
changeset
|
70 |
else vn :: index_vnames(vns, occupied) |
15531 | 71 |
| SOME(i) => (vn^(string_of_int (i+1))) |
1819
245721624c8d
minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents:
1637
diff
changeset
|
72 |
:: index_vnames(vns,(vn,i+1)::occupied)) |
245721624c8d
minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents:
1637
diff
changeset
|
73 |
| index_vnames([],occupied) = []; |
4122 | 74 |
in index_vnames(map nonreserved ids, [("O",0),("o",0)]) end; |
1274 | 75 |
|
7652 | 76 |
fun pcpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, pcpoS); |
4008 | 77 |
fun string_of_typ sg = Sign.string_of_typ sg o Sign.certify_typ sg; |
78 |
fun str2typ sg = typ_of o read_ctyp sg; |
|
1274 | 79 |
|
80 |
(* ----- constructor list handling ----- *) |
|
81 |
||
18085
ec9e36ece6bb
improve support for mutual recursion: now generates correct copy constant and induction theorem for mutually-recursive types; initial support for indirect recursion
huffman
parents:
18083
diff
changeset
|
82 |
type cons = (string * (* operator name of constr *) |
ec9e36ece6bb
improve support for mutual recursion: now generates correct copy constant and induction theorem for mutually-recursive types; initial support for indirect recursion
huffman
parents:
18083
diff
changeset
|
83 |
((bool*int*DatatypeAux.dtyp)* (* (lazy,recursive element or ~1) *) |
ec9e36ece6bb
improve support for mutual recursion: now generates correct copy constant and induction theorem for mutually-recursive types; initial support for indirect recursion
huffman
parents:
18083
diff
changeset
|
84 |
string option* (* selector name *) |
ec9e36ece6bb
improve support for mutual recursion: now generates correct copy constant and induction theorem for mutually-recursive types; initial support for indirect recursion
huffman
parents:
18083
diff
changeset
|
85 |
string) (* argument name *) |
ec9e36ece6bb
improve support for mutual recursion: now generates correct copy constant and induction theorem for mutually-recursive types; initial support for indirect recursion
huffman
parents:
18083
diff
changeset
|
86 |
list); (* argument list *) |
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
87 |
type eq = (string * (* name of abstracted type *) |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
88 |
typ list) * (* arguments of abstracted type *) |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
89 |
cons list; (* represented type, as a constructor list *) |
1274 | 90 |
|
18085
ec9e36ece6bb
improve support for mutual recursion: now generates correct copy constant and induction theorem for mutually-recursive types; initial support for indirect recursion
huffman
parents:
18083
diff
changeset
|
91 |
fun rec_of arg = second (first arg); |
ec9e36ece6bb
improve support for mutual recursion: now generates correct copy constant and induction theorem for mutually-recursive types; initial support for indirect recursion
huffman
parents:
18083
diff
changeset
|
92 |
fun is_lazy arg = first (first arg); |
1274 | 93 |
val sel_of = second; |
94 |
val vname = third; |
|
95 |
val upd_vname = upd_third; |
|
96 |
fun is_rec arg = rec_of arg >=0; |
|
97 |
fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg); |
|
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
98 |
fun nonlazy args = map vname (filter_out is_lazy args); |
15570 | 99 |
fun nonlazy_rec args = map vname (List.filter is_nonlazy_rec args); |
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
100 |
|
16778 | 101 |
(* ----- qualified names of HOLCF constants ----- *) |
102 |
||
17810
3bdf516d93d8
cleaned up; renamed "Porder.op <<" to "Porder.<<"
huffman
parents:
17325
diff
changeset
|
103 |
val lessN = "Porder.<<" |
16778 | 104 |
val UU_N = "Pcpo.UU"; |
105 |
val admN = "Adm.adm"; |
|
17840
11bcd77cfa22
domain package generates compactness lemmas for new constructors
huffman
parents:
17811
diff
changeset
|
106 |
val compactN = "Adm.compact"; |
16778 | 107 |
val Rep_CFunN = "Cfun.Rep_CFun"; |
108 |
val Abs_CFunN = "Cfun.Abs_CFun"; |
|
109 |
val ID_N = "Cfun.ID"; |
|
110 |
val cfcompN = "Cfun.cfcomp"; |
|
111 |
val strictifyN = "Cfun.strictify"; |
|
112 |
val cpairN = "Cprod.cpair"; |
|
113 |
val cfstN = "Cprod.cfst"; |
|
114 |
val csndN = "Cprod.csnd"; |
|
115 |
val csplitN = "Cprod.csplit"; |
|
116 |
val spairN = "Sprod.spair"; |
|
117 |
val sfstN = "Sprod.sfst"; |
|
118 |
val ssndN = "Sprod.ssnd"; |
|
119 |
val ssplitN = "Sprod.ssplit"; |
|
120 |
val sinlN = "Ssum.sinl"; |
|
121 |
val sinrN = "Ssum.sinr"; |
|
122 |
val sscaseN = "Ssum.sscase"; |
|
123 |
val upN = "Up.up"; |
|
124 |
val fupN = "Up.fup"; |
|
125 |
val ONE_N = "One.ONE"; |
|
126 |
val TT_N = "Tr.TT"; |
|
127 |
val FF_N = "Tr.FF"; |
|
128 |
val iterateN = "Fix.iterate"; |
|
129 |
val fixN = "Fix.fix"; |
|
130 |
val returnN = "Fixrec.return"; |
|
131 |
val failN = "Fixrec.fail"; |
|
18113 | 132 |
val cpair_patN = "Fixrec.cpair_pat"; |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18113
diff
changeset
|
133 |
val branchN = "Fixrec.branch"; |
16778 | 134 |
|
135 |
val pcpoN = "Pcpo.pcpo" |
|
136 |
val pcpoS = [pcpoN]; |
|
137 |
||
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
138 |
(* ----- support for type and mixfix expressions ----- *) |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
139 |
|
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
140 |
infixr 5 -->; |
1274 | 141 |
|
142 |
(* ----- support for term expressions ----- *) |
|
143 |
||
11531
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
144 |
fun %: s = Free(s,dummyT); |
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
145 |
fun %# arg = %:(vname arg); |
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
146 |
fun %%: s = Const(s,dummyT); |
1274 | 147 |
|
148 |
local open HOLogic in |
|
149 |
val mk_trp = mk_Trueprop; |
|
150 |
fun mk_conj (S,T) = conj $ S $ T; |
|
151 |
fun mk_disj (S,T) = disj $ S $ T; |
|
152 |
fun mk_imp (S,T) = imp $ S $ T; |
|
153 |
fun mk_lam (x,T) = Abs(x,dummyT,T); |
|
154 |
fun mk_all (x,P) = HOLogic.mk_all (x,dummyT,P); |
|
16778 | 155 |
fun mk_ex (x,P) = mk_exists (x,dummyT,P); |
1274 | 156 |
local |
11531
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
157 |
fun sg [s] = %:s |
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
158 |
| sg (s::ss) = %%:"_classes" $ %:s $ sg ss |
2445
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents:
2276
diff
changeset
|
159 |
| sg _ = Imposs "library:sg"; |
11531
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
160 |
fun sf [] = %%:"_emptysort" |
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
161 |
| sf s = %%:"_sort" $ sg s |
15570 | 162 |
fun tf (Type (s,args)) = Library.foldl (op $) (%:s,map tf args) |
11531
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
163 |
| tf (TFree(s,sort)) = %%:"_ofsort" $ %:s $ sf sort |
2445
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents:
2276
diff
changeset
|
164 |
| tf _ = Imposs "library:tf"; |
1274 | 165 |
in |
11531
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
166 |
fun mk_constrain (typ,T) = %%:"_constrain" $ T $ tf typ; |
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
167 |
fun mk_constrainall (x,typ,P) = %%:"All" $ (%%:"_constrainAbs" $ mk_lam(x,P) $ tf typ); |
1274 | 168 |
end; |
169 |
end; |
|
170 |
||
11531
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
171 |
fun mk_All (x,P) = %%:"all" $ mk_lam(x,P); (* meta universal quantification *) |
1274 | 172 |
|
16778 | 173 |
infixr 0 ===>; fun S ===> T = %%:"==>" $ S $ T; |
174 |
infixr 0 ==>; fun S ==> T = mk_trp S ===> mk_trp T; |
|
175 |
infix 0 ==; fun S == T = %%:"==" $ S $ T; |
|
176 |
infix 1 ===; fun S === T = %%:"op =" $ S $ T; |
|
16842 | 177 |
infix 1 ~=; fun S ~= T = HOLogic.mk_not (S === T); |
16778 | 178 |
infix 1 <<; fun S << T = %%:lessN $ S $ T; |
16842 | 179 |
infix 1 ~<<; fun S ~<< T = HOLogic.mk_not (S << T); |
1274 | 180 |
|
16778 | 181 |
infix 9 ` ; fun f` x = %%:Rep_CFunN $ f $ x; |
11531
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
182 |
infix 9 `% ; fun f`% s = f` %: s; |
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
183 |
infix 9 `%%; fun f`%%s = f` %%:s; |
18083
cf7669049df5
cleaned up; renamed library function mk_cRep_CFun to list_ccomb; replaced rep_TFree with dest_TFree; added functions spair, mk_stuple
huffman
parents:
17840
diff
changeset
|
184 |
val list_ccomb = Library.foldl (op `); (* continuous version of list_comb *) |
cf7669049df5
cleaned up; renamed library function mk_cRep_CFun to list_ccomb; replaced rep_TFree with dest_TFree; added functions spair, mk_stuple
huffman
parents:
17840
diff
changeset
|
185 |
fun con_app2 con f args = list_ccomb(%%:con,map f args); |
1274 | 186 |
fun con_app con = con_app2 con %#; |
187 |
fun if_rec arg f y = if is_rec arg then f (rec_of arg) else y; |
|
16842 | 188 |
fun app_rec_arg p arg = if_rec arg (fn n => fn x => (p n)`x) I (%# arg); |
4754 | 189 |
fun prj _ _ x ( _::[]) _ = x |
190 |
| prj f1 _ x (_::y::ys) 0 = f1 x y |
|
191 |
| prj f1 f2 x (y:: ys) j = prj f1 f2 (f2 x y) ys (j-1); |
|
11531
d038246a62f2
Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents:
7652
diff
changeset
|
192 |
fun proj x = prj (fn S => K(%%:"fst" $S)) (fn S => K(%%:"snd" $S)) x; |
16778 | 193 |
fun cproj x = prj (fn S => K(%%:cfstN`S)) (fn S => K(%%:csndN`S)) x; |
15570 | 194 |
fun lift tfn = Library.foldr (fn (x,t)=> (mk_trp(tfn x) ===> t)); |
1274 | 195 |
|
16778 | 196 |
fun /\ v T = %%:Abs_CFunN $ mk_lam(v,T); |
1274 | 197 |
fun /\# (arg,T) = /\ (vname arg) T; |
16778 | 198 |
infixr 9 oo; fun S oo T = %%:cfcompN`S`T; |
199 |
val UU = %%:UU_N; |
|
1274 | 200 |
fun strict f = f`UU === UU; |
201 |
fun defined t = t ~= UU; |
|
18083
cf7669049df5
cleaned up; renamed library function mk_cRep_CFun to list_ccomb; replaced rep_TFree with dest_TFree; added functions spair, mk_stuple
huffman
parents:
17840
diff
changeset
|
202 |
fun cpair (t,u) = %%:cpairN`t`u; |
cf7669049df5
cleaned up; renamed library function mk_cRep_CFun to list_ccomb; replaced rep_TFree with dest_TFree; added functions spair, mk_stuple
huffman
parents:
17840
diff
changeset
|
203 |
fun spair (t,u) = %%:spairN`t`u; |
16778 | 204 |
fun mk_ctuple [] = HOLogic.unit (* used in match_defs *) |
18083
cf7669049df5
cleaned up; renamed library function mk_cRep_CFun to list_ccomb; replaced rep_TFree with dest_TFree; added functions spair, mk_stuple
huffman
parents:
17840
diff
changeset
|
205 |
| mk_ctuple ts = foldr1 cpair ts; |
cf7669049df5
cleaned up; renamed library function mk_cRep_CFun to list_ccomb; replaced rep_TFree with dest_TFree; added functions spair, mk_stuple
huffman
parents:
17840
diff
changeset
|
206 |
fun mk_stuple [] = %%:ONE_N |
cf7669049df5
cleaned up; renamed library function mk_cRep_CFun to list_ccomb; replaced rep_TFree with dest_TFree; added functions spair, mk_stuple
huffman
parents:
17840
diff
changeset
|
207 |
| mk_stuple ts = foldr1 spair ts; |
16224
57094b83774e
Domain package generates match functions for new datatypes, for use with the fixrec package
huffman
parents:
15638
diff
changeset
|
208 |
fun mk_ctupleT [] = HOLogic.unitT (* used in match_defs *) |
18083
cf7669049df5
cleaned up; renamed library function mk_cRep_CFun to list_ccomb; replaced rep_TFree with dest_TFree; added functions spair, mk_stuple
huffman
parents:
17840
diff
changeset
|
209 |
| mk_ctupleT Ts = foldr1 HOLogic.mk_prodT Ts; |
19092
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
210 |
fun mk_maybeT T = Type ("Fixrec.maybe",[T]); |
18113 | 211 |
fun cpair_pat (p1,p2) = %%:cpair_patN $ p1 $ p2; |
1274 | 212 |
fun lift_defined f = lift (fn x => defined (f x)); |
4188
1025a27b08f9
changed libraray function find to find_index_eq, currying it
oheimb
parents:
4122
diff
changeset
|
213 |
fun bound_arg vns v = Bound(length vns -find_index_eq v vns -1); |
1274 | 214 |
|
16778 | 215 |
fun cont_eta_contract (Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body)) = |
1274 | 216 |
(case cont_eta_contract body of |
16778 | 217 |
body' as (Const("Cfun.Rep_CFun",Ta) $ f $ Bound 0) => |
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
218 |
if not (0 mem loose_bnos f) then incr_boundvars ~1 f |
16778 | 219 |
else Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body') |
220 |
| body' => Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body')) |
|
1274 | 221 |
| cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t |
222 |
| cont_eta_contract t = t; |
|
223 |
||
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
224 |
fun idx_name dnames s n = s^(if length dnames = 1 then "" else string_of_int n); |
1834 | 225 |
fun when_funs cons = if length cons = 1 then ["f"] |
226 |
else mapn (fn n => K("f"^(string_of_int n))) 1 cons; |
|
1274 | 227 |
fun when_body cons funarg = let |
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
228 |
fun one_fun n (_,[] ) = /\ "dummy" (funarg(1,n)) |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
229 |
| one_fun n (_,args) = let |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
230 |
val l2 = length args; |
16778 | 231 |
fun idxs m arg = (if is_lazy arg then fn x=> %%:fupN` %%:ID_N`x |
16842 | 232 |
else I) (Bound(l2-m)); |
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
233 |
in cont_eta_contract (foldr'' |
16778 | 234 |
(fn (a,t) => %%:ssplitN`(/\# (a,t))) |
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
235 |
(args, |
18083
cf7669049df5
cleaned up; renamed library function mk_cRep_CFun to list_ccomb; replaced rep_TFree with dest_TFree; added functions spair, mk_stuple
huffman
parents:
17840
diff
changeset
|
236 |
fn a=> /\#(a,(list_ccomb(funarg(l2,n),mapn idxs 1 args)))) |
1637
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
237 |
) end; |
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents:
1461
diff
changeset
|
238 |
in (if length cons = 1 andalso length(snd(hd cons)) <= 1 |
16842 | 239 |
then fn t => %%:strictifyN`t else I) |
17811 | 240 |
(foldr1 (fn (x,y)=> %%:sscaseN`x`y) (mapn one_fun 1 cons)) end; |
1274 | 241 |
end; (* struct *) |