1 (* Title: HOL/Tools/recdef_package.ML
3 Author: Markus Wenzel, TU Muenchen
5 Wrapper module for Konrad Slind's TFL package.
8 signature RECDEF_PACKAGE =
10 val quiet_mode: bool ref
11 val get_recdef: theory -> string
12 -> {simps: thm list, rules: thm list list, induct: thm, tcs: term list} option
13 val get_hints: Proof.context -> {simps: thm list, congs: (string * thm) list, wfs: thm list}
14 val simp_add: attribute
15 val simp_del: attribute
16 val cong_add: attribute
17 val cong_del: attribute
20 val add_recdef: bool -> xstring -> string -> ((bstring * string) * Attrib.src list) list ->
21 Attrib.src option -> theory -> theory
22 * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
23 val add_recdef_i: bool -> xstring -> term -> ((bstring * term) * attribute list) list ->
24 theory -> theory * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
25 val defer_recdef: xstring -> string list -> (thmref * Attrib.src list) list
26 -> theory -> theory * {induct_rules: thm}
27 val defer_recdef_i: xstring -> term list -> (thm list * attribute list) list
28 -> theory -> theory * {induct_rules: thm}
29 val recdef_tc: bstring * Attrib.src list -> xstring -> int option -> local_theory -> Proof.state
30 val recdef_tc_i: bstring * Attrib.src list -> string -> int option -> local_theory -> Proof.state
31 val setup: theory -> theory
34 structure RecdefPackage: RECDEF_PACKAGE =
37 val quiet_mode = Tfl.quiet_mode;
38 val message = Tfl.message;
45 type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
47 fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints;
48 fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs));
50 fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs));
51 fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs));
52 fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs));
54 fun pretty_hints ({simps, congs, wfs}: hints) =
55 [Pretty.big_list "recdef simp hints:" (map Display.pretty_thm simps),
56 Pretty.big_list "recdef cong hints:" (map Display.pretty_thm (map snd congs)),
57 Pretty.big_list "recdef wf hints:" (map Display.pretty_thm wfs)];
60 (* congruence rules *)
65 fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of;
67 fun prep_cong raw_thm =
68 let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end;
72 fun add_cong raw_thm congs =
74 val (c, thm) = prep_cong raw_thm;
75 val _ = if AList.defined (op =) congs c
76 then warning ("Overwriting recdef congruence rule for " ^ quote c)
78 in AList.update (op =) (c, thm) congs end;
80 fun del_cong raw_thm congs =
82 val (c, thm) = prep_cong raw_thm;
83 val _ = if AList.defined (op =) congs c
85 else warning ("No recdef congruence rule for " ^ quote c);
86 in AList.delete (op =) c congs end;
92 (** global and local recdef data **)
96 type recdef_info = {simps: thm list, rules: thm list list, induct: thm, tcs: term list};
98 structure GlobalRecdefData = TheoryDataFun
100 type T = recdef_info Symtab.table * hints;
101 val empty = (Symtab.empty, mk_hints ([], [], [])): T;
105 ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
106 (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T =
107 (Symtab.merge (K true) (tab1, tab2),
108 mk_hints (Drule.merge_rules (simps1, simps2),
109 AList.merge (op =) Thm.eq_thm (congs1, congs2),
110 Drule.merge_rules (wfs1, wfs2)));
113 val get_recdef = Symtab.lookup o #1 o GlobalRecdefData.get;
115 fun put_recdef name info thy =
117 val (tab, hints) = GlobalRecdefData.get thy;
118 val tab' = Symtab.update_new (name, info) tab
119 handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name);
120 in GlobalRecdefData.put (tab', hints) thy end;
122 val get_global_hints = #2 o GlobalRecdefData.get;
127 structure LocalRecdefData = ProofDataFun
130 val init = get_global_hints;
133 val get_hints = LocalRecdefData.get;
134 fun map_hints f = Context.mapping (GlobalRecdefData.map (apsnd f)) (LocalRecdefData.map f);
139 fun attrib f = Thm.declaration_attribute (map_hints o f);
141 val simp_add = attrib (map_simps o Drule.add_rule);
142 val simp_del = attrib (map_simps o Drule.del_rule);
143 val cong_add = attrib (map_congs o add_cong);
144 val cong_del = attrib (map_congs o del_cong);
145 val wf_add = attrib (map_wfs o Drule.add_rule);
146 val wf_del = attrib (map_wfs o Drule.del_rule);
151 val recdef_simpN = "recdef_simp";
152 val recdef_congN = "recdef_cong";
153 val recdef_wfN = "recdef_wf";
155 val recdef_modifiers =
156 [Args.$$$ recdef_simpN -- Args.colon >> K ((I, simp_add): Method.modifier),
157 Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (I, simp_add),
158 Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (I, simp_del),
159 Args.$$$ recdef_congN -- Args.colon >> K (I, cong_add),
160 Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (I, cong_add),
161 Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (I, cong_del),
162 Args.$$$ recdef_wfN -- Args.colon >> K (I, wf_add),
163 Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (I, wf_add),
164 Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (I, wf_del)] @
165 Clasimp.clasimp_modifiers;
169 (** prepare_hints(_i) **)
171 fun prepare_hints thy opt_src =
173 val ctxt0 = ProofContext.init thy;
177 | SOME src => Method.only_sectioned_args recdef_modifiers I src ctxt0);
178 val {simps, congs, wfs} = get_hints ctxt;
179 val cs = local_claset_of ctxt;
180 val ss = local_simpset_of ctxt addsimps simps;
181 in (cs, ss, rev (map snd congs), wfs) end;
183 fun prepare_hints_i thy () =
185 val ctxt0 = ProofContext.init thy;
186 val {simps, congs, wfs} = get_global_hints thy;
187 in (local_claset_of ctxt0, local_simpset_of ctxt0 addsimps simps, rev (map snd congs), wfs) end;
191 (** add_recdef(_i) **)
193 fun requires_recdef thy = Theory.requires thy "Recdef" "recursive functions";
195 fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy =
197 val _ = requires_recdef thy;
199 val name = Sign.intern_const thy raw_name;
200 val bname = Sign.base_name name;
201 val _ = message ("Defining recursive function " ^ quote name ^ " ...");
203 val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs);
204 val eq_atts = map (map (prep_att thy)) raw_eq_atts;
206 val (cs, ss, congs, wfs) = prep_hints thy hints;
207 (*We must remove imp_cong to prevent looping when the induction rule
208 is simplified. Many induction rules have nested implications that would
209 give rise to looping conditional rewriting.*)
210 val (thy, {rules = rules_idx, induct, tcs}) =
211 tfl_fn not_permissive thy cs (ss delcongs [imp_cong])
212 congs wfs name R eqs;
213 val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx);
214 val simp_att = if null tcs then [Simplifier.simp_add, RecfunCodegen.add NONE] else [];
216 val ((simps' :: rules', [induct']), thy) =
218 |> Theory.add_path bname
220 ((("simps", List.concat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
221 ||>> PureThy.add_thms [(("induct", induct), [])];
222 val result = {simps = simps', rules = rules', induct = induct', tcs = tcs};
225 |> put_recdef name result
226 |> Theory.parent_path;
227 in (thy, result) end;
229 val add_recdef = gen_add_recdef Tfl.define Attrib.attribute prepare_hints;
230 fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w ();
234 (** defer_recdef(_i) **)
236 fun gen_defer_recdef tfl_fn app_thms raw_name eqs raw_congs thy =
238 val name = Sign.intern_const thy raw_name;
239 val bname = Sign.base_name name;
241 val _ = requires_recdef thy;
242 val _ = message ("Deferred recursive function " ^ quote name ^ " ...");
244 val (congs, thy1) = thy |> app_thms raw_congs;
245 val (thy2, induct_rules) = tfl_fn thy1 congs name eqs;
246 val ([induct_rules'], thy3) =
248 |> Theory.add_path bname
249 |> PureThy.add_thms [(("induct_rules", induct_rules), [])]
250 ||> Theory.parent_path;
251 in (thy3, {induct_rules = induct_rules'}) end;
253 val defer_recdef = gen_defer_recdef Tfl.defer IsarCmd.apply_theorems;
254 val defer_recdef_i = gen_defer_recdef Tfl.defer_i IsarCmd.apply_theorems_i;
258 (** recdef_tc(_i) **)
260 fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i lthy =
262 val thy = ProofContext.theory_of lthy;
263 val name = prep_name thy raw_name;
264 val atts = map (prep_att thy) raw_atts;
266 (case get_recdef thy name of
267 NONE => error ("No recdef definition of constant: " ^ quote name)
268 | SOME {tcs, ...} => tcs);
269 val i = the_default 1 opt_i;
270 val tc = nth tcs (i - 1) handle Subscript =>
271 error ("No termination condition #" ^ string_of_int i ^
272 " in recdef definition of " ^ quote name);
274 Specification.theorem_i Thm.internalK NONE (K I) (bname, atts)
275 [] (Element.Shows [(("", []), [(HOLogic.mk_Trueprop tc, [])])]) lthy
278 val recdef_tc = gen_recdef_tc Attrib.intern_src Sign.intern_const;
279 val recdef_tc_i = gen_recdef_tc (K I) (K I);
283 (** package setup **)
288 Attrib.add_attributes
289 [(recdef_simpN, Attrib.add_del_args simp_add simp_del, "declaration of recdef simp rule"),
290 (recdef_congN, Attrib.add_del_args cong_add cong_del, "declaration of recdef cong rule"),
291 (recdef_wfN, Attrib.add_del_args wf_add wf_del, "declaration of recdef wf rule")];
296 local structure P = OuterParse and K = OuterKeyword in
299 P.$$$ "(" |-- P.!!! (P.position (P.$$$ "hints" -- P.arguments) --| P.$$$ ")") >> Args.src;
302 Scan.optional (P.$$$ "(" -- P.!!! (P.$$$ "permissive" -- P.$$$ ")") >> K false) true --
303 P.name -- P.term -- Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop) -- Scan.option hints
304 >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map P.triple_swap eqs) src);
307 OuterSyntax.command "recdef" "define general recursive functions (TFL)" K.thy_decl
308 (recdef_decl >> Toplevel.theory);
311 val defer_recdef_decl =
312 P.name -- Scan.repeat1 P.prop --
313 Scan.optional (P.$$$ "(" |-- P.$$$ "congs" |-- P.!!! (SpecParse.xthms1 --| P.$$$ ")")) []
314 >> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs);
317 OuterSyntax.command "defer_recdef" "defer general recursive functions (TFL)" K.thy_decl
318 (defer_recdef_decl >> Toplevel.theory);
321 OuterSyntax.command "recdef_tc" "recommence proof of termination condition (TFL)" K.thy_goal
323 SpecParse.opt_thm_name ":" -- P.xname -- Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")")
324 >> (fn (((loc, thm_name), name), i) =>
325 Toplevel.print o Toplevel.local_theory_to_proof loc (recdef_tc thm_name name i)));
328 val _ = OuterSyntax.add_keywords ["permissive", "congs", "hints"];
329 val _ = OuterSyntax.add_parsers [recdefP, defer_recdefP, recdef_tcP];