src/HOLCF/Tools/domain/domain_library.ML
changeset 32127 631546213601
parent 32112 6da9c2a49fed
parent 32126 a5042f260440
child 32128 59be4804c9ae
equal deleted inserted replaced
32112:6da9c2a49fed 32127:631546213601
     1 (*  Title:      HOLCF/Tools/domain/domain_library.ML
       
     2     Author:     David von Oheimb
       
     3 
       
     4 Library for domain command.
       
     5 *)
       
     6 
       
     7 
       
     8 (* ----- general support ---------------------------------------------------- *)
       
     9 
       
    10 fun mapn f n []      = []
       
    11   | mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs;
       
    12 
       
    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;
       
    18 
       
    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;
       
    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 
       
    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;
       
    37 
       
    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 
       
    55 (* ----- specific support for domain ---------------------------------------- *)
       
    56 
       
    57 signature DOMAIN_LIBRARY =
       
    58 sig
       
    59   val Imposs : string -> 'a;
       
    60   val cpo_type : theory -> typ -> bool;
       
    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;
       
    86   val ID : term;
       
    87   val oo : term * term -> term;
       
    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 cproj : term -> 'a list -> int -> term;
       
    98   val list_ccomb : term * term list -> term;
       
    99   (*
       
   100    val con_app : string -> ('a * 'b * string) list -> term;
       
   101    *)
       
   102   val con_app2 : string -> ('a -> term) -> 'a list -> term;
       
   103   val proj : term -> 'a list -> int -> term;
       
   104   val prj : ('a -> 'b -> 'a) -> ('a -> 'b -> 'a) -> 'a -> 'b list -> int -> 'a;
       
   105   val mk_ctuple_pat : term list -> term;
       
   106   val mk_branch : term -> term;
       
   107 
       
   108   (* Creating propositions *)
       
   109   val mk_conj : term * term -> term;
       
   110   val mk_disj : term * term -> term;
       
   111   val mk_imp : term * term -> term;
       
   112   val mk_lam : string * term -> term;
       
   113   val mk_all : string * term -> term;
       
   114   val mk_ex : string * term -> term;
       
   115   val mk_constrain : typ * term -> term;
       
   116   val mk_constrainall : string * typ * term -> term;
       
   117   val === : term * term -> term;
       
   118   val << : term * term -> term;
       
   119   val ~<< : term * term -> term;
       
   120   val strict : term -> term;
       
   121   val defined : term -> term;
       
   122   val mk_adm : term -> term;
       
   123   val mk_compact : term -> term;
       
   124   val lift : ('a -> term) -> 'a list * term -> term;
       
   125   val lift_defined : ('a -> term) -> 'a list * term -> term;
       
   126 
       
   127   (* Creating meta-propositions *)
       
   128   val mk_trp : term -> term; (* HOLogic.mk_Trueprop *)
       
   129   val == : term * term -> term;
       
   130   val ===> : term * term -> term;
       
   131   val ==> : term * term -> term;
       
   132   val mk_All : string * term -> term;
       
   133 
       
   134       (* Domain specifications *)
       
   135       eqtype arg;
       
   136   type cons = string * arg list;
       
   137   type eq = (string * typ list) * cons list;
       
   138   val mk_arg : (bool * Datatype.dtyp) * string option * string -> arg;
       
   139   val is_lazy : arg -> bool;
       
   140   val rec_of : arg -> int;
       
   141   val dtyp_of : arg -> Datatype.dtyp;
       
   142   val sel_of : arg -> string option;
       
   143   val vname : arg -> string;
       
   144   val upd_vname : (string -> string) -> arg -> arg;
       
   145   val is_rec : arg -> bool;
       
   146   val is_nonlazy_rec : arg -> bool;
       
   147   val nonlazy : arg list -> string list;
       
   148   val nonlazy_rec : arg list -> string list;
       
   149   val %# : arg -> term;
       
   150   val /\# : arg * term -> term;
       
   151   val when_body : cons list -> (int * int -> term) -> term;
       
   152   val when_funs : 'a list -> string list;
       
   153   val bound_arg : ''a list -> ''a -> term; (* ''a = arg or string *)
       
   154   val idx_name : 'a list -> string -> int -> string;
       
   155   val app_rec_arg : (int -> term) -> arg -> term;
       
   156   val con_app : string -> arg list -> term;
       
   157   val dtyp_of_eq : eq -> Datatype.dtyp;
       
   158 
       
   159 
       
   160   (* Name mangling *)
       
   161   val strip_esc : string -> string;
       
   162   val extern_name : string -> string;
       
   163   val dis_name : string -> string;
       
   164   val mat_name : string -> string;
       
   165   val pat_name : string -> string;
       
   166   val mk_var_names : string list -> string list;
       
   167 end;
       
   168 
       
   169 structure Domain_Library :> DOMAIN_LIBRARY =
       
   170 struct
       
   171 
       
   172 exception Impossible of string;
       
   173 fun Imposs msg = raise Impossible ("Domain:"^msg);
       
   174 
       
   175 (* ----- name handling ----- *)
       
   176 
       
   177 val strip_esc =
       
   178     let fun strip ("'" :: c :: cs) = c :: strip cs
       
   179           | strip ["'"] = []
       
   180           | strip (c :: cs) = c :: strip cs
       
   181           | strip [] = [];
       
   182     in implode o strip o Symbol.explode end;
       
   183 
       
   184 fun extern_name con =
       
   185     case Symbol.explode con of 
       
   186       ("o"::"p"::" "::rest) => implode rest
       
   187     | _ => con;
       
   188 fun dis_name  con = "is_"^ (extern_name con);
       
   189 fun dis_name_ con = "is_"^ (strip_esc   con);
       
   190 fun mat_name  con = "match_"^ (extern_name con);
       
   191 fun mat_name_ con = "match_"^ (strip_esc   con);
       
   192 fun pat_name  con = (extern_name con) ^ "_pat";
       
   193 fun pat_name_ con = (strip_esc   con) ^ "_pat";
       
   194 
       
   195 (* make distinct names out of the type list, 
       
   196                                        forbidding "o","n..","x..","f..","P.." as names *)
       
   197 (* a number string is added if necessary *)
       
   198 fun mk_var_names ids : string list =
       
   199     let
       
   200       fun nonreserved s = if s mem ["n","x","f","P"] then s^"'" else s;
       
   201       fun index_vnames(vn::vns,occupied) =
       
   202           (case AList.lookup (op =) occupied vn of
       
   203              NONE => if vn mem vns
       
   204                      then (vn^"1") :: index_vnames(vns,(vn,1)  ::occupied)
       
   205                      else  vn      :: index_vnames(vns,          occupied)
       
   206            | SOME(i) => (vn^(string_of_int (i+1)))
       
   207                         :: index_vnames(vns,(vn,i+1)::occupied))
       
   208         | index_vnames([],occupied) = [];
       
   209     in index_vnames(map nonreserved ids, [("O",0),("o",0)]) end;
       
   210 
       
   211 fun cpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort cpo});
       
   212 fun pcpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort pcpo});
       
   213 fun string_of_typ sg = Syntax.string_of_typ_global sg o Sign.certify_typ sg;
       
   214 
       
   215 (* ----- constructor list handling ----- *)
       
   216 
       
   217 type arg =
       
   218      (bool * Datatype.dtyp) *   (*  (lazy, recursive element) *)
       
   219      string option *               (*   selector name    *)
       
   220      string;                       (*   argument name    *)
       
   221 
       
   222 type cons =
       
   223      string *         (* operator name of constr *)
       
   224      arg list;        (* argument list      *)
       
   225 
       
   226 type eq =
       
   227      (string *        (* name      of abstracted type *)
       
   228       typ list) *     (* arguments of abstracted type *)
       
   229      cons list;       (* represented type, as a constructor list *)
       
   230 
       
   231 val mk_arg = I;
       
   232 
       
   233 fun rec_of ((_,dtyp),_,_) =
       
   234     case dtyp of DatatypeAux.DtRec i => i | _ => ~1;
       
   235 (* FIXME: what about indirect recursion? *)
       
   236 
       
   237 fun is_lazy arg = fst (first arg);
       
   238 fun dtyp_of arg = snd (first arg);
       
   239 val sel_of    =       second;
       
   240 val     vname =       third;
       
   241 val upd_vname =   upd_third;
       
   242 fun is_rec         arg = rec_of arg >=0;
       
   243 fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg);
       
   244 fun nonlazy     args   = map vname (filter_out is_lazy    args);
       
   245 fun nonlazy_rec args   = map vname (List.filter is_nonlazy_rec args);
       
   246 
       
   247 
       
   248 (* ----- combinators for making dtyps ----- *)
       
   249 
       
   250 fun mk_uD T = DatatypeAux.DtType(@{type_name "u"}, [T]);
       
   251 fun mk_sprodD (T, U) = DatatypeAux.DtType(@{type_name "**"}, [T, U]);
       
   252 fun mk_ssumD (T, U) = DatatypeAux.DtType(@{type_name "++"}, [T, U]);
       
   253 fun mk_liftD T = DatatypeAux.DtType(@{type_name "lift"}, [T]);
       
   254 val unitD = DatatypeAux.DtType(@{type_name "unit"}, []);
       
   255 val boolD = DatatypeAux.DtType(@{type_name "bool"}, []);
       
   256 val oneD = mk_liftD unitD;
       
   257 val trD = mk_liftD boolD;
       
   258 fun big_sprodD ds = case ds of [] => oneD | _ => foldr1 mk_sprodD ds;
       
   259 fun big_ssumD ds = case ds of [] => unitD | _ => foldr1 mk_ssumD ds;
       
   260 
       
   261 fun dtyp_of_arg ((lazy, D), _, _) = if lazy then mk_uD D else D;
       
   262 fun dtyp_of_cons (_, args) = big_sprodD (map dtyp_of_arg args);
       
   263 fun dtyp_of_eq (_, cons) = big_ssumD (map dtyp_of_cons cons);
       
   264 
       
   265 
       
   266 (* ----- support for type and mixfix expressions ----- *)
       
   267 
       
   268 fun mk_uT T = Type(@{type_name "u"}, [T]);
       
   269 fun mk_cfunT (T, U) = Type(@{type_name "->"}, [T, U]);
       
   270 fun mk_sprodT (T, U) = Type(@{type_name "**"}, [T, U]);
       
   271 fun mk_ssumT (T, U) = Type(@{type_name "++"}, [T, U]);
       
   272 val oneT = @{typ one};
       
   273 val trT = @{typ tr};
       
   274 
       
   275 val op ->> = mk_cfunT;
       
   276 
       
   277 fun mk_TFree s = TFree ("'" ^ s, @{sort pcpo});
       
   278 
       
   279 (* ----- support for term expressions ----- *)
       
   280 
       
   281 fun %: s = Free(s,dummyT);
       
   282 fun %# arg = %:(vname arg);
       
   283 fun %%: s = Const(s,dummyT);
       
   284 
       
   285 local open HOLogic in
       
   286 val mk_trp = mk_Trueprop;
       
   287 fun mk_conj (S,T) = conj $ S $ T;
       
   288 fun mk_disj (S,T) = disj $ S $ T;
       
   289 fun mk_imp  (S,T) = imp  $ S $ T;
       
   290 fun mk_lam  (x,T) = Abs(x,dummyT,T);
       
   291 fun mk_all  (x,P) = HOLogic.mk_all (x,dummyT,P);
       
   292 fun mk_ex   (x,P) = mk_exists (x,dummyT,P);
       
   293 val mk_constrain = uncurry TypeInfer.constrain;
       
   294 fun mk_constrainall (x,typ,P) = %%:"All" $ (TypeInfer.constrain (typ --> boolT) (mk_lam(x,P)));
       
   295 end
       
   296 
       
   297 fun mk_All  (x,P) = %%:"all" $ mk_lam(x,P); (* meta universal quantification *)
       
   298 
       
   299 infixr 0 ===>;  fun S ===> T = %%:"==>" $ S $ T;
       
   300 infixr 0 ==>;   fun S ==> T = mk_trp S ===> mk_trp T;
       
   301 infix 0 ==;     fun S ==  T = %%:"==" $ S $ T;
       
   302 infix 1 ===;    fun S === T = %%:"op =" $ S $ T;
       
   303 infix 1 ~=;     fun S ~=  T = HOLogic.mk_not (S === T);
       
   304 infix 1 <<;     fun S <<  T = %%: @{const_name Porder.below} $ S $ T;
       
   305 infix 1 ~<<;    fun S ~<< T = HOLogic.mk_not (S << T);
       
   306 
       
   307 infix 9 `  ; fun f ` x = %%: @{const_name Rep_CFun} $ f $ x;
       
   308 infix 9 `% ; fun f`% s = f` %: s;
       
   309 infix 9 `%%; fun f`%%s = f` %%:s;
       
   310 
       
   311 fun mk_adm t = %%: @{const_name adm} $ t;
       
   312 fun mk_compact t = %%: @{const_name compact} $ t;
       
   313 val ID = %%: @{const_name ID};
       
   314 fun mk_strictify t = %%: @{const_name strictify}`t;
       
   315 fun mk_cfst t = %%: @{const_name cfst}`t;
       
   316 fun mk_csnd t = %%: @{const_name csnd}`t;
       
   317 (*val csplitN    = "Cprod.csplit";*)
       
   318 (*val sfstN      = "Sprod.sfst";*)
       
   319 (*val ssndN      = "Sprod.ssnd";*)
       
   320 fun mk_ssplit t = %%: @{const_name ssplit}`t;
       
   321 fun mk_sinl t = %%: @{const_name sinl}`t;
       
   322 fun mk_sinr t = %%: @{const_name sinr}`t;
       
   323 fun mk_sscase (x, y) = %%: @{const_name sscase}`x`y;
       
   324 fun mk_up t = %%: @{const_name up}`t;
       
   325 fun mk_fup (t,u) = %%: @{const_name fup} ` t ` u;
       
   326 val ONE = @{term ONE};
       
   327 val TT = @{term TT};
       
   328 val FF = @{term FF};
       
   329 fun mk_iterate (n,f,z) = %%: @{const_name iterate} $ n ` f ` z;
       
   330 fun mk_fix t = %%: @{const_name fix}`t;
       
   331 fun mk_return t = %%: @{const_name Fixrec.return}`t;
       
   332 val mk_fail = %%: @{const_name Fixrec.fail};
       
   333 
       
   334 fun mk_branch t = %%: @{const_name Fixrec.branch} $ t;
       
   335 
       
   336 val pcpoS = @{sort pcpo};
       
   337 
       
   338 val list_ccomb = Library.foldl (op `); (* continuous version of list_comb *)
       
   339 fun con_app2 con f args = list_ccomb(%%:con,map f args);
       
   340 fun con_app con = con_app2 con %#;
       
   341 fun if_rec  arg f y   = if is_rec arg then f (rec_of arg) else y;
       
   342 fun app_rec_arg p arg = if_rec arg (fn n => fn x => (p n)`x) I (%# arg);
       
   343 fun prj _  _  x (   _::[]) _ = x
       
   344   | prj f1 _  x (_::y::ys) 0 = f1 x y
       
   345   | prj f1 f2 x (y::   ys) j = prj f1 f2 (f2 x y) ys (j-1);
       
   346 fun  proj x      = prj (fn S => K(%%:"fst" $S)) (fn S => K(%%:"snd" $S)) x;
       
   347 fun cproj x      = prj (fn S => K(mk_cfst S)) (fn S => K(mk_csnd S)) x;
       
   348 fun lift tfn = Library.foldr (fn (x,t)=> (mk_trp(tfn x) ===> t));
       
   349 
       
   350 fun /\ v T = %%: @{const_name Abs_CFun} $ mk_lam(v,T);
       
   351 fun /\# (arg,T) = /\ (vname arg) T;
       
   352 infixr 9 oo; fun S oo T = %%: @{const_name cfcomp}`S`T;
       
   353 val UU = %%: @{const_name UU};
       
   354 fun strict f = f`UU === UU;
       
   355 fun defined t = t ~= UU;
       
   356 fun cpair (t,u) = %%: @{const_name cpair}`t`u;
       
   357 fun spair (t,u) = %%: @{const_name spair}`t`u;
       
   358 fun mk_ctuple [] = HOLogic.unit (* used in match_defs *)
       
   359   | mk_ctuple ts = foldr1 cpair ts;
       
   360 fun mk_stuple [] = ONE
       
   361   | mk_stuple ts = foldr1 spair ts;
       
   362 fun mk_ctupleT [] = HOLogic.unitT   (* used in match_defs *)
       
   363   | mk_ctupleT Ts = foldr1 HOLogic.mk_prodT Ts;
       
   364 fun mk_maybeT T = Type ("Fixrec.maybe",[T]);
       
   365 fun cpair_pat (p1,p2) = %%: @{const_name cpair_pat} $ p1 $ p2;
       
   366 val mk_ctuple_pat = foldr1 cpair_pat;
       
   367 fun lift_defined f = lift (fn x => defined (f x));
       
   368 fun bound_arg vns v = Bound (length vns - find_index (fn v' => v' = v) vns - 1);
       
   369 
       
   370 fun cont_eta_contract (Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body)) = 
       
   371     (case cont_eta_contract body  of
       
   372        body' as (Const("Cfun.Rep_CFun",Ta) $ f $ Bound 0) => 
       
   373        if not (0 mem loose_bnos f) then incr_boundvars ~1 f 
       
   374        else   Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body')
       
   375      | body' => Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body'))
       
   376   | cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t
       
   377   | cont_eta_contract t    = t;
       
   378 
       
   379 fun idx_name dnames s n = s^(if length dnames = 1 then "" else string_of_int n);
       
   380 fun when_funs cons = if length cons = 1 then ["f"] 
       
   381                      else mapn (fn n => K("f"^(string_of_int n))) 1 cons;
       
   382 fun when_body cons funarg =
       
   383     let
       
   384       fun one_fun n (_,[]  ) = /\ "dummy" (funarg(1,n))
       
   385         | one_fun n (_,args) = let
       
   386             val l2 = length args;
       
   387             fun idxs m arg = (if is_lazy arg then (fn t => mk_fup (ID, t))
       
   388                               else I) (Bound(l2-m));
       
   389           in cont_eta_contract
       
   390                (foldr'' 
       
   391                   (fn (a,t) => mk_ssplit (/\# (a,t)))
       
   392                   (args,
       
   393                 fn a=> /\#(a,(list_ccomb(funarg(l2,n),mapn idxs 1 args))))
       
   394                ) end;
       
   395     in (if length cons = 1 andalso length(snd(hd cons)) <= 1
       
   396         then mk_strictify else I)
       
   397          (foldr1 mk_sscase (mapn one_fun 1 cons)) end;
       
   398 
       
   399 end; (* struct *)