src/HOLCF/Tools/domain/domain_axioms.ML
changeset 23152 9497234a2743
child 24712 64ed05609568
equal deleted inserted replaced
23151:ed3f6685ff90 23152:9497234a2743
       
     1 (*  Title:      HOLCF/Tools/domain/domain_axioms.ML
       
     2     ID:         $Id$
       
     3     Author:     David von Oheimb
       
     4 
       
     5 Syntax generator for domain command.
       
     6 *)
       
     7 
       
     8 structure Domain_Axioms = struct
       
     9 
       
    10 local
       
    11 
       
    12 open Domain_Library;
       
    13 infixr 0 ===>;infixr 0 ==>;infix 0 == ; 
       
    14 infix 1 ===; infix 1 ~= ; infix 1 <<; infix 1 ~<<;
       
    15 infix 9 `   ; infix 9 `% ; infix 9 `%%; infixr 9 oo;
       
    16 
       
    17 fun calc_axioms comp_dname (eqs : eq list) n (((dname,_),cons) : eq)=
       
    18 let
       
    19 
       
    20 (* ----- axioms and definitions concerning the isomorphism ------------------ *)
       
    21 
       
    22   val dc_abs = %%:(dname^"_abs");
       
    23   val dc_rep = %%:(dname^"_rep");
       
    24   val x_name'= "x";
       
    25   val x_name = idx_name eqs x_name' (n+1);
       
    26   val dnam = Sign.base_name dname;
       
    27 
       
    28   val abs_iso_ax = ("abs_iso", mk_trp(dc_rep`(dc_abs`%x_name') === %:x_name'));
       
    29   val rep_iso_ax = ("rep_iso", mk_trp(dc_abs`(dc_rep`%x_name') === %:x_name'));
       
    30 
       
    31   val when_def = ("when_def",%%:(dname^"_when") == 
       
    32      foldr (uncurry /\ ) (/\x_name'((when_body cons (fn (x,y) =>
       
    33 				Bound(1+length cons+x-y)))`(dc_rep`Bound 0))) (when_funs cons));
       
    34   
       
    35   val copy_def = let
       
    36     fun idxs z x arg = if is_rec arg
       
    37 			 then (cproj (Bound z) eqs (rec_of arg))`Bound(z-x)
       
    38 			 else Bound(z-x);
       
    39     fun one_con (con,args) =
       
    40         foldr /\# (list_ccomb (%%:con, mapn (idxs (length args)) 1 args)) args;
       
    41   in ("copy_def", %%:(dname^"_copy") ==
       
    42        /\"f" (list_ccomb (%%:(dname^"_when"), map one_con cons))) end;
       
    43 
       
    44 (* -- definitions concerning the constructors, discriminators and selectors - *)
       
    45 
       
    46   fun con_def m n (_,args) = let
       
    47     fun idxs z x arg = (if is_lazy arg then fn t => %%:upN`t else I) (Bound(z-x));
       
    48     fun parms vs = mk_stuple (mapn (idxs(length vs)) 1 vs);
       
    49     fun inj y 1 _ = y
       
    50     |   inj y _ 0 = %%:sinlN`y
       
    51     |   inj y i j = %%:sinrN`(inj y (i-1) (j-1));
       
    52   in foldr /\# (dc_abs`(inj (parms args) m n)) args end;
       
    53   
       
    54   val con_defs = mapn (fn n => fn (con,args) =>
       
    55     (extern_name con ^"_def", %%:con == con_def (length cons) n (con,args))) 0 cons;
       
    56   
       
    57   val dis_defs = let
       
    58 	fun ddef (con,_) = (dis_name con ^"_def",%%:(dis_name con) == 
       
    59 		 list_ccomb(%%:(dname^"_when"),map 
       
    60 			(fn (con',args) => (foldr /\#
       
    61 			   (if con'=con then %%:TT_N else %%:FF_N) args)) cons))
       
    62 	in map ddef cons end;
       
    63 
       
    64   val mat_defs = let
       
    65 	fun mdef (con,_) = (mat_name con ^"_def",%%:(mat_name con) == 
       
    66 		 list_ccomb(%%:(dname^"_when"),map 
       
    67 			(fn (con',args) => (foldr /\#
       
    68 			   (if con'=con
       
    69                                then %%:returnN`(mk_ctuple (map (bound_arg args) args))
       
    70                                else %%:failN) args)) cons))
       
    71 	in map mdef cons end;
       
    72 
       
    73   val pat_defs =
       
    74     let
       
    75       fun pdef (con,args) =
       
    76         let
       
    77           val ps = mapn (fn n => fn _ => %:("pat" ^ string_of_int n)) 1 args;
       
    78           val xs = map (bound_arg args) args;
       
    79           val r = Bound (length args);
       
    80           val rhs = case args of [] => %%:returnN ` HOLogic.unit
       
    81                                 | _ => foldr1 cpair_pat ps ` mk_ctuple xs;
       
    82           fun one_con (con',args') = foldr /\# (if con'=con then rhs else %%:failN) args';
       
    83         in (pat_name con ^"_def", list_comb (%%:(pat_name con), ps) == 
       
    84                list_ccomb(%%:(dname^"_when"), map one_con cons))
       
    85         end
       
    86     in map pdef cons end;
       
    87 
       
    88   val sel_defs = let
       
    89 	fun sdef con n arg = Option.map (fn sel => (sel^"_def",%%:sel == 
       
    90 		 list_ccomb(%%:(dname^"_when"),map 
       
    91 			(fn (con',args) => if con'<>con then UU else
       
    92 			 foldr /\# (Bound (length args - n)) args) cons))) (sel_of arg);
       
    93 	in List.mapPartial I (List.concat(map (fn (con,args) => mapn (sdef con) 1 args) cons)) end;
       
    94 
       
    95 
       
    96 (* ----- axiom and definitions concerning induction ------------------------- *)
       
    97 
       
    98   val reach_ax = ("reach", mk_trp(cproj (%%:fixN`%%(comp_dname^"_copy")) eqs n
       
    99 					`%x_name === %:x_name));
       
   100   val take_def = ("take_def",%%:(dname^"_take") == mk_lam("n",cproj
       
   101 	     (%%:iterateN $ Bound 0 ` %%:(comp_dname^"_copy") ` UU) eqs n));
       
   102   val finite_def = ("finite_def",%%:(dname^"_finite") == mk_lam(x_name,
       
   103 	mk_ex("n",(%%:(dname^"_take") $ Bound 0)`Bound 1 === Bound 1)));
       
   104 
       
   105 in (dnam,
       
   106     [abs_iso_ax, rep_iso_ax, reach_ax],
       
   107     [when_def, copy_def] @
       
   108      con_defs @ dis_defs @ mat_defs @ pat_defs @ sel_defs @
       
   109     [take_def, finite_def])
       
   110 end; (* let *)
       
   111 
       
   112 fun infer_props thy = map (apsnd (FixrecPackage.legacy_infer_prop thy));
       
   113 
       
   114 fun add_axioms_i x = snd o PureThy.add_axioms_i (map Thm.no_attributes x);
       
   115 fun add_axioms_infer axms thy = add_axioms_i (infer_props thy axms) thy;
       
   116 
       
   117 fun add_defs_i x = snd o (PureThy.add_defs_i false) (map Thm.no_attributes x);
       
   118 fun add_defs_infer defs thy = add_defs_i (infer_props thy defs) thy;
       
   119 
       
   120 in (* local *)
       
   121 
       
   122 fun add_axioms (comp_dnam, eqs : eq list) thy' = let
       
   123   val comp_dname = Sign.full_name thy' comp_dnam;
       
   124   val dnames = map (fst o fst) eqs;
       
   125   val x_name = idx_name dnames "x"; 
       
   126   fun copy_app dname = %%:(dname^"_copy")`Bound 0;
       
   127   val copy_def = ("copy_def" , %%:(comp_dname^"_copy") ==
       
   128 				    /\"f"(foldr1 cpair (map copy_app dnames)));
       
   129   val bisim_def = ("bisim_def",%%:(comp_dname^"_bisim")==mk_lam("R",
       
   130     let
       
   131       fun one_con (con,args) = let
       
   132 	val nonrec_args = filter_out is_rec args;
       
   133 	val    rec_args = List.filter     is_rec args;
       
   134 	val    recs_cnt = length rec_args;
       
   135 	val allargs     = nonrec_args @ rec_args
       
   136 				      @ map (upd_vname (fn s=> s^"'")) rec_args;
       
   137 	val allvns      = map vname allargs;
       
   138 	fun vname_arg s arg = if is_rec arg then vname arg^s else vname arg;
       
   139 	val vns1        = map (vname_arg "" ) args;
       
   140 	val vns2        = map (vname_arg "'") args;
       
   141 	val allargs_cnt = length nonrec_args + 2*recs_cnt;
       
   142 	val rec_idxs    = (recs_cnt-1) downto 0;
       
   143 	val nonlazy_idxs = map snd (filter_out (fn (arg,_) => is_lazy arg)
       
   144 					 (allargs~~((allargs_cnt-1) downto 0)));
       
   145 	fun rel_app i ra = proj (Bound(allargs_cnt+2)) eqs (rec_of ra) $ 
       
   146 			   Bound (2*recs_cnt-i) $ Bound (recs_cnt-i);
       
   147 	val capps = foldr mk_conj (mk_conj(
       
   148 	   Bound(allargs_cnt+1)===list_ccomb(%%:con,map (bound_arg allvns) vns1),
       
   149 	   Bound(allargs_cnt+0)===list_ccomb(%%:con,map (bound_arg allvns) vns2)))
       
   150            (mapn rel_app 1 rec_args);
       
   151         in foldr mk_ex (Library.foldr mk_conj 
       
   152 			      (map (defined o Bound) nonlazy_idxs,capps)) allvns end;
       
   153       fun one_comp n (_,cons) =mk_all(x_name(n+1),mk_all(x_name(n+1)^"'",mk_imp(
       
   154 	 		proj (Bound 2) eqs n $ Bound 1 $ Bound 0,
       
   155          		foldr1 mk_disj (mk_conj(Bound 1 === UU,Bound 0 === UU)
       
   156 					::map one_con cons))));
       
   157     in foldr1 mk_conj (mapn one_comp 0 eqs)end ));
       
   158   fun add_one (thy,(dnam,axs,dfs)) = thy
       
   159 	|> Theory.add_path dnam
       
   160 	|> add_defs_infer dfs
       
   161 	|> add_axioms_infer axs
       
   162 	|> Theory.parent_path;
       
   163   val thy = Library.foldl add_one (thy', mapn (calc_axioms comp_dname eqs) 0 eqs);
       
   164 in thy |> Theory.add_path comp_dnam  
       
   165        |> add_defs_infer (bisim_def::(if length eqs>1 then [copy_def] else []))
       
   166        |> Theory.parent_path
       
   167 end;
       
   168 
       
   169 end; (* local *)
       
   170 end; (* struct *)