src/HOL/Tools/recdef.ML
author wenzelm
Thu, 15 Mar 2012 20:07:00 +0100
changeset 47823 94aa7b81bcf6
parent 47821 b8c7eb0c2f89
child 47836 5c6955f487e5
permissions -rw-r--r--
prefer formally checked @{keyword} parser;
     1 (*  Title:      HOL/Tools/recdef.ML
     2     Author:     Markus Wenzel, TU Muenchen
     3 
     4 Wrapper module for Konrad Slind's TFL package.
     5 *)
     6 
     7 signature RECDEF =
     8 sig
     9   val get_recdef: theory -> string
    10     -> {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list} option
    11   val get_hints: Proof.context -> {simps: thm list, congs: (string * thm) list, wfs: thm list}
    12   val simp_add: attribute
    13   val simp_del: attribute
    14   val cong_add: attribute
    15   val cong_del: attribute
    16   val wf_add: attribute
    17   val wf_del: attribute
    18   val add_recdef: bool -> xstring -> string -> ((binding * string) * Attrib.src list) list ->
    19     Attrib.src option -> theory -> theory
    20       * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
    21   val add_recdef_i: bool -> xstring -> term -> ((binding * term) * attribute list) list ->
    22     theory -> theory * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
    23   val defer_recdef: xstring -> string list -> (Facts.ref * Attrib.src list) list
    24     -> theory -> theory * {induct_rules: thm}
    25   val defer_recdef_i: xstring -> term list -> thm list -> theory -> theory * {induct_rules: thm}
    26   val recdef_tc: bstring * Attrib.src list -> xstring -> int option -> bool ->
    27     local_theory -> Proof.state
    28   val recdef_tc_i: bstring * Attrib.src list -> string -> int option -> bool ->
    29     local_theory -> Proof.state
    30   val setup: theory -> theory
    31 end;
    32 
    33 structure Recdef: RECDEF =
    34 struct
    35 
    36 
    37 (** recdef hints **)
    38 
    39 (* type hints *)
    40 
    41 type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
    42 
    43 fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints;
    44 fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs));
    45 
    46 fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs));
    47 fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs));
    48 fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs));
    49 
    50 
    51 (* congruence rules *)
    52 
    53 local
    54 
    55 val cong_head =
    56   fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of;
    57 
    58 fun prep_cong raw_thm =
    59   let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end;
    60 
    61 in
    62 
    63 fun add_cong raw_thm congs =
    64   let
    65     val (c, thm) = prep_cong raw_thm;
    66     val _ = if AList.defined (op =) congs c
    67       then warning ("Overwriting recdef congruence rule for " ^ quote c)
    68       else ();
    69   in AList.update (op =) (c, thm) congs end;
    70 
    71 fun del_cong raw_thm congs =
    72   let
    73     val (c, thm) = prep_cong raw_thm;
    74     val _ = if AList.defined (op =) congs c
    75       then ()
    76       else warning ("No recdef congruence rule for " ^ quote c);
    77   in AList.delete (op =) c congs end;
    78 
    79 end;
    80 
    81 
    82 
    83 (** global and local recdef data **)
    84 
    85 (* theory data *)
    86 
    87 type recdef_info = {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list};
    88 
    89 structure GlobalRecdefData = Theory_Data
    90 (
    91   type T = recdef_info Symtab.table * hints;
    92   val empty = (Symtab.empty, mk_hints ([], [], [])): T;
    93   val extend = I;
    94   fun merge
    95    ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
    96     (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T =
    97       (Symtab.merge (K true) (tab1, tab2),
    98         mk_hints (Thm.merge_thms (simps1, simps2),
    99           AList.merge (op =) (K true) (congs1, congs2),
   100           Thm.merge_thms (wfs1, wfs2)));
   101 );
   102 
   103 val get_recdef = Symtab.lookup o #1 o GlobalRecdefData.get;
   104 
   105 fun put_recdef name info thy =
   106   let
   107     val (tab, hints) = GlobalRecdefData.get thy;
   108     val tab' = Symtab.update_new (name, info) tab
   109       handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name);
   110   in GlobalRecdefData.put (tab', hints) thy end;
   111 
   112 val get_global_hints = #2 o GlobalRecdefData.get;
   113 
   114 
   115 (* proof data *)
   116 
   117 structure LocalRecdefData = Proof_Data
   118 (
   119   type T = hints;
   120   val init = get_global_hints;
   121 );
   122 
   123 val get_hints = LocalRecdefData.get;
   124 fun map_hints f = Context.mapping (GlobalRecdefData.map (apsnd f)) (LocalRecdefData.map f);
   125 
   126 
   127 (* attributes *)
   128 
   129 fun attrib f = Thm.declaration_attribute (map_hints o f);
   130 
   131 val simp_add = attrib (map_simps o Thm.add_thm);
   132 val simp_del = attrib (map_simps o Thm.del_thm);
   133 val cong_add = attrib (map_congs o add_cong);
   134 val cong_del = attrib (map_congs o del_cong);
   135 val wf_add = attrib (map_wfs o Thm.add_thm);
   136 val wf_del = attrib (map_wfs o Thm.del_thm);
   137 
   138 
   139 (* modifiers *)
   140 
   141 val recdef_simpN = "recdef_simp";
   142 val recdef_congN = "recdef_cong";
   143 val recdef_wfN = "recdef_wf";
   144 
   145 val recdef_modifiers =
   146  [Args.$$$ recdef_simpN -- Args.colon >> K ((I, simp_add): Method.modifier),
   147   Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (I, simp_add),
   148   Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (I, simp_del),
   149   Args.$$$ recdef_congN -- Args.colon >> K (I, cong_add),
   150   Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (I, cong_add),
   151   Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (I, cong_del),
   152   Args.$$$ recdef_wfN -- Args.colon >> K (I, wf_add),
   153   Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (I, wf_add),
   154   Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (I, wf_del)] @
   155   Clasimp.clasimp_modifiers;
   156 
   157 
   158 
   159 (** prepare_hints(_i) **)
   160 
   161 fun prepare_hints thy opt_src =
   162   let
   163     val ctxt0 = Proof_Context.init_global thy;
   164     val ctxt =
   165       (case opt_src of
   166         NONE => ctxt0
   167       | SOME src => #2 (Method.syntax (Method.sections recdef_modifiers) src ctxt0));
   168     val {simps, congs, wfs} = get_hints ctxt;
   169     val ctxt' = ctxt
   170       |> Simplifier.map_simpset (fn ss => ss addsimps simps |> Simplifier.del_cong @{thm imp_cong});
   171   in (ctxt', rev (map snd congs), wfs) end;
   172 
   173 fun prepare_hints_i thy () =
   174   let
   175     val ctxt = Proof_Context.init_global thy;
   176     val {simps, congs, wfs} = get_global_hints thy;
   177     val ctxt' = ctxt
   178       |> Simplifier.map_simpset (fn ss => ss addsimps simps |> Simplifier.del_cong @{thm imp_cong});
   179   in (ctxt', rev (map snd congs), wfs) end;
   180 
   181 
   182 
   183 (** add_recdef(_i) **)
   184 
   185 fun requires_recdef thy = Theory.requires thy "Old_Recdef" "recursive functions";
   186 
   187 fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy =
   188   let
   189     val _ = legacy_feature "Old 'recdef' command -- use 'fun' or 'function' instead";
   190     val _ = requires_recdef thy;
   191 
   192     val name = Sign.intern_const thy raw_name;
   193     val bname = Long_Name.base_name name;
   194     val _ = writeln ("Defining recursive function " ^ quote name ^ " ...");
   195 
   196     val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs);
   197     val eq_atts = map (map (prep_att thy)) raw_eq_atts;
   198 
   199     val (ctxt, congs, wfs) = prep_hints thy hints;
   200     (*We must remove imp_cong to prevent looping when the induction rule
   201       is simplified. Many induction rules have nested implications that would
   202       give rise to looping conditional rewriting.*)
   203     val ({lhs, rules = rules_idx, induct, tcs}, thy) =
   204       tfl_fn not_permissive ctxt congs wfs name R eqs thy;
   205     val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx);
   206     val simp_att =
   207       if null tcs then [Simplifier.simp_add, Nitpick_Simps.add, Code.add_default_eqn_attribute]
   208       else [];
   209     val ((simps' :: rules', [induct']), thy) =
   210       thy
   211       |> Sign.add_path bname
   212       |> Global_Theory.add_thmss
   213         (((Binding.name "simps", flat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
   214       ||>> Global_Theory.add_thms [((Binding.name "induct", induct), [])]
   215       ||> Spec_Rules.add_global Spec_Rules.Equational ([lhs], flat rules);
   216     val result = {lhs = lhs, simps = simps', rules = rules', induct = induct', tcs = tcs};
   217     val thy =
   218       thy
   219       |> put_recdef name result
   220       |> Sign.parent_path;
   221   in (thy, result) end;
   222 
   223 val add_recdef = gen_add_recdef Tfl.define Attrib.attribute prepare_hints;
   224 fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w ();
   225 
   226 
   227 
   228 (** defer_recdef(_i) **)
   229 
   230 fun gen_defer_recdef tfl_fn eval_thms raw_name eqs raw_congs thy =
   231   let
   232     val name = Sign.intern_const thy raw_name;
   233     val bname = Long_Name.base_name name;
   234 
   235     val _ = requires_recdef thy;
   236     val _ = writeln ("Deferred recursive function " ^ quote name ^ " ...");
   237 
   238     val congs = eval_thms (Proof_Context.init_global thy) raw_congs;
   239     val (induct_rules, thy2) = tfl_fn congs name eqs thy;
   240     val ([induct_rules'], thy3) =
   241       thy2
   242       |> Sign.add_path bname
   243       |> Global_Theory.add_thms [((Binding.name "induct_rules", induct_rules), [])]
   244       ||> Sign.parent_path;
   245   in (thy3, {induct_rules = induct_rules'}) end;
   246 
   247 val defer_recdef = gen_defer_recdef Tfl.defer Attrib.eval_thms;
   248 val defer_recdef_i = gen_defer_recdef Tfl.defer_i (K I);
   249 
   250 
   251 
   252 (** recdef_tc(_i) **)
   253 
   254 fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i int lthy =
   255   let
   256     val thy = Proof_Context.theory_of lthy;
   257     val name = prep_name thy raw_name;
   258     val atts = map (prep_att thy) raw_atts;
   259     val tcs =
   260       (case get_recdef thy name of
   261         NONE => error ("No recdef definition of constant: " ^ quote name)
   262       | SOME {tcs, ...} => tcs);
   263     val i = the_default 1 opt_i;
   264     val tc = nth tcs (i - 1) handle General.Subscript =>
   265       error ("No termination condition #" ^ string_of_int i ^
   266         " in recdef definition of " ^ quote name);
   267   in
   268     Specification.theorem "" NONE (K I)
   269       (Binding.conceal (Binding.name bname), atts) []
   270       (Element.Shows [(Attrib.empty_binding, [(HOLogic.mk_Trueprop tc, [])])]) int lthy
   271   end;
   272 
   273 val recdef_tc = gen_recdef_tc Attrib.intern_src Sign.intern_const;
   274 val recdef_tc_i = gen_recdef_tc (K I) (K I);
   275 
   276 
   277 
   278 (** package setup **)
   279 
   280 (* setup theory *)
   281 
   282 val setup =
   283   Attrib.setup @{binding recdef_simp} (Attrib.add_del simp_add simp_del)
   284     "declaration of recdef simp rule" #>
   285   Attrib.setup @{binding recdef_cong} (Attrib.add_del cong_add cong_del)
   286     "declaration of recdef cong rule" #>
   287   Attrib.setup @{binding recdef_wf} (Attrib.add_del wf_add wf_del)
   288     "declaration of recdef wf rule";
   289 
   290 
   291 (* outer syntax *)
   292 
   293 val hints =
   294   @{keyword "("} |--
   295     Parse.!!! (Parse.position (@{keyword "hints"} -- Args.parse) --| @{keyword ")"}) >> Args.src;
   296 
   297 val recdef_decl =
   298   Scan.optional
   299     (@{keyword "("} -- Parse.!!! (@{keyword "permissive"} -- @{keyword ")"}) >> K false) true --
   300   Parse.name -- Parse.term -- Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)
   301     -- Scan.option hints
   302   >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map Parse.triple_swap eqs) src);
   303 
   304 val _ =
   305   Outer_Syntax.command "recdef" "define general recursive functions (TFL)" Keyword.thy_decl
   306     (recdef_decl >> Toplevel.theory);
   307 
   308 
   309 val defer_recdef_decl =
   310   Parse.name -- Scan.repeat1 Parse.prop --
   311   Scan.optional
   312     (@{keyword "("} |-- @{keyword "congs"} |-- Parse.!!! (Parse_Spec.xthms1 --| @{keyword ")"})) []
   313   >> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs);
   314 
   315 val _ =
   316   Outer_Syntax.command "defer_recdef" "defer general recursive functions (TFL)" Keyword.thy_decl
   317     (defer_recdef_decl >> Toplevel.theory);
   318 
   319 val _ =
   320   Outer_Syntax.local_theory_to_proof' "recdef_tc" "recommence proof of termination condition (TFL)"
   321     Keyword.thy_goal
   322     ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.xname --
   323         Scan.option (@{keyword "("} |-- Parse.nat --| @{keyword ")"})
   324       >> (fn ((thm_name, name), i) => recdef_tc thm_name name i));
   325 
   326 end;