src/HOL/Tools/recdef_package.ML
author wenzelm
Wed, 19 Oct 2005 21:52:32 +0200
changeset 17920 5c3e9415c653
parent 17593 01870f76478c
child 18377 0e1d025d57b3
permissions -rw-r--r--
removed obsolete add_recdef_old;
tuned;
     1 (*  Title:      HOL/Tools/recdef_package.ML
     2     ID:         $Id$
     3     Author:     Markus Wenzel, TU Muenchen
     4 
     5 Wrapper module for Konrad Slind's TFL package.
     6 *)
     7 
     8 signature RECDEF_PACKAGE =
     9 sig
    10   val quiet_mode: bool ref
    11   val print_recdefs: theory -> unit
    12   val get_recdef: theory -> string
    13     -> {simps: thm list, rules: thm list list, induct: thm, tcs: term list} option
    14   val simp_add_global: theory attribute
    15   val simp_del_global: theory attribute
    16   val cong_add_global: theory attribute
    17   val cong_del_global: theory attribute
    18   val wf_add_global: theory attribute
    19   val wf_del_global: theory attribute
    20   val simp_add_local: Proof.context attribute
    21   val simp_del_local: Proof.context attribute
    22   val cong_add_local: Proof.context attribute
    23   val cong_del_local: Proof.context attribute
    24   val wf_add_local: Proof.context attribute
    25   val wf_del_local: Proof.context attribute
    26   val add_recdef: bool -> xstring -> string -> ((bstring * string) * Attrib.src list) list ->
    27     Attrib.src option -> theory -> theory
    28       * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
    29   val add_recdef_i: bool -> xstring -> term -> ((bstring * term) * theory attribute list) list ->
    30     theory -> theory * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
    31   val defer_recdef: xstring -> string list -> (thmref * Attrib.src list) list
    32     -> theory -> theory * {induct_rules: thm}
    33   val defer_recdef_i: xstring -> term list -> (thm list * theory attribute list) list
    34     -> theory -> theory * {induct_rules: thm}
    35   val recdef_tc: bstring * Attrib.src list -> xstring -> int option -> theory -> Proof.state
    36   val recdef_tc_i: bstring * theory attribute list -> string -> int option
    37     -> theory -> Proof.state
    38   val setup: (theory -> theory) list
    39 end;
    40 
    41 structure RecdefPackage: RECDEF_PACKAGE =
    42 struct
    43 
    44 
    45 val quiet_mode = Tfl.quiet_mode;
    46 val message = Tfl.message;
    47 
    48 
    49 (** recdef hints **)
    50 
    51 (* type hints *)
    52 
    53 type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
    54 
    55 fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints;
    56 fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs));
    57 
    58 fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs));
    59 fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs));
    60 fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs));
    61 
    62 fun pretty_hints ({simps, congs, wfs}: hints) =
    63  [Pretty.big_list "recdef simp hints:" (map Display.pretty_thm simps),
    64   Pretty.big_list "recdef cong hints:" (map Display.pretty_thm (map #2 congs)),
    65   Pretty.big_list "recdef wf hints:" (map Display.pretty_thm wfs)];
    66 
    67 
    68 (* congruence rules *)
    69 
    70 local
    71 
    72 val cong_head =
    73   fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of;
    74 
    75 fun prep_cong raw_thm =
    76   let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end;
    77 
    78 in
    79 
    80 fun add_cong raw_thm congs =
    81   let val (c, thm) = prep_cong raw_thm
    82   in overwrite_warn (congs, (c, thm)) ("Overwriting recdef congruence rule for " ^ quote c) end;
    83 
    84 fun del_cong raw_thm congs =
    85   let
    86     val (c, thm) = prep_cong raw_thm;
    87     val (del, rest) = Library.partition (Library.equal c o fst) congs;
    88   in if null del then (warning ("No recdef congruence rule for " ^ quote c); congs) else rest end;
    89 
    90 val add_congs = foldr (uncurry add_cong);
    91 
    92 end;
    93 
    94 
    95 
    96 (** global and local recdef data **)
    97 
    98 (* theory data *)
    99 
   100 type recdef_info = {simps: thm list, rules: thm list list, induct: thm, tcs: term list};
   101 
   102 structure GlobalRecdefData = TheoryDataFun
   103 (struct
   104   val name = "HOL/recdef";
   105   type T = recdef_info Symtab.table * hints;
   106 
   107   val empty = (Symtab.empty, mk_hints ([], [], [])): T;
   108   val copy = I;
   109   val extend = I;
   110   fun merge _
   111    ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
   112     (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T =
   113       (Symtab.merge (K true) (tab1, tab2),
   114         mk_hints (Drule.merge_rules (simps1, simps2),
   115           Library.merge_alists congs1 congs2,
   116           Drule.merge_rules (wfs1, wfs2)));
   117 
   118   fun print thy (tab, hints) =
   119     (Pretty.strs ("recdefs:" :: map #1 (NameSpace.extern_table (Sign.const_space thy, tab))) ::
   120       pretty_hints hints) |> Pretty.chunks |> Pretty.writeln;
   121 end);
   122 
   123 val print_recdefs = GlobalRecdefData.print;
   124 
   125 
   126 val get_recdef = Symtab.lookup o #1 o GlobalRecdefData.get;
   127 
   128 fun put_recdef name info thy =
   129   let
   130     val (tab, hints) = GlobalRecdefData.get thy;
   131     val tab' = Symtab.update_new (name, info) tab
   132       handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name);
   133   in GlobalRecdefData.put (tab', hints) thy end;
   134 
   135 val get_global_hints = #2 o GlobalRecdefData.get;
   136 val map_global_hints = GlobalRecdefData.map o apsnd;
   137 
   138 
   139 (* proof data *)
   140 
   141 structure LocalRecdefData = ProofDataFun
   142 (struct
   143   val name = "HOL/recdef";
   144   type T = hints;
   145   val init = get_global_hints;
   146   fun print _ hints = pretty_hints hints |> Pretty.chunks |> Pretty.writeln;
   147 end);
   148 
   149 val get_local_hints = LocalRecdefData.get;
   150 val map_local_hints = LocalRecdefData.map;
   151 
   152 
   153 (* attributes *)
   154 
   155 local
   156 
   157 fun global_local f g =
   158  (fn (thy, thm) => (map_global_hints (f (g thm)) thy, thm),
   159   fn (ctxt, thm) => (map_local_hints (f (g thm)) ctxt, thm));
   160 
   161 fun mk_attr (add1, add2) (del1, del2) =
   162  (Attrib.add_del_args add1 del1, Attrib.add_del_args add2 del2);
   163 
   164 in
   165 
   166 val (simp_add_global, simp_add_local) = global_local map_simps Drule.add_rule;
   167 val (simp_del_global, simp_del_local) = global_local map_simps Drule.del_rule;
   168 val (cong_add_global, cong_add_local) = global_local map_congs add_cong;
   169 val (cong_del_global, cong_del_local) = global_local map_congs del_cong;
   170 val (wf_add_global, wf_add_local) = global_local map_wfs Drule.add_rule;
   171 val (wf_del_global, wf_del_local) = global_local map_wfs Drule.del_rule;
   172 
   173 val simp_attr = mk_attr (simp_add_global, simp_add_local) (simp_del_global, simp_del_local);
   174 val cong_attr = mk_attr (cong_add_global, cong_add_local) (cong_del_global, cong_del_local);
   175 val wf_attr = mk_attr (wf_add_global, wf_add_local) (wf_del_global, wf_del_local);
   176 
   177 end;
   178 
   179 
   180 (* modifiers *)
   181 
   182 val recdef_simpN = "recdef_simp";
   183 val recdef_congN = "recdef_cong";
   184 val recdef_wfN = "recdef_wf";
   185 
   186 val recdef_modifiers =
   187  [Args.$$$ recdef_simpN -- Args.colon >> K ((I, simp_add_local): Method.modifier),
   188   Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (I, simp_add_local),
   189   Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (I, simp_del_local),
   190   Args.$$$ recdef_congN -- Args.colon >> K (I, cong_add_local),
   191   Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (I, cong_add_local),
   192   Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (I, cong_del_local),
   193   Args.$$$ recdef_wfN -- Args.colon >> K (I, wf_add_local),
   194   Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (I, wf_add_local),
   195   Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (I, wf_del_local)] @
   196   Clasimp.clasimp_modifiers;
   197 
   198 
   199 
   200 (** prepare_hints(_i) **)
   201 
   202 fun prepare_hints thy opt_src =
   203   let
   204     val ctxt0 = ProofContext.init thy;
   205     val ctxt =
   206       (case opt_src of
   207         NONE => ctxt0
   208       | SOME src => Method.only_sectioned_args recdef_modifiers I src ctxt0);
   209     val {simps, congs, wfs} = get_local_hints ctxt;
   210     val cs = local_claset_of ctxt;
   211     val ss = local_simpset_of ctxt addsimps simps;
   212   in (cs, ss, map #2 congs, wfs) end;
   213 
   214 fun prepare_hints_i thy () =
   215   let
   216     val ctxt0 = ProofContext.init thy;
   217     val {simps, congs, wfs} = get_global_hints thy;
   218   in (local_claset_of ctxt0, local_simpset_of ctxt0 addsimps simps, map #2 congs, wfs) end;
   219 
   220 
   221 
   222 (** add_recdef(_i) **)
   223 
   224 fun requires_recdef thy = Theory.requires thy "Recdef" "recursive functions";
   225 
   226 fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy =
   227   let
   228     val _ = requires_recdef thy;
   229 
   230     val name = Sign.intern_const thy raw_name;
   231     val bname = Sign.base_name name;
   232     val _ = message ("Defining recursive function " ^ quote name ^ " ...");
   233 
   234     val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs);
   235     val eq_atts = map (map (prep_att thy)) raw_eq_atts;
   236 
   237     val (cs, ss, congs, wfs) = prep_hints thy hints;
   238     (*We must remove imp_cong to prevent looping when the induction rule
   239       is simplified. Many induction rules have nested implications that would
   240       give rise to looping conditional rewriting.*)
   241     val (thy, {rules = rules_idx, induct, tcs}) =
   242         tfl_fn not_permissive thy cs (ss delcongs [imp_cong])
   243                congs wfs name R eqs;
   244     val rules = map (map #1) (Library.partition_eq (Library.eq_snd (op =)) rules_idx);
   245     val simp_att = if null tcs then
   246       [Simplifier.simp_add_global, RecfunCodegen.add NONE] else [];
   247 
   248     val (thy, (simps' :: rules', [induct'])) =
   249       thy
   250       |> Theory.add_path bname
   251       |> PureThy.add_thmss ((("simps", List.concat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
   252       |>>> PureThy.add_thms [(("induct", induct), [])];
   253     val result = {simps = simps', rules = rules', induct = induct', tcs = tcs};
   254     val thy =
   255       thy
   256       |> put_recdef name result
   257       |> Theory.parent_path;
   258   in (thy, result) end;
   259 
   260 val add_recdef = gen_add_recdef Tfl.define Attrib.global_attribute prepare_hints;
   261 fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w ();
   262 
   263 
   264 
   265 (** defer_recdef(_i) **)
   266 
   267 fun gen_defer_recdef tfl_fn app_thms raw_name eqs raw_congs thy =
   268   let
   269     val name = Sign.intern_const thy raw_name;
   270     val bname = Sign.base_name name;
   271 
   272     val _ = requires_recdef thy;
   273     val _ = message ("Deferred recursive function " ^ quote name ^ " ...");
   274 
   275     val (thy1, congs) = thy |> app_thms raw_congs;
   276     val (thy2, induct_rules) = tfl_fn thy1 congs name eqs;
   277     val (thy3, [induct_rules']) =
   278       thy2
   279       |> Theory.add_path bname
   280       |> PureThy.add_thms [(("induct_rules", induct_rules), [])]
   281       |>> Theory.parent_path;
   282   in (thy3, {induct_rules = induct_rules'}) end;
   283 
   284 val defer_recdef = gen_defer_recdef Tfl.defer IsarThy.apply_theorems;
   285 val defer_recdef_i = gen_defer_recdef Tfl.defer_i IsarThy.apply_theorems_i;
   286 
   287 
   288 
   289 (** recdef_tc(_i) **)
   290 
   291 fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i thy =
   292   let
   293     val name = prep_name thy raw_name;
   294     val atts = map (prep_att thy) raw_atts;
   295     val tcs =
   296       (case get_recdef thy name of
   297         NONE => error ("No recdef definition of constant: " ^ quote name)
   298       | SOME {tcs, ...} => tcs);
   299     val i = getOpt (opt_i, 1);
   300     val tc = List.nth (tcs, i - 1) handle Subscript =>
   301       error ("No termination condition #" ^ string_of_int i ^
   302         " in recdef definition of " ^ quote name);
   303   in IsarThy.theorem_i Drule.internalK (bname, atts) (HOLogic.mk_Trueprop tc, ([], [])) thy end;
   304 
   305 val recdef_tc = gen_recdef_tc Attrib.global_attribute Sign.intern_const;
   306 val recdef_tc_i = gen_recdef_tc (K I) (K I);
   307 
   308 
   309 
   310 (** package setup **)
   311 
   312 (* setup theory *)
   313 
   314 val setup =
   315  [GlobalRecdefData.init, LocalRecdefData.init,
   316   Attrib.add_attributes
   317    [(recdef_simpN, simp_attr, "declaration of recdef simp rule"),
   318     (recdef_congN, cong_attr, "declaration of recdef cong rule"),
   319     (recdef_wfN, wf_attr, "declaration of recdef wf rule")]];
   320 
   321 
   322 (* outer syntax *)
   323 
   324 local structure P = OuterParse and K = OuterKeyword in
   325 
   326 val hints =
   327   P.$$$ "(" |-- P.!!! (P.position (P.$$$ "hints" -- P.arguments) --| P.$$$ ")") >> Args.src;
   328 
   329 val recdef_decl =
   330   Scan.optional (P.$$$ "(" -- P.!!! (P.$$$ "permissive" -- P.$$$ ")") >> K false) true --
   331   P.name -- P.term -- Scan.repeat1 (P.opt_thm_name ":" -- P.prop) -- Scan.option hints
   332   >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map P.triple_swap eqs) src);
   333 
   334 val recdefP =
   335   OuterSyntax.command "recdef" "define general recursive functions (TFL)" K.thy_decl
   336     (recdef_decl >> Toplevel.theory);
   337 
   338 
   339 val defer_recdef_decl =
   340   P.name -- Scan.repeat1 P.prop --
   341   Scan.optional (P.$$$ "(" |-- P.$$$ "congs" |-- P.!!! (P.xthms1 --| P.$$$ ")")) []
   342   >> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs);
   343 
   344 val defer_recdefP =
   345   OuterSyntax.command "defer_recdef" "defer general recursive functions (TFL)" K.thy_decl
   346     (defer_recdef_decl >> Toplevel.theory);
   347 
   348 val recdef_tcP =
   349   OuterSyntax.command "recdef_tc" "recommence proof of termination condition (TFL)" K.thy_goal
   350     (P.opt_thm_name ":" -- P.xname -- Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")")
   351       >> (fn ((thm_name, name), i) =>
   352         Toplevel.print o Toplevel.theory_to_proof (recdef_tc thm_name name i)));
   353 
   354 
   355 val _ = OuterSyntax.add_keywords ["permissive", "congs", "hints"];
   356 val _ = OuterSyntax.add_parsers [recdefP, defer_recdefP, recdef_tcP];
   357 
   358 end;
   359 
   360 end;