src/HOLCF/Tools/Domain/domain_library.ML
changeset 40026 8f8f18a88685
parent 40025 876689e6bbdf
child 40027 98f2d8280eb4
equal deleted inserted replaced
40025:876689e6bbdf 40026:8f8f18a88685
     1 (*  Title:      HOLCF/Tools/Domain/domain_library.ML
       
     2     Author:     David von Oheimb
       
     3 
       
     4 Library for domain command.
       
     5 *)
       
     6 
       
     7 
       
     8 (* infix syntax *)
       
     9 
       
    10 infixr 5 -->;
       
    11 infixr 6 ->>;
       
    12 infixr 0 ===>;
       
    13 infixr 0 ==>;
       
    14 infix 0 ==;
       
    15 infix 1 ===;
       
    16 infix 1 ~=;
       
    17 
       
    18 infix 9 `  ;
       
    19 infix 9 `% ;
       
    20 infix 9 `%%;
       
    21 
       
    22 
       
    23 (* ----- specific support for domain ---------------------------------------- *)
       
    24 
       
    25 signature DOMAIN_LIBRARY =
       
    26 sig
       
    27   val first  : 'a * 'b * 'c -> 'a
       
    28   val second : 'a * 'b * 'c -> 'b
       
    29   val third  : 'a * 'b * 'c -> 'c
       
    30   val upd_second : ('b -> 'd) -> 'a * 'b * 'c -> 'a * 'd * 'c
       
    31   val upd_third  : ('c -> 'd) -> 'a * 'b * 'c -> 'a * 'b * 'd
       
    32   val mapn : (int -> 'a -> 'b) -> int -> 'a list -> 'b list
       
    33   val atomize : Proof.context -> thm -> thm list
       
    34 
       
    35   val Imposs : string -> 'a;
       
    36   val cpo_type : theory -> typ -> bool;
       
    37   val pcpo_type : theory -> typ -> bool;
       
    38   val string_of_typ : theory -> typ -> string;
       
    39 
       
    40   (* Creating HOLCF types *)
       
    41   val mk_ssumT : typ * typ -> typ;
       
    42   val mk_sprodT : typ * typ -> typ;
       
    43   val mk_uT : typ -> typ;
       
    44   val oneT : typ;
       
    45   val pcpoS : sort;
       
    46 
       
    47   (* Creating HOLCF terms *)
       
    48   val %: : string -> term;
       
    49   val %%: : string -> term;
       
    50   val ` : term * term -> term;
       
    51   val `% : term * string -> term;
       
    52   val UU : term;
       
    53   val ID : term;
       
    54   val list_ccomb : term * term list -> term;
       
    55   val con_app2 : string -> ('a -> term) -> 'a list -> term;
       
    56   val prj : ('a -> 'b -> 'a) -> ('a -> 'b -> 'a) -> 'a -> 'b list -> int -> 'a
       
    57   val proj : term -> 'a list -> int -> term;
       
    58 
       
    59   (* Creating propositions *)
       
    60   val mk_conj : term * term -> term;
       
    61   val mk_disj : term * term -> term;
       
    62   val mk_imp : term * term -> term;
       
    63   val mk_lam : string * term -> term;
       
    64   val mk_all : string * term -> term;
       
    65   val mk_ex : string * term -> term;
       
    66   val mk_constrainall : string * typ * term -> term;
       
    67   val === : term * term -> term;
       
    68   val defined : term -> term;
       
    69   val mk_adm : term -> term;
       
    70   val lift : ('a -> term) -> 'a list * term -> term;
       
    71   val lift_defined : ('a -> term) -> 'a list * term -> term;
       
    72 
       
    73   (* Creating meta-propositions *)
       
    74   val mk_trp : term -> term; (* HOLogic.mk_Trueprop *)
       
    75   val == : term * term -> term;
       
    76   val ===> : term * term -> term;
       
    77   val mk_All : string * term -> term;
       
    78 
       
    79       (* Domain specifications *)
       
    80       eqtype arg;
       
    81   type cons = string * arg list;
       
    82   type eq = (string * typ list) * cons list;
       
    83   val mk_arg : (bool * Datatype.dtyp) * string -> arg;
       
    84   val is_lazy : arg -> bool;
       
    85   val rec_of : arg -> int;
       
    86   val dtyp_of : arg -> Datatype.dtyp;
       
    87   val vname : arg -> string;
       
    88   val upd_vname : (string -> string) -> arg -> arg;
       
    89   val is_rec : arg -> bool;
       
    90   val is_nonlazy_rec : arg -> bool;
       
    91   val nonlazy : arg list -> string list;
       
    92   val nonlazy_rec : arg list -> string list;
       
    93   val %# : arg -> term;
       
    94   val bound_arg : ''a list -> ''a -> term; (* ''a = arg or string *)
       
    95   val idx_name : 'a list -> string -> int -> string;
       
    96   val con_app : string -> arg list -> term;
       
    97 end;
       
    98 
       
    99 structure Domain_Library :> DOMAIN_LIBRARY =
       
   100 struct
       
   101 
       
   102 fun first  (x,_,_) = x;
       
   103 fun second (_,x,_) = x;
       
   104 fun third  (_,_,x) = x;
       
   105 
       
   106 fun upd_first  f (x,y,z) = (f x,   y,   z);
       
   107 fun upd_second f (x,y,z) = (  x, f y,   z);
       
   108 fun upd_third  f (x,y,z) = (  x,   y, f z);
       
   109 
       
   110 fun mapn f n []      = []
       
   111   | mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs;
       
   112 
       
   113 fun foldr'' f (l,f2) =
       
   114     let fun itr []  = raise Fail "foldr''" 
       
   115           | itr [a] = f2 a
       
   116           | itr (a::l) = f(a, itr l)
       
   117     in  itr l  end;
       
   118 
       
   119 fun atomize ctxt thm =
       
   120     let
       
   121       val r_inst = read_instantiate ctxt;
       
   122       fun at thm =
       
   123           case concl_of thm of
       
   124             _$(Const(@{const_name HOL.conj},_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
       
   125           | _$(Const(@{const_name All} ,_)$Abs(s,_,_))=> at(thm RS (r_inst [(("x", 0), "?" ^ s)] spec))
       
   126           | _                             => [thm];
       
   127     in map zero_var_indexes (at thm) end;
       
   128 
       
   129 exception Impossible of string;
       
   130 fun Imposs msg = raise Impossible ("Domain:"^msg);
       
   131 
       
   132 fun cpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort cpo});
       
   133 fun pcpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort pcpo});
       
   134 fun string_of_typ sg = Syntax.string_of_typ_global sg o Sign.certify_typ sg;
       
   135 
       
   136 (* ----- constructor list handling ----- *)
       
   137 
       
   138 type arg =
       
   139      (bool * Datatype.dtyp) *   (*  (lazy, recursive element) *)
       
   140      string;                       (*   argument name    *)
       
   141 
       
   142 type cons =
       
   143      string *         (* operator name of constr *)
       
   144      arg list;        (* argument list      *)
       
   145 
       
   146 type eq =
       
   147      (string *        (* name      of abstracted type *)
       
   148       typ list) *     (* arguments of abstracted type *)
       
   149      cons list;       (* represented type, as a constructor list *)
       
   150 
       
   151 val mk_arg = I;
       
   152 
       
   153 fun rec_of ((_,dtyp),_) =
       
   154     case dtyp of Datatype_Aux.DtRec i => i | _ => ~1;
       
   155 (* FIXME: what about indirect recursion? *)
       
   156 
       
   157 fun is_lazy arg = fst (fst arg);
       
   158 fun dtyp_of arg = snd (fst arg);
       
   159 val     vname =       snd;
       
   160 val upd_vname =   apsnd;
       
   161 fun is_rec         arg = rec_of arg >=0;
       
   162 fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg);
       
   163 fun nonlazy     args   = map vname (filter_out is_lazy args);
       
   164 fun nonlazy_rec args   = map vname (filter is_nonlazy_rec args);
       
   165 
       
   166 
       
   167 (* ----- support for type and mixfix expressions ----- *)
       
   168 
       
   169 fun mk_uT T = Type(@{type_name "u"}, [T]);
       
   170 fun mk_sprodT (T, U) = Type(@{type_name sprod}, [T, U]);
       
   171 fun mk_ssumT (T, U) = Type(@{type_name ssum}, [T, U]);
       
   172 val oneT = @{typ one};
       
   173 
       
   174 (* ----- support for term expressions ----- *)
       
   175 
       
   176 fun %: s = Free(s,dummyT);
       
   177 fun %# arg = %:(vname arg);
       
   178 fun %%: s = Const(s,dummyT);
       
   179 
       
   180 local open HOLogic in
       
   181 val mk_trp = mk_Trueprop;
       
   182 fun mk_conj (S,T) = conj $ S $ T;
       
   183 fun mk_disj (S,T) = disj $ S $ T;
       
   184 fun mk_imp  (S,T) = imp  $ S $ T;
       
   185 fun mk_lam  (x,T) = Abs(x,dummyT,T);
       
   186 fun mk_all  (x,P) = HOLogic.mk_all (x,dummyT,P);
       
   187 fun mk_ex   (x,P) = mk_exists (x,dummyT,P);
       
   188 fun mk_constrainall (x,typ,P) = %%: @{const_name All} $ (Type.constraint (typ --> boolT) (mk_lam(x,P)));
       
   189 end
       
   190 
       
   191 fun mk_All  (x,P) = %%:"all" $ mk_lam(x,P); (* meta universal quantification *)
       
   192 
       
   193 infixr 0 ===>;  fun S ===> T = %%: "==>" $ S $ T;
       
   194 infix 0 ==;     fun S ==  T = %%: "==" $ S $ T;
       
   195 infix 1 ===;    fun S === T = %%: @{const_name HOL.eq} $ S $ T;
       
   196 infix 1 ~=;     fun S ~=  T = HOLogic.mk_not (S === T);
       
   197 
       
   198 infix 9 `  ; fun f ` x = %%: @{const_name Rep_CFun} $ f $ x;
       
   199 infix 9 `% ; fun f`% s = f` %: s;
       
   200 infix 9 `%%; fun f`%%s = f` %%:s;
       
   201 
       
   202 fun mk_adm t = %%: @{const_name adm} $ t;
       
   203 val ID = %%: @{const_name ID};
       
   204 fun mk_strictify t = %%: @{const_name strictify}`t;
       
   205 fun mk_ssplit t = %%: @{const_name ssplit}`t;
       
   206 fun mk_sscase (x, y) = %%: @{const_name sscase}`x`y;
       
   207 fun mk_fup (t,u) = %%: @{const_name fup} ` t ` u;
       
   208 
       
   209 val pcpoS = @{sort pcpo};
       
   210 
       
   211 val list_ccomb = Library.foldl (op `); (* continuous version of list_comb *)
       
   212 fun con_app2 con f args = list_ccomb(%%:con,map f args);
       
   213 fun con_app con = con_app2 con %#;
       
   214 fun prj _  _  x (   _::[]) _ = x
       
   215   | prj _  _  _ []         _ = raise Fail "Domain_Library.prj: empty list"
       
   216   | prj f1 _  x (_::y::ys) 0 = f1 x y
       
   217   | prj f1 f2 x (y::   ys) j = prj f1 f2 (f2 x y) ys (j-1);
       
   218 fun proj x = prj (fn S => K (%%: @{const_name fst} $ S)) (fn S => K (%%: @{const_name snd} $ S)) x;
       
   219 fun lift tfn = Library.foldr (fn (x,t)=> (mk_trp(tfn x) ===> t));
       
   220 
       
   221 val UU = %%: @{const_name UU};
       
   222 fun defined t = t ~= UU;
       
   223 fun cpair (t,u) = %%: @{const_name Pair} $ t $ u;
       
   224 fun spair (t,u) = %%: @{const_name spair}`t`u;
       
   225 fun lift_defined f = lift (fn x => defined (f x));
       
   226 fun bound_arg vns v = Bound (length vns - find_index (fn v' => v' = v) vns - 1);
       
   227 
       
   228 fun cont_eta_contract (Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body)) = 
       
   229     (case cont_eta_contract body  of
       
   230        body' as (Const("Cfun.Rep_CFun",Ta) $ f $ Bound 0) => 
       
   231        if not (member (op =) (loose_bnos f) 0) then incr_boundvars ~1 f 
       
   232        else   Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body')
       
   233      | body' => Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body'))
       
   234   | cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t
       
   235   | cont_eta_contract t    = t;
       
   236 
       
   237 fun idx_name dnames s n = s^(if length dnames = 1 then "" else string_of_int n);
       
   238 
       
   239 end; (* struct *)