src/HOL/Tools/primrec_package.ML
author wenzelm
Sat, 06 Oct 2007 16:50:04 +0200
changeset 24867 e5b55d7be9bb
parent 24712 64ed05609568
child 25557 ea6b11021e79
permissions -rw-r--r--
simplified interfaces for outer syntax;
     1 (*  Title:      HOL/Tools/primrec_package.ML
     2     ID:         $Id$
     3     Author:     Stefan Berghofer, TU Muenchen and Norbert Voelker, FernUni Hagen
     4 
     5 Package for defining functions on datatypes by primitive recursion.
     6 *)
     7 
     8 signature PRIMREC_PACKAGE =
     9 sig
    10   val quiet_mode: bool ref
    11   val unify_consts: theory -> term list -> term list -> term list * term list
    12   val add_primrec: string -> ((bstring * string) * Attrib.src list) list
    13     -> theory -> thm list * theory
    14   val add_primrec_unchecked: string -> ((bstring * string) * Attrib.src list) list
    15     -> theory -> thm list * theory
    16   val add_primrec_i: string -> ((bstring * term) * attribute list) list
    17     -> theory -> thm list * theory
    18   val add_primrec_unchecked_i: string -> ((bstring * term) * attribute list) list
    19     -> theory -> thm list * theory
    20   (* FIXME !? *)
    21   val gen_primrec: ((bstring * attribute list) * thm list -> theory -> (bstring * thm list) * theory)
    22     -> ((bstring * attribute list) * term -> theory -> (bstring * thm) * theory)
    23     -> string -> ((bstring * attribute list) * term) list
    24     -> theory -> thm list * theory;
    25 end;
    26 
    27 structure PrimrecPackage : PRIMREC_PACKAGE =
    28 struct
    29 
    30 open DatatypeAux;
    31 
    32 exception RecError of string;
    33 
    34 fun primrec_err s = error ("Primrec definition error:\n" ^ s);
    35 fun primrec_eq_err thy s eq =
    36   primrec_err (s ^ "\nin\n" ^ quote (Sign.string_of_term thy eq));
    37 
    38 
    39 (* messages *)
    40 
    41 val quiet_mode = ref false;
    42 fun message s = if ! quiet_mode then () else writeln s;
    43 
    44 
    45 (*the following code ensures that each recursive set always has the
    46   same type in all introduction rules*)
    47 fun unify_consts thy cs intr_ts =
    48   (let
    49     val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
    50     fun varify (t, (i, ts)) =
    51       let val t' = map_types (Logic.incr_tvar (i + 1)) (snd (Type.varify [] t))
    52       in (maxidx_of_term t', t'::ts) end;
    53     val (i, cs') = foldr varify (~1, []) cs;
    54     val (i', intr_ts') = foldr varify (i, []) intr_ts;
    55     val rec_consts = fold add_term_consts_2 cs' [];
    56     val intr_consts = fold add_term_consts_2 intr_ts' [];
    57     fun unify (cname, cT) =
    58       let val consts = map snd (filter (fn (c, _) => c = cname) intr_consts)
    59       in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
    60     val (env, _) = fold unify rec_consts (Vartab.empty, i');
    61     val subst = Type.freeze o map_types (Envir.norm_type env)
    62 
    63   in (map subst cs', map subst intr_ts')
    64   end) handle Type.TUNIFY =>
    65     (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
    66 
    67 
    68 (* preprocessing of equations *)
    69 
    70 fun process_eqn thy eq rec_fns =
    71   let
    72     val (lhs, rhs) =
    73       if null (term_vars eq) then
    74         HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
    75         handle TERM _ => raise RecError "not a proper equation"
    76       else raise RecError "illegal schematic variable(s)";
    77 
    78     val (recfun, args) = strip_comb lhs;
    79     val fnameT = dest_Const recfun handle TERM _ =>
    80       raise RecError "function is not declared as constant in theory";
    81 
    82     val (ls', rest)  = take_prefix is_Free args;
    83     val (middle, rs') = take_suffix is_Free rest;
    84     val rpos = length ls';
    85 
    86     val (constr, cargs') = if null middle then raise RecError "constructor missing"
    87       else strip_comb (hd middle);
    88     val (cname, T) = dest_Const constr
    89       handle TERM _ => raise RecError "ill-formed constructor";
    90     val (tname, _) = dest_Type (body_type T) handle TYPE _ =>
    91       raise RecError "cannot determine datatype associated with function"
    92 
    93     val (ls, cargs, rs) =
    94       (map dest_Free ls', map dest_Free cargs', map dest_Free rs')
    95       handle TERM _ => raise RecError "illegal argument in pattern";
    96     val lfrees = ls @ rs @ cargs;
    97 
    98     fun check_vars _ [] = ()
    99       | check_vars s vars = raise RecError (s ^ commas_quote (map fst vars))
   100   in
   101     if length middle > 1 then
   102       raise RecError "more than one non-variable in pattern"
   103     else
   104      (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
   105       check_vars "extra variables on rhs: "
   106         (map dest_Free (term_frees rhs) \\ lfrees);
   107       case AList.lookup (op =) rec_fns fnameT of
   108         NONE =>
   109           (fnameT, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns
   110       | SOME (_, rpos', eqns) =>
   111           if AList.defined (op =) eqns cname then
   112             raise RecError "constructor already occurred as pattern"
   113           else if rpos <> rpos' then
   114             raise RecError "position of recursive argument inconsistent"
   115           else
   116             AList.update (op =) (fnameT, (tname, rpos, (cname, (ls, cargs, rs, rhs, eq))::eqns))
   117               rec_fns)
   118   end
   119   handle RecError s => primrec_eq_err thy s eq;
   120 
   121 fun process_fun thy descr rec_eqns (i, fnameT as (fname, _)) (fnameTs, fnss) =
   122   let
   123     val (_, (tname, _, constrs)) = List.nth (descr, i);
   124 
   125     (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *)
   126 
   127     fun subst [] t fs = (t, fs)
   128       | subst subs (Abs (a, T, t)) fs =
   129           fs
   130           |> subst subs t
   131           |-> (fn t' => pair (Abs (a, T, t')))
   132       | subst subs (t as (_ $ _)) fs =
   133           let
   134             val (f, ts) = strip_comb t;
   135           in
   136             if is_Const f andalso dest_Const f mem map fst rec_eqns then
   137               let
   138                 val fnameT' as (fname', _) = dest_Const f;
   139                 val (_, rpos, _) = the (AList.lookup (op =) rec_eqns fnameT');
   140                 val ls = Library.take (rpos, ts);
   141                 val rest = Library.drop (rpos, ts);
   142                 val (x', rs) = (hd rest, tl rest)
   143                   handle Empty => raise RecError ("not enough arguments\
   144                    \ in recursive application\nof function " ^ quote fname' ^ " on rhs");
   145                 val (x, xs) = strip_comb x'
   146               in case AList.lookup (op =) subs x
   147                of NONE =>
   148                     fs
   149                     |> fold_map (subst subs) ts
   150                     |-> (fn ts' => pair (list_comb (f, ts')))
   151                 | SOME (i', y) =>
   152                     fs
   153                     |> fold_map (subst subs) (xs @ ls @ rs)
   154                     ||> process_fun thy descr rec_eqns (i', fnameT')
   155                     |-> (fn ts' => pair (list_comb (y, ts')))
   156               end
   157             else
   158               fs
   159               |> fold_map (subst subs) (f :: ts)
   160               |-> (fn (f'::ts') => pair (list_comb (f', ts')))
   161           end
   162       | subst _ t fs = (t, fs);
   163 
   164     (* translate rec equations into function arguments suitable for rec comb *)
   165 
   166     fun trans eqns (cname, cargs) (fnameTs', fnss', fns) =
   167       (case AList.lookup (op =) eqns cname of
   168           NONE => (warning ("No equation for constructor " ^ quote cname ^
   169             "\nin definition of function " ^ quote fname);
   170               (fnameTs', fnss', (Const ("HOL.undefined", dummyT))::fns))
   171         | SOME (ls, cargs', rs, rhs, eq) =>
   172             let
   173               val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
   174               val rargs = map fst recs;
   175               val subs = map (rpair dummyT o fst)
   176                 (rev (rename_wrt_term rhs rargs));
   177               val (rhs', (fnameTs'', fnss'')) =
   178                   (subst (map (fn ((x, y), z) =>
   179                                (Free x, (body_index y, Free z)))
   180                           (recs ~~ subs)) rhs (fnameTs', fnss'))
   181                   handle RecError s => primrec_eq_err thy s eq
   182             in (fnameTs'', fnss'',
   183                 (list_abs_free (cargs' @ subs @ ls @ rs, rhs'))::fns)
   184             end)
   185 
   186   in (case AList.lookup (op =) fnameTs i of
   187       NONE =>
   188         if exists (equal fnameT o snd) fnameTs then
   189           raise RecError ("inconsistent functions for datatype " ^ quote tname)
   190         else
   191           let
   192             val (_, _, eqns) = the (AList.lookup (op =) rec_eqns fnameT);
   193             val (fnameTs', fnss', fns) = fold_rev (trans eqns) constrs
   194               ((i, fnameT)::fnameTs, fnss, [])
   195           in
   196             (fnameTs', (i, (fname, #1 (snd (hd eqns)), fns))::fnss')
   197           end
   198     | SOME fnameT' =>
   199         if fnameT = fnameT' then (fnameTs, fnss)
   200         else raise RecError ("inconsistent functions for datatype " ^ quote tname))
   201   end;
   202 
   203 
   204 (* prepare functions needed for definitions *)
   205 
   206 fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) =
   207   case AList.lookup (op =) fns i of
   208      NONE =>
   209        let
   210          val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined",
   211            replicate ((length cargs) + (length (List.filter is_rec_type cargs)))
   212              dummyT ---> HOLogic.unitT)) constrs;
   213          val _ = warning ("No function definition for datatype " ^ quote tname)
   214        in
   215          (dummy_fns @ fs, defs)
   216        end
   217    | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs);
   218 
   219 
   220 (* make definition *)
   221 
   222 fun make_def thy fs (fname, ls, rec_name, tname) =
   223   let
   224     val rhs = fold_rev (fn T => fn t => Abs ("", T, t))
   225                     ((map snd ls) @ [dummyT])
   226                     (list_comb (Const (rec_name, dummyT),
   227                                 fs @ map Bound (0 ::(length ls downto 1))))
   228     val def_name = Sign.base_name fname ^ "_" ^ Sign.base_name tname ^ "_def";
   229     val def_prop =
   230       singleton (Syntax.check_terms (ProofContext.init thy))
   231         (Logic.mk_equals (Const (fname, dummyT), rhs));
   232   in (def_name, def_prop) end;
   233 
   234 
   235 (* find datatypes which contain all datatypes in tnames' *)
   236 
   237 fun find_dts (dt_info : datatype_info Symtab.table) _ [] = []
   238   | find_dts dt_info tnames' (tname::tnames) =
   239       (case Symtab.lookup dt_info tname of
   240           NONE => primrec_err (quote tname ^ " is not a datatype")
   241         | SOME dt =>
   242             if tnames' subset (map (#1 o snd) (#descr dt)) then
   243               (tname, dt)::(find_dts dt_info tnames' tnames)
   244             else find_dts dt_info tnames' tnames);
   245 
   246 fun prepare_induct ({descr, induction, ...}: datatype_info) rec_eqns =
   247   let
   248     fun constrs_of (_, (_, _, cs)) =
   249       map (fn (cname:string, (_, cargs, _, _, _)) => (cname, map fst cargs)) cs;
   250     val params_of = these o AList.lookup (op =) (List.concat (map constrs_of rec_eqns));
   251   in
   252     induction
   253     |> RuleCases.rename_params (map params_of (List.concat (map (map #1 o #3 o #2) descr)))
   254     |> RuleCases.save induction
   255   end;
   256 
   257 local
   258 
   259 fun gen_primrec_i note def alt_name eqns_atts thy =
   260   let
   261     val (eqns, atts) = split_list eqns_atts;
   262     val dt_info = DatatypePackage.get_datatypes thy;
   263     val rec_eqns = fold_rev (process_eqn thy o snd) eqns [] ;
   264     val tnames = distinct (op =) (map (#1 o snd) rec_eqns);
   265     val dts = find_dts dt_info tnames tnames;
   266     val main_fns =
   267       map (fn (tname, {index, ...}) =>
   268         (index,
   269           (fst o the o find_first (fn f => (#1 o snd) f = tname)) rec_eqns))
   270       dts;
   271     val {descr, rec_names, rec_rewrites, ...} =
   272       if null dts then
   273         primrec_err ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
   274       else snd (hd dts);
   275     val (fnameTs, fnss) =
   276       fold_rev (process_fun thy descr rec_eqns) main_fns ([], []);
   277     val (fs, defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []);
   278     val defs' = map (make_def thy fs) defs;
   279     val nameTs1 = map snd fnameTs;
   280     val nameTs2 = map fst rec_eqns;
   281     val _ = if gen_eq_set (op =) (nameTs1, nameTs2) then ()
   282             else primrec_err ("functions " ^ commas_quote (map fst nameTs2) ^
   283               "\nare not mutually recursive");
   284     val primrec_name =
   285       if alt_name = "" then (space_implode "_" (map (Sign.base_name o #1) defs)) else alt_name;
   286     val (defs_thms', thy') =
   287       thy
   288       |> Sign.add_path primrec_name
   289       |> fold_map def (map (fn (name, t) => ((name, []), t)) defs');
   290     val rewrites = (map mk_meta_eq rec_rewrites) @ map snd defs_thms';
   291     val _ = message ("Proving equations for primrec function(s) " ^
   292       commas_quote (map fst nameTs1) ^ " ...");
   293     val simps = map (fn (_, t) => Goal.prove_global thy' [] [] t
   294         (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns;
   295     val (simps', thy'') =
   296       thy'
   297       |> fold_map note ((map fst eqns ~~ atts) ~~ map single simps);
   298     val simps'' = maps snd simps';
   299   in
   300     thy''
   301     |> note (("simps", [Simplifier.simp_add, RecfunCodegen.add_default]), simps'')
   302     |> snd
   303     |> note (("induct", []), [prepare_induct (#2 (hd dts)) rec_eqns])
   304     |> snd
   305     |> Sign.parent_path
   306     |> pair simps''
   307   end;
   308 
   309 fun gen_primrec note def alt_name eqns thy =
   310   let
   311     val ((names, strings), srcss) = apfst split_list (split_list eqns);
   312     val atts = map (map (Attrib.attribute thy)) srcss;
   313     val eqn_ts = map (fn s => Syntax.read_prop_global thy s
   314       handle ERROR msg => cat_error msg ("The error(s) above occurred for " ^ s)) strings;
   315     val rec_ts = map (fn eq => head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop eq)))
   316       handle TERM _ => primrec_eq_err thy "not a proper equation" eq) eqn_ts;
   317     val (_, eqn_ts') = unify_consts thy rec_ts eqn_ts
   318   in
   319     gen_primrec_i note def alt_name (names ~~ eqn_ts' ~~ atts) thy
   320   end;
   321 
   322 fun thy_note ((name, atts), thms) =
   323   PureThy.add_thmss [((name, thms), atts)] #-> (fn [thms] => pair (name, thms));
   324 fun thy_def false ((name, atts), t) =
   325       PureThy.add_defs_i false [((name, t), atts)] #-> (fn [thm] => pair (name, thm))
   326   | thy_def true ((name, atts), t) =
   327       PureThy.add_defs_unchecked_i false [((name, t), atts)] #-> (fn [thm] => pair (name, thm));
   328 
   329 in
   330 
   331 val add_primrec = gen_primrec thy_note (thy_def false);
   332 val add_primrec_unchecked = gen_primrec thy_note (thy_def true);
   333 val add_primrec_i = gen_primrec_i thy_note (thy_def false);
   334 val add_primrec_unchecked_i = gen_primrec_i thy_note (thy_def true);
   335 fun gen_primrec note def alt_name specs =
   336   gen_primrec_i note def alt_name (map (fn ((name, t), atts) => ((name, atts), t)) specs);
   337 
   338 end;
   339 
   340 
   341 (* outer syntax *)
   342 
   343 local structure P = OuterParse and K = OuterKeyword in
   344 
   345 val opt_unchecked_name =
   346   Scan.optional (P.$$$ "(" |-- P.!!!
   347     (((P.$$$ "unchecked" >> K true) -- Scan.optional P.name "" ||
   348       P.name >> pair false) --| P.$$$ ")")) (false, "");
   349 
   350 val primrec_decl =
   351   opt_unchecked_name -- Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop);
   352 
   353 val _ =
   354   OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl
   355     (primrec_decl >> (fn ((unchecked, alt_name), eqns) =>
   356       Toplevel.theory (snd o
   357         (if unchecked then add_primrec_unchecked else add_primrec) alt_name
   358           (map P.triple_swap eqns))));
   359 
   360 end;
   361 
   362 end;