| author | haftmann | 
| Mon, 30 Nov 2009 08:08:31 +0100 | |
| changeset 33966 | b863967f23ea | 
| parent 33396 | 45c5c3c51918 | 
| child 33971 | 9c7fa7f76950 | 
| permissions | -rw-r--r-- | 
| 32126 | 1 | (* Title: HOLCF/Tools/Domain/domain_library.ML | 
| 23152 | 2 | Author: David von Oheimb | 
| 3 | ||
| 4 | Library for domain command. | |
| 5 | *) | |
| 6 | ||
| 7 | ||
| 8 | (* ----- general support ---------------------------------------------------- *) | |
| 9 | ||
| 10 | fun mapn f n [] = [] | |
| 31288 | 11 | | mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs; | 
| 23152 | 12 | |
| 31288 | 13 | fun foldr'' f (l,f2) = | 
| 14 | let fun itr [] = raise Fail "foldr''" | |
| 15 | | itr [a] = f2 a | |
| 16 | | itr (a::l) = f(a, itr l) | |
| 17 | in itr l end; | |
| 23152 | 18 | |
| 31288 | 19 | fun map_cumulr f start xs = | 
| 20 | List.foldr (fn (x,(ys,res))=>case f(x,res) of (y,res2) => | |
| 21 | (y::ys,res2)) ([],start) xs; | |
| 23152 | 22 | |
| 23 | fun first (x,_,_) = x; fun second (_,x,_) = x; fun third (_,_,x) = x; | |
| 24 | fun upd_first f (x,y,z) = (f x, y, z); | |
| 25 | fun upd_second f (x,y,z) = ( x, f y, z); | |
| 26 | fun upd_third f (x,y,z) = ( x, y, f z); | |
| 27 | ||
| 31288 | 28 | fun atomize ctxt thm = | 
| 29 | let | |
| 30 | val r_inst = read_instantiate ctxt; | |
| 31 | fun at thm = | |
| 32 | case concl_of thm of | |
| 33 |             _$(Const("op &",_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
 | |
| 34 |           | _$(Const("All" ,_)$Abs(s,_,_))=> at(thm RS (r_inst [(("x", 0), "?" ^ s)] spec))
 | |
| 35 | | _ => [thm]; | |
| 36 | in map zero_var_indexes (at thm) end; | |
| 23152 | 37 | |
| 31006 | 38 | (* infix syntax *) | 
| 39 | ||
| 40 | infixr 5 -->; | |
| 41 | infixr 6 ->>; | |
| 42 | infixr 0 ===>; | |
| 43 | infixr 0 ==>; | |
| 44 | infix 0 ==; | |
| 45 | infix 1 ===; | |
| 46 | infix 1 ~=; | |
| 47 | infix 1 <<; | |
| 48 | infix 1 ~<<; | |
| 49 | ||
| 50 | infix 9 ` ; | |
| 51 | infix 9 `% ; | |
| 52 | infix 9 `%%; | |
| 53 | ||
| 54 | ||
| 23152 | 55 | (* ----- specific support for domain ---------------------------------------- *) | 
| 56 | ||
| 31006 | 57 | signature DOMAIN_LIBRARY = | 
| 58 | sig | |
| 59 | val Imposs : string -> 'a; | |
| 31162 | 60 | val cpo_type : theory -> typ -> bool; | 
| 31006 | 61 | val pcpo_type : theory -> typ -> bool; | 
| 62 | val string_of_typ : theory -> typ -> string; | |
| 63 | ||
| 64 | (* Creating HOLCF types *) | |
| 65 | val mk_cfunT : typ * typ -> typ; | |
| 66 | val ->> : typ * typ -> typ; | |
| 67 | val mk_ssumT : typ * typ -> typ; | |
| 68 | val mk_sprodT : typ * typ -> typ; | |
| 69 | val mk_uT : typ -> typ; | |
| 70 | val oneT : typ; | |
| 71 | val trT : typ; | |
| 72 | val mk_maybeT : typ -> typ; | |
| 73 | val mk_ctupleT : typ list -> typ; | |
| 74 | val mk_TFree : string -> typ; | |
| 75 | val pcpoS : sort; | |
| 76 | ||
| 77 | (* Creating HOLCF terms *) | |
| 78 | val %: : string -> term; | |
| 79 | val %%: : string -> term; | |
| 80 | val ` : term * term -> term; | |
| 81 | val `% : term * string -> term; | |
| 82 | val /\ : string -> term -> term; | |
| 83 | val UU : term; | |
| 84 | val TT : term; | |
| 85 | val FF : term; | |
| 31231 | 86 | val ID : term; | 
| 87 | val oo : term * term -> term; | |
| 31006 | 88 | val mk_up : term -> term; | 
| 89 | val mk_sinl : term -> term; | |
| 90 | val mk_sinr : term -> term; | |
| 91 | val mk_stuple : term list -> term; | |
| 92 | val mk_ctuple : term list -> term; | |
| 93 | val mk_fix : term -> term; | |
| 94 | val mk_iterate : term * term * term -> term; | |
| 95 | val mk_fail : term; | |
| 96 | val mk_return : term -> term; | |
| 97 | val list_ccomb : term * term list -> term; | |
| 31288 | 98 | (* | 
| 99 |    val con_app : string -> ('a * 'b * string) list -> term;
 | |
| 100 | *) | |
| 31006 | 101 |   val con_app2 : string -> ('a -> term) -> 'a list -> term;
 | 
| 102 | val proj : term -> 'a list -> int -> term; | |
| 103 |   val prj : ('a -> 'b -> 'a) -> ('a -> 'b -> 'a) -> 'a -> 'b list -> int -> 'a;
 | |
| 104 | val mk_ctuple_pat : term list -> term; | |
| 105 | val mk_branch : term -> term; | |
| 106 | ||
| 107 | (* Creating propositions *) | |
| 108 | val mk_conj : term * term -> term; | |
| 109 | val mk_disj : term * term -> term; | |
| 110 | val mk_imp : term * term -> term; | |
| 111 | val mk_lam : string * term -> term; | |
| 112 | val mk_all : string * term -> term; | |
| 113 | val mk_ex : string * term -> term; | |
| 114 | val mk_constrain : typ * term -> term; | |
| 115 | val mk_constrainall : string * typ * term -> term; | |
| 116 | val === : term * term -> term; | |
| 117 | val << : term * term -> term; | |
| 118 | val ~<< : term * term -> term; | |
| 119 | val strict : term -> term; | |
| 120 | val defined : term -> term; | |
| 121 | val mk_adm : term -> term; | |
| 122 | val mk_compact : term -> term; | |
| 123 |   val lift : ('a -> term) -> 'a list * term -> term;
 | |
| 124 |   val lift_defined : ('a -> term) -> 'a list * term -> term;
 | |
| 125 | ||
| 126 | (* Creating meta-propositions *) | |
| 127 | val mk_trp : term -> term; (* HOLogic.mk_Trueprop *) | |
| 128 | val == : term * term -> term; | |
| 129 | val ===> : term * term -> term; | |
| 130 | val ==> : term * term -> term; | |
| 131 | val mk_All : string * term -> term; | |
| 132 | ||
| 31288 | 133 | (* Domain specifications *) | 
| 134 | eqtype arg; | |
| 31006 | 135 | type cons = string * arg list; | 
| 136 | type eq = (string * typ list) * cons list; | |
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31288diff
changeset | 137 | val mk_arg : (bool * Datatype.dtyp) * string option * string -> arg; | 
| 31006 | 138 | val is_lazy : arg -> bool; | 
| 139 | val rec_of : arg -> int; | |
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31288diff
changeset | 140 | val dtyp_of : arg -> Datatype.dtyp; | 
| 31006 | 141 | val sel_of : arg -> string option; | 
| 142 | val vname : arg -> string; | |
| 143 | val upd_vname : (string -> string) -> arg -> arg; | |
| 144 | val is_rec : arg -> bool; | |
| 145 | val is_nonlazy_rec : arg -> bool; | |
| 146 | val nonlazy : arg list -> string list; | |
| 147 | val nonlazy_rec : arg list -> string list; | |
| 148 | val %# : arg -> term; | |
| 149 | val /\# : arg * term -> term; | |
| 150 | val when_body : cons list -> (int * int -> term) -> term; | |
| 151 | val when_funs : 'a list -> string list; | |
| 152 | val bound_arg : ''a list -> ''a -> term; (* ''a = arg or string *) | |
| 153 | val idx_name : 'a list -> string -> int -> string; | |
| 154 | val app_rec_arg : (int -> term) -> arg -> term; | |
| 31228 | 155 | val con_app : string -> arg list -> term; | 
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31288diff
changeset | 156 | val dtyp_of_eq : eq -> Datatype.dtyp; | 
| 31228 | 157 | |
| 31006 | 158 | |
| 159 | (* Name mangling *) | |
| 160 | val strip_esc : string -> string; | |
| 161 | val extern_name : string -> string; | |
| 162 | val dis_name : string -> string; | |
| 163 | val mat_name : string -> string; | |
| 164 | val pat_name : string -> string; | |
| 165 | val mk_var_names : string list -> string list; | |
| 166 | end; | |
| 167 | ||
| 31023 | 168 | structure Domain_Library :> DOMAIN_LIBRARY = | 
| 31006 | 169 | struct | 
| 23152 | 170 | |
| 171 | exception Impossible of string; | |
| 172 | fun Imposs msg = raise Impossible ("Domain:"^msg);
 | |
| 173 | ||
| 174 | (* ----- name handling ----- *) | |
| 175 | ||
| 31288 | 176 | val strip_esc = | 
| 177 |     let fun strip ("'" :: c :: cs) = c :: strip cs
 | |
| 178 | | strip ["'"] = [] | |
| 179 | | strip (c :: cs) = c :: strip cs | |
| 180 | | strip [] = []; | |
| 181 | in implode o strip o Symbol.explode end; | |
| 23152 | 182 | |
| 31288 | 183 | fun extern_name con = | 
| 184 | case Symbol.explode con of | |
| 185 |       ("o"::"p"::" "::rest) => implode rest
 | |
| 186 | | _ => con; | |
| 23152 | 187 | fun dis_name con = "is_"^ (extern_name con); | 
| 188 | fun dis_name_ con = "is_"^ (strip_esc con); | |
| 189 | fun mat_name con = "match_"^ (extern_name con); | |
| 190 | fun mat_name_ con = "match_"^ (strip_esc con); | |
| 191 | fun pat_name con = (extern_name con) ^ "_pat"; | |
| 192 | fun pat_name_ con = (strip_esc con) ^ "_pat"; | |
| 193 | ||
| 194 | (* make distinct names out of the type list, | |
| 31288 | 195 | forbidding "o","n..","x..","f..","P.." as names *) | 
| 23152 | 196 | (* a number string is added if necessary *) | 
| 31288 | 197 | fun mk_var_names ids : string list = | 
| 198 | let | |
| 199 | fun nonreserved s = if s mem ["n","x","f","P"] then s^"'" else s; | |
| 200 | fun index_vnames(vn::vns,occupied) = | |
| 23152 | 201 | (case AList.lookup (op =) occupied vn of | 
| 202 | NONE => if vn mem vns | |
| 203 | then (vn^"1") :: index_vnames(vns,(vn,1) ::occupied) | |
| 204 | else vn :: index_vnames(vns, occupied) | |
| 205 | | SOME(i) => (vn^(string_of_int (i+1))) | |
| 31288 | 206 | :: index_vnames(vns,(vn,i+1)::occupied)) | 
| 207 | | index_vnames([],occupied) = []; | |
| 208 |     in index_vnames(map nonreserved ids, [("O",0),("o",0)]) end;
 | |
| 23152 | 209 | |
| 31162 | 210 | fun cpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort cpo});
 | 
| 30910 | 211 | fun pcpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort pcpo});
 | 
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26012diff
changeset | 212 | fun string_of_typ sg = Syntax.string_of_typ_global sg o Sign.certify_typ sg; | 
| 23152 | 213 | |
| 214 | (* ----- constructor list handling ----- *) | |
| 215 | ||
| 31006 | 216 | type arg = | 
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31288diff
changeset | 217 | (bool * Datatype.dtyp) * (* (lazy, recursive element) *) | 
| 31288 | 218 | string option * (* selector name *) | 
| 219 | string; (* argument name *) | |
| 31006 | 220 | |
| 221 | type cons = | |
| 31288 | 222 | string * (* operator name of constr *) | 
| 223 | arg list; (* argument list *) | |
| 31006 | 224 | |
| 225 | type eq = | |
| 31288 | 226 | (string * (* name of abstracted type *) | 
| 227 | typ list) * (* arguments of abstracted type *) | |
| 228 | cons list; (* represented type, as a constructor list *) | |
| 23152 | 229 | |
| 31228 | 230 | val mk_arg = I; | 
| 31229 | 231 | |
| 232 | fun rec_of ((_,dtyp),_,_) = | |
| 31288 | 233 | case dtyp of DatatypeAux.DtRec i => i | _ => ~1; | 
| 31229 | 234 | (* FIXME: what about indirect recursion? *) | 
| 235 | ||
| 236 | fun is_lazy arg = fst (first arg); | |
| 31231 | 237 | fun dtyp_of arg = snd (first arg); | 
| 23152 | 238 | val sel_of = second; | 
| 239 | val vname = third; | |
| 240 | val upd_vname = upd_third; | |
| 241 | fun is_rec arg = rec_of arg >=0; | |
| 242 | fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg); | |
| 33317 | 243 | fun nonlazy args = map vname (filter_out is_lazy args); | 
| 244 | fun nonlazy_rec args = map vname (filter is_nonlazy_rec args); | |
| 23152 | 245 | |
| 31231 | 246 | |
| 247 | (* ----- combinators for making dtyps ----- *) | |
| 248 | ||
| 249 | fun mk_uD T = DatatypeAux.DtType(@{type_name "u"}, [T]);
 | |
| 250 | fun mk_sprodD (T, U) = DatatypeAux.DtType(@{type_name "**"}, [T, U]);
 | |
| 251 | fun mk_ssumD (T, U) = DatatypeAux.DtType(@{type_name "++"}, [T, U]);
 | |
| 252 | fun mk_liftD T = DatatypeAux.DtType(@{type_name "lift"}, [T]);
 | |
| 253 | val unitD = DatatypeAux.DtType(@{type_name "unit"}, []);
 | |
| 254 | val boolD = DatatypeAux.DtType(@{type_name "bool"}, []);
 | |
| 255 | val oneD = mk_liftD unitD; | |
| 256 | val trD = mk_liftD boolD; | |
| 257 | fun big_sprodD ds = case ds of [] => oneD | _ => foldr1 mk_sprodD ds; | |
| 258 | fun big_ssumD ds = case ds of [] => unitD | _ => foldr1 mk_ssumD ds; | |
| 259 | ||
| 260 | fun dtyp_of_arg ((lazy, D), _, _) = if lazy then mk_uD D else D; | |
| 261 | fun dtyp_of_cons (_, args) = big_sprodD (map dtyp_of_arg args); | |
| 262 | fun dtyp_of_eq (_, cons) = big_ssumD (map dtyp_of_cons cons); | |
| 263 | ||
| 264 | ||
| 23152 | 265 | (* ----- support for type and mixfix expressions ----- *) | 
| 266 | ||
| 30910 | 267 | fun mk_uT T = Type(@{type_name "u"}, [T]);
 | 
| 268 | fun mk_cfunT (T, U) = Type(@{type_name "->"}, [T, U]);
 | |
| 269 | fun mk_sprodT (T, U) = Type(@{type_name "**"}, [T, U]);
 | |
| 270 | fun mk_ssumT (T, U) = Type(@{type_name "++"}, [T, U]);
 | |
| 271 | val oneT = @{typ one};
 | |
| 272 | val trT = @{typ tr};
 | |
| 273 | ||
| 274 | val op ->> = mk_cfunT; | |
| 275 | ||
| 276 | fun mk_TFree s = TFree ("'" ^ s, @{sort pcpo});
 | |
| 23152 | 277 | |
| 278 | (* ----- support for term expressions ----- *) | |
| 279 | ||
| 280 | fun %: s = Free(s,dummyT); | |
| 281 | fun %# arg = %:(vname arg); | |
| 282 | fun %%: s = Const(s,dummyT); | |
| 283 | ||
| 284 | local open HOLogic in | |
| 285 | val mk_trp = mk_Trueprop; | |
| 286 | fun mk_conj (S,T) = conj $ S $ T; | |
| 287 | fun mk_disj (S,T) = disj $ S $ T; | |
| 288 | fun mk_imp (S,T) = imp $ S $ T; | |
| 289 | fun mk_lam (x,T) = Abs(x,dummyT,T); | |
| 290 | fun mk_all (x,P) = HOLogic.mk_all (x,dummyT,P); | |
| 291 | fun mk_ex (x,P) = mk_exists (x,dummyT,P); | |
| 24680 | 292 | val mk_constrain = uncurry TypeInfer.constrain; | 
| 293 | fun mk_constrainall (x,typ,P) = %%:"All" $ (TypeInfer.constrain (typ --> boolT) (mk_lam(x,P))); | |
| 23152 | 294 | end | 
| 295 | ||
| 296 | fun mk_All (x,P) = %%:"all" $ mk_lam(x,P); (* meta universal quantification *) | |
| 297 | ||
| 298 | infixr 0 ===>; fun S ===> T = %%:"==>" $ S $ T; | |
| 299 | infixr 0 ==>; fun S ==> T = mk_trp S ===> mk_trp T; | |
| 300 | infix 0 ==; fun S == T = %%:"==" $ S $ T; | |
| 301 | infix 1 ===; fun S === T = %%:"op =" $ S $ T; | |
| 302 | infix 1 ~=; fun S ~= T = HOLogic.mk_not (S === T); | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31023diff
changeset | 303 | infix 1 <<;     fun S <<  T = %%: @{const_name Porder.below} $ S $ T;
 | 
| 23152 | 304 | infix 1 ~<<; fun S ~<< T = HOLogic.mk_not (S << T); | 
| 305 | ||
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 306 | infix 9 `  ; fun f ` x = %%: @{const_name Rep_CFun} $ f $ x;
 | 
| 23152 | 307 | infix 9 `% ; fun f`% s = f` %: s; | 
| 308 | infix 9 `%%; fun f`%%s = f` %%:s; | |
| 26012 | 309 | |
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 310 | fun mk_adm t = %%: @{const_name adm} $ t;
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 311 | fun mk_compact t = %%: @{const_name compact} $ t;
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 312 | val ID = %%: @{const_name ID};
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 313 | fun mk_strictify t = %%: @{const_name strictify}`t;
 | 
| 26012 | 314 | (*val csplitN = "Cprod.csplit";*) | 
| 315 | (*val sfstN = "Sprod.sfst";*) | |
| 316 | (*val ssndN = "Sprod.ssnd";*) | |
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 317 | fun mk_ssplit t = %%: @{const_name ssplit}`t;
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 318 | fun mk_sinl t = %%: @{const_name sinl}`t;
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 319 | fun mk_sinr t = %%: @{const_name sinr}`t;
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 320 | fun mk_sscase (x, y) = %%: @{const_name sscase}`x`y;
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 321 | fun mk_up t = %%: @{const_name up}`t;
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 322 | fun mk_fup (t,u) = %%: @{const_name fup} ` t ` u;
 | 
| 26012 | 323 | val ONE = @{term ONE};
 | 
| 324 | val TT = @{term TT};
 | |
| 325 | val FF = @{term FF};
 | |
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 326 | fun mk_iterate (n,f,z) = %%: @{const_name iterate} $ n ` f ` z;
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 327 | fun mk_fix t = %%: @{const_name fix}`t;
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 328 | fun mk_return t = %%: @{const_name Fixrec.return}`t;
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 329 | val mk_fail = %%: @{const_name Fixrec.fail};
 | 
| 26012 | 330 | |
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 331 | fun mk_branch t = %%: @{const_name Fixrec.branch} $ t;
 | 
| 26012 | 332 | |
| 333 | val pcpoS = @{sort pcpo};
 | |
| 334 | ||
| 23152 | 335 | val list_ccomb = Library.foldl (op `); (* continuous version of list_comb *) | 
| 336 | fun con_app2 con f args = list_ccomb(%%:con,map f args); | |
| 337 | fun con_app con = con_app2 con %#; | |
| 338 | fun if_rec arg f y = if is_rec arg then f (rec_of arg) else y; | |
| 339 | fun app_rec_arg p arg = if_rec arg (fn n => fn x => (p n)`x) I (%# arg); | |
| 340 | fun prj _ _ x ( _::[]) _ = x | |
| 31288 | 341 | | prj f1 _ x (_::y::ys) 0 = f1 x y | 
| 342 | | prj f1 f2 x (y:: ys) j = prj f1 f2 (f2 x y) ys (j-1); | |
| 23152 | 343 | fun proj x = prj (fn S => K(%%:"fst" $S)) (fn S => K(%%:"snd" $S)) x; | 
| 344 | fun lift tfn = Library.foldr (fn (x,t)=> (mk_trp(tfn x) ===> t)); | |
| 345 | ||
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 346 | fun /\ v T = %%: @{const_name Abs_CFun} $ mk_lam(v,T);
 | 
| 23152 | 347 | fun /\# (arg,T) = /\ (vname arg) T; | 
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 348 | infixr 9 oo; fun S oo T = %%: @{const_name cfcomp}`S`T;
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 349 | val UU = %%: @{const_name UU};
 | 
| 23152 | 350 | fun strict f = f`UU === UU; | 
| 351 | fun defined t = t ~= UU; | |
| 33396 | 352 | fun cpair (t,u) = %%: @{const_name Pair} $ t $ u;
 | 
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 353 | fun spair (t,u) = %%: @{const_name spair}`t`u;
 | 
| 23152 | 354 | fun mk_ctuple [] = HOLogic.unit (* used in match_defs *) | 
| 31288 | 355 | | mk_ctuple ts = foldr1 cpair ts; | 
| 26012 | 356 | fun mk_stuple [] = ONE | 
| 31288 | 357 | | mk_stuple ts = foldr1 spair ts; | 
| 23152 | 358 | fun mk_ctupleT [] = HOLogic.unitT (* used in match_defs *) | 
| 31288 | 359 | | mk_ctupleT Ts = foldr1 HOLogic.mk_prodT Ts; | 
| 23152 | 360 | fun mk_maybeT T = Type ("Fixrec.maybe",[T]);
 | 
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30190diff
changeset | 361 | fun cpair_pat (p1,p2) = %%: @{const_name cpair_pat} $ p1 $ p2;
 | 
| 26012 | 362 | val mk_ctuple_pat = foldr1 cpair_pat; | 
| 23152 | 363 | fun lift_defined f = lift (fn x => defined (f x)); | 
| 31986 | 364 | fun bound_arg vns v = Bound (length vns - find_index (fn v' => v' = v) vns - 1); | 
| 23152 | 365 | |
| 366 | fun cont_eta_contract (Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body)) = 
 | |
| 31288 | 367 | (case cont_eta_contract body of | 
| 368 |        body' as (Const("Cfun.Rep_CFun",Ta) $ f $ Bound 0) => 
 | |
| 369 | if not (0 mem loose_bnos f) then incr_boundvars ~1 f | |
| 370 |        else   Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body')
 | |
| 371 |      | body' => Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body'))
 | |
| 372 | | cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t | |
| 373 | | cont_eta_contract t = t; | |
| 23152 | 374 | |
| 375 | fun idx_name dnames s n = s^(if length dnames = 1 then "" else string_of_int n); | |
| 376 | fun when_funs cons = if length cons = 1 then ["f"] | |
| 377 |                      else mapn (fn n => K("f"^(string_of_int n))) 1 cons;
 | |
| 31288 | 378 | fun when_body cons funarg = | 
| 379 | let | |
| 380 | fun one_fun n (_,[] ) = /\ "dummy" (funarg(1,n)) | |
| 381 | | one_fun n (_,args) = let | |
| 382 | val l2 = length args; | |
| 383 | fun idxs m arg = (if is_lazy arg then (fn t => mk_fup (ID, t)) | |
| 384 | else I) (Bound(l2-m)); | |
| 385 | in cont_eta_contract | |
| 386 | (foldr'' | |
| 387 | (fn (a,t) => mk_ssplit (/\# (a,t))) | |
| 388 | (args, | |
| 389 | fn a=> /\#(a,(list_ccomb(funarg(l2,n),mapn idxs 1 args)))) | |
| 390 | ) end; | |
| 391 | in (if length cons = 1 andalso length(snd(hd cons)) <= 1 | |
| 392 | then mk_strictify else I) | |
| 393 | (foldr1 mk_sscase (mapn one_fun 1 cons)) end; | |
| 394 | ||
| 23152 | 395 | end; (* struct *) |