oheimb@2276: (* Title: HOLCF/domain/library.ML oheimb@2445: ID: $Id$ wenzelm@12030: Author: David von Oheimb wenzelm@12030: License: GPL (GNU GENERAL PUBLIC LICENSE) oheimb@2445: wenzelm@12030: Library for domain section. regensbu@1274: *) regensbu@1274: oheimb@2446: oheimb@1637: (* ----- general support ---------------------------------------------------- *) regensbu@1274: regensbu@1274: fun Id x = x; regensbu@1274: regensbu@1274: fun mapn f n [] = [] regensbu@1274: | mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs; regensbu@1274: oheimb@1637: fun foldr'' f (l,f2) = let fun itr [] = raise LIST "foldr''" oheimb@1637: | itr [a] = f2 a oheimb@1637: | itr (a::l) = f(a, itr l) oheimb@1637: in itr l end; regensbu@1274: fun foldr' f l = foldr'' f (l,Id); oheimb@1637: fun map_cumulr f start xs = foldr (fn (x,(ys,res))=>case f(x,res) of (y,res2) => oheimb@1637: (y::ys,res2)) (xs,([],start)); regensbu@1274: regensbu@1274: regensbu@1274: fun first (x,_,_) = x; fun second (_,x,_) = x; fun third (_,_,x) = x; regensbu@1274: fun upd_first f (x,y,z) = (f x, y, z); regensbu@1274: fun upd_second f (x,y,z) = ( x, f y, z); regensbu@1274: fun upd_third f (x,y,z) = ( x, y, f z); regensbu@1274: oheimb@1637: fun atomize thm = let val r_inst = read_instantiate; oheimb@1637: fun at thm = case concl_of thm of oheimb@1637: _$(Const("op &",_)$_$_) => at(thm RS conjunct1)@at(thm RS conjunct2) oheimb@1637: | _$(Const("All" ,_)$Abs(s,_,_))=> at(thm RS (r_inst [("x","?"^s)] spec)) oheimb@1637: | _ => [thm]; oheimb@1637: in map zero_var_indexes (at thm) end; regensbu@1274: oheimb@1637: (* ----- specific support for domain ---------------------------------------- *) regensbu@1274: regensbu@1274: structure Domain_Library = struct regensbu@1274: oheimb@4122: open HOLCFLogic; oheimb@4122: regensbu@1274: exception Impossible of string; regensbu@1274: fun Imposs msg = raise Impossible ("Domain:"^msg); regensbu@1274: regensbu@1274: (* ----- name handling ----- *) regensbu@1274: oheimb@1637: val strip_esc = let fun strip ("'" :: c :: cs) = c :: strip cs oheimb@1637: | strip ["'"] = [] oheimb@1637: | strip (c :: cs) = c :: strip cs oheimb@1637: | strip [] = []; wenzelm@4709: in implode o strip o Symbol.explode end; regensbu@1274: wenzelm@4709: fun extern_name con = case Symbol.explode con of oheimb@1637: ("o"::"p"::" "::rest) => implode rest oheimb@1637: | _ => con; regensbu@1274: fun dis_name con = "is_"^ (extern_name con); regensbu@1274: fun dis_name_ con = "is_"^ (strip_esc con); regensbu@1274: oheimb@1637: (* make distinct names out of the type list, oheimb@5292: forbidding "o","n..","x..","f..","P.." as names *) oheimb@1637: (* a number string is added if necessary *) oheimb@4122: fun mk_var_names ids : string list = let oheimb@5292: fun nonreserved s = if s mem ["n","x","f","P"] then s^"'" else s; oheimb@1819: fun index_vnames(vn::vns,occupied) = oheimb@1819: (case assoc(occupied,vn) of regensbu@1274: None => if vn mem vns oheimb@1819: then (vn^"1") :: index_vnames(vns,(vn,1) ::occupied) oheimb@1819: else vn :: index_vnames(vns, occupied) oheimb@1819: | Some(i) => (vn^(string_of_int (i+1))) oheimb@1819: :: index_vnames(vns,(vn,i+1)::occupied)) oheimb@1819: | index_vnames([],occupied) = []; oheimb@4122: in index_vnames(map nonreserved ids, [("O",0),("o",0)]) end; regensbu@1274: oheimb@2445: fun rep_Type (Type x) = x | rep_Type _ = Imposs "library:rep_Type"; oheimb@2445: fun rep_TFree (TFree x) = x | rep_TFree _ = Imposs "library:rep_TFree"; oheimb@4008: val tsig_of = #tsig o Sign.rep_sg; oheimb@4122: wenzelm@7652: fun pcpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, pcpoS); oheimb@4008: fun string_of_typ sg = Sign.string_of_typ sg o Sign.certify_typ sg; oheimb@4008: fun str2typ sg = typ_of o read_ctyp sg; regensbu@1274: regensbu@1274: (* ----- constructor list handling ----- *) regensbu@1274: oheimb@1637: type cons = (string * (* operator name of constr *) oheimb@1637: ((bool*int)* (* (lazy,recursive element or ~1) *) oheimb@1637: string* (* selector name *) oheimb@1637: string) (* argument name *) oheimb@1637: list); (* argument list *) oheimb@1637: type eq = (string * (* name of abstracted type *) oheimb@1637: typ list) * (* arguments of abstracted type *) oheimb@1637: cons list; (* represented type, as a constructor list *) regensbu@1274: paulson@2238: fun rec_of arg = snd (first arg); paulson@2238: fun is_lazy arg = fst (first arg); regensbu@1274: val sel_of = second; regensbu@1274: val vname = third; regensbu@1274: val upd_vname = upd_third; regensbu@1274: fun is_rec arg = rec_of arg >=0; regensbu@1274: fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg); oheimb@1637: fun nonlazy args = map vname (filter_out is_lazy args); oheimb@1637: fun nonlazy_rec args = map vname (filter is_nonlazy_rec args); oheimb@1637: oheimb@1637: (* ----- support for type and mixfix expressions ----- *) oheimb@1637: oheimb@1637: infixr 5 -->; regensbu@1274: regensbu@1274: (* ----- support for term expressions ----- *) regensbu@1274: berghofe@11531: fun %: s = Free(s,dummyT); berghofe@11531: fun %# arg = %:(vname arg); berghofe@11531: fun %%: s = Const(s,dummyT); regensbu@1274: regensbu@1274: local open HOLogic in regensbu@1274: val mk_trp = mk_Trueprop; regensbu@1274: fun mk_conj (S,T) = conj $ S $ T; regensbu@1274: fun mk_disj (S,T) = disj $ S $ T; regensbu@1274: fun mk_imp (S,T) = imp $ S $ T; regensbu@1274: fun mk_lam (x,T) = Abs(x,dummyT,T); regensbu@1274: fun mk_all (x,P) = HOLogic.mk_all (x,dummyT,P); regensbu@1274: local berghofe@11531: fun sg [s] = %:s berghofe@11531: | sg (s::ss) = %%:"_classes" $ %:s $ sg ss oheimb@2445: | sg _ = Imposs "library:sg"; berghofe@11531: fun sf [] = %%:"_emptysort" berghofe@11531: | sf s = %%:"_sort" $ sg s berghofe@11531: fun tf (Type (s,args)) = foldl (op $) (%:s,map tf args) berghofe@11531: | tf (TFree(s,sort)) = %%:"_ofsort" $ %:s $ sf sort oheimb@2445: | tf _ = Imposs "library:tf"; regensbu@1274: in berghofe@11531: fun mk_constrain (typ,T) = %%:"_constrain" $ T $ tf typ; berghofe@11531: fun mk_constrainall (x,typ,P) = %%:"All" $ (%%:"_constrainAbs" $ mk_lam(x,P) $ tf typ); regensbu@1274: end; regensbu@1274: fun mk_ex (x,P) = mk_exists (x,dummyT,P); regensbu@1274: end; regensbu@1274: berghofe@11531: fun mk_All (x,P) = %%:"all" $ mk_lam(x,P); (* meta universal quantification *) regensbu@1274: berghofe@11531: infixr 0 ===>;fun S ===> T = %%:"==>" $ S $ T; regensbu@1274: infixr 0 ==>;fun S ==> T = mk_trp S ===> mk_trp T; berghofe@11531: infix 0 ==; fun S == T = %%:"==" $ S $ T; berghofe@11531: infix 1 ===; fun S === T = %%:"op =" $ S $ T; regensbu@1274: infix 1 ~=; fun S ~= T = mk_not (S === T); berghofe@11531: infix 1 <<; fun S << T = %%:"op <<" $ S $ T; regensbu@1274: infix 1 ~<<; fun S ~<< T = mk_not (S << T); regensbu@1274: berghofe@11531: infix 9 ` ; fun f` x = %%:"Rep_CFun" $ f $ x; berghofe@11531: infix 9 `% ; fun f`% s = f` %: s; berghofe@11531: infix 9 `%%; fun f`%%s = f` %%:s; slotosch@5291: fun mk_cRep_CFun (F,As) = foldl (op `) (F,As); berghofe@11531: fun con_app2 con f args = mk_cRep_CFun(%%:con,map f args); regensbu@1274: fun con_app con = con_app2 con %#; regensbu@1274: fun if_rec arg f y = if is_rec arg then f (rec_of arg) else y; regensbu@1274: fun app_rec_arg p arg = if_rec arg (fn n => fn x => (p n)`x) Id (%# arg); oheimb@4754: fun prj _ _ x ( _::[]) _ = x oheimb@4754: | prj f1 _ x (_::y::ys) 0 = f1 x y oheimb@4754: | prj f1 f2 x (y:: ys) j = prj f1 f2 (f2 x y) ys (j-1); oheimb@4754: fun fix_tp (tn, args) = (tn, map (K oneT) args); (* instantiate type to oheimb@4754: avoid type varaibles *) berghofe@11531: fun proj x = prj (fn S => K(%%:"fst" $S)) (fn S => K(%%:"snd" $S)) x; berghofe@11531: fun cproj x = prj (fn S => K(%%:"cfst"`S)) (fn S => K(%%:"csnd"`S)) x; oheimb@4754: fun cproj' T eqs = prj oheimb@4754: (fn S => fn t => Const("cfst",mk_prodT(dummyT,t)->>dummyT)`S) oheimb@4754: (fn S => fn t => Const("csnd",mk_prodT(t,dummyT)->>dummyT)`S) oheimb@4754: T (map ((fn tp => tp ->> tp) o Type o fix_tp o fst) eqs); regensbu@1274: fun lift tfn = foldr (fn (x,t)=> (mk_trp(tfn x) ===> t)); regensbu@1274: berghofe@11531: fun /\ v T = %%:"Abs_CFun" $ mk_lam(v,T); regensbu@1274: fun /\# (arg,T) = /\ (vname arg) T; berghofe@11531: infixr 9 oo; fun S oo T = %%:"cfcomp"`S`T; berghofe@11531: val UU = %%:"UU"; regensbu@1274: fun strict f = f`UU === UU; regensbu@1274: fun defined t = t ~= UU; berghofe@11531: fun cpair (S,T) = %%:"cpair"`S`T; regensbu@1274: fun lift_defined f = lift (fn x => defined (f x)); oheimb@4188: fun bound_arg vns v = Bound(length vns -find_index_eq v vns -1); regensbu@1274: slotosch@5291: fun cont_eta_contract (Const("Abs_CFun",TT) $ Abs(a,T,body)) = regensbu@1274: (case cont_eta_contract body of slotosch@5291: body' as (Const("Rep_CFun",Ta) $ f $ Bound 0) => oheimb@1637: if not (0 mem loose_bnos f) then incr_boundvars ~1 f slotosch@5291: else Const("Abs_CFun",TT) $ Abs(a,T,body') slotosch@5291: | body' => Const("Abs_CFun",TT) $ Abs(a,T,body')) regensbu@1274: | cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t regensbu@1274: | cont_eta_contract t = t; regensbu@1274: oheimb@1637: fun idx_name dnames s n = s^(if length dnames = 1 then "" else string_of_int n); oheimb@1834: fun when_funs cons = if length cons = 1 then ["f"] oheimb@1834: else mapn (fn n => K("f"^(string_of_int n))) 1 cons; regensbu@1274: fun when_body cons funarg = let oheimb@1637: fun one_fun n (_,[] ) = /\ "dummy" (funarg(1,n)) oheimb@1637: | one_fun n (_,args) = let oheimb@1637: val l2 = length args; berghofe@11531: fun idxs m arg = (if is_lazy arg then fn x=> %%:"fup"`%%"ID"`x oheimb@1637: else Id) (Bound(l2-m)); oheimb@1637: in cont_eta_contract (foldr'' berghofe@11531: (fn (a,t) => %%:"ssplit"`(/\# (a,t))) oheimb@1637: (args, slotosch@5291: fn a=> /\#(a,(mk_cRep_CFun(funarg(l2,n),mapn idxs 1 args)))) oheimb@1637: ) end; oheimb@1637: in (if length cons = 1 andalso length(snd(hd cons)) <= 1 berghofe@11531: then fn t => %%:"strictify"`t else Id) berghofe@11531: (foldr' (fn (x,y)=> %%:"sscase"`x`y) (mapn one_fun 1 cons)) end; regensbu@1274: end; (* struct *)