primitive definitions of bot/top/inf/sup for bool and fun are named with canonical suffix `_def` rather than `_eq`
1 (* Title: HOL/Tools/recdef.ML
2 Author: Markus Wenzel, TU Muenchen
4 Wrapper module for Konrad Slind's TFL package.
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
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
33 structure Recdef: RECDEF =
41 type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
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));
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));
51 (* congruence rules *)
56 fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of;
58 fun prep_cong raw_thm =
59 let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end;
63 fun add_cong raw_thm congs =
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)
69 in AList.update (op =) (c, thm) congs end;
71 fun del_cong raw_thm congs =
73 val (c, thm) = prep_cong raw_thm;
74 val _ = if AList.defined (op =) congs c
76 else warning ("No recdef congruence rule for " ^ quote c);
77 in AList.delete (op =) c congs end;
83 (** global and local recdef data **)
87 type recdef_info = {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list};
89 structure GlobalRecdefData = Theory_Data
91 type T = recdef_info Symtab.table * hints;
92 val empty = (Symtab.empty, mk_hints ([], [], [])): T;
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)));
103 val get_recdef = Symtab.lookup o #1 o GlobalRecdefData.get;
105 fun put_recdef name info thy =
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;
112 val get_global_hints = #2 o GlobalRecdefData.get;
117 structure LocalRecdefData = Proof_Data
120 val init = get_global_hints;
123 val get_hints = LocalRecdefData.get;
124 fun map_hints f = Context.mapping (GlobalRecdefData.map (apsnd f)) (LocalRecdefData.map f);
129 fun attrib f = Thm.declaration_attribute (map_hints o f);
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);
141 val recdef_simpN = "recdef_simp";
142 val recdef_congN = "recdef_cong";
143 val recdef_wfN = "recdef_wf";
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;
159 (** prepare_hints(_i) **)
161 fun prepare_hints thy opt_src =
163 val ctxt0 = ProofContext.init_global thy;
167 | SOME src => #2 (Method.syntax (Method.sections recdef_modifiers) src ctxt0));
168 val {simps, congs, wfs} = get_hints ctxt;
169 val cs = claset_of ctxt;
170 val ss = simpset_of ctxt addsimps simps;
171 in (cs, ss, rev (map snd congs), wfs) end;
173 fun prepare_hints_i thy () =
175 val ctxt0 = ProofContext.init_global thy;
176 val {simps, congs, wfs} = get_global_hints thy;
177 in (claset_of ctxt0, simpset_of ctxt0 addsimps simps, rev (map snd congs), wfs) end;
181 (** add_recdef(_i) **)
183 fun requires_recdef thy = Theory.requires thy "Recdef" "recursive functions";
185 fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy =
187 val _ = legacy_feature "Old 'recdef' command -- use 'fun' or 'function' instead";
188 val _ = requires_recdef thy;
190 val name = Sign.intern_const thy raw_name;
191 val bname = Long_Name.base_name name;
192 val _ = writeln ("Defining recursive function " ^ quote name ^ " ...");
194 val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs);
195 val eq_atts = map (map (prep_att thy)) raw_eq_atts;
197 val (cs, ss, congs, wfs) = prep_hints thy hints;
198 (*We must remove imp_cong to prevent looping when the induction rule
199 is simplified. Many induction rules have nested implications that would
200 give rise to looping conditional rewriting.*)
201 val (thy, {lhs, rules = rules_idx, induct, tcs}) =
202 tfl_fn not_permissive thy cs (ss delcongs [imp_cong])
203 congs wfs name R eqs;
204 val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx);
206 if null tcs then [Simplifier.simp_add, Nitpick_Simps.add, Code.add_default_eqn_attribute]
208 val ((simps' :: rules', [induct']), thy) =
210 |> Sign.add_path bname
211 |> Global_Theory.add_thmss
212 (((Binding.name "simps", flat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
213 ||>> Global_Theory.add_thms [((Binding.name "induct", induct), [])]
214 ||> Spec_Rules.add_global Spec_Rules.Equational ([lhs], flat rules);
215 val result = {lhs = lhs, simps = simps', rules = rules', induct = induct', tcs = tcs};
218 |> put_recdef name result
220 in (thy, result) end;
222 val add_recdef = gen_add_recdef Tfl.define Attrib.attribute prepare_hints;
223 fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w ();
227 (** defer_recdef(_i) **)
229 fun gen_defer_recdef tfl_fn eval_thms raw_name eqs raw_congs thy =
231 val name = Sign.intern_const thy raw_name;
232 val bname = Long_Name.base_name name;
234 val _ = requires_recdef thy;
235 val _ = writeln ("Deferred recursive function " ^ quote name ^ " ...");
237 val congs = eval_thms (ProofContext.init_global thy) raw_congs;
238 val (thy2, induct_rules) = tfl_fn thy congs name eqs;
239 val ([induct_rules'], thy3) =
241 |> Sign.add_path bname
242 |> Global_Theory.add_thms [((Binding.name "induct_rules", induct_rules), [])]
243 ||> Sign.parent_path;
244 in (thy3, {induct_rules = induct_rules'}) end;
246 val defer_recdef = gen_defer_recdef Tfl.defer Attrib.eval_thms;
247 val defer_recdef_i = gen_defer_recdef Tfl.defer_i (K I);
251 (** recdef_tc(_i) **)
253 fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i int lthy =
255 val thy = ProofContext.theory_of lthy;
256 val name = prep_name thy raw_name;
257 val atts = map (prep_att thy) raw_atts;
259 (case get_recdef thy name of
260 NONE => error ("No recdef definition of constant: " ^ quote name)
261 | SOME {tcs, ...} => tcs);
262 val i = the_default 1 opt_i;
263 val tc = nth tcs (i - 1) handle Subscript =>
264 error ("No termination condition #" ^ string_of_int i ^
265 " in recdef definition of " ^ quote name);
267 Specification.theorem "" NONE (K I)
268 (Binding.conceal (Binding.name bname), atts) []
269 (Element.Shows [(Attrib.empty_binding, [(HOLogic.mk_Trueprop tc, [])])]) int lthy
272 val recdef_tc = gen_recdef_tc Attrib.intern_src Sign.intern_const;
273 val recdef_tc_i = gen_recdef_tc (K I) (K I);
277 (** package setup **)
282 Attrib.setup @{binding recdef_simp} (Attrib.add_del simp_add simp_del)
283 "declaration of recdef simp rule" #>
284 Attrib.setup @{binding recdef_cong} (Attrib.add_del cong_add cong_del)
285 "declaration of recdef cong rule" #>
286 Attrib.setup @{binding recdef_wf} (Attrib.add_del wf_add wf_del)
287 "declaration of recdef wf rule";
292 val _ = List.app Keyword.keyword ["permissive", "congs", "hints"];
296 Parse.!!! (Parse.position (Parse.$$$ "hints" -- Args.parse) --| Parse.$$$ ")") >> Args.src;
300 (Parse.$$$ "(" -- Parse.!!! (Parse.$$$ "permissive" -- Parse.$$$ ")") >> K false) true --
301 Parse.name -- Parse.term -- Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)
303 >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map Parse.triple_swap eqs) src);
306 Outer_Syntax.command "recdef" "define general recursive functions (TFL)" Keyword.thy_decl
307 (recdef_decl >> Toplevel.theory);
310 val defer_recdef_decl =
311 Parse.name -- Scan.repeat1 Parse.prop --
313 (Parse.$$$ "(" |-- Parse.$$$ "congs" |-- Parse.!!! (Parse_Spec.xthms1 --| Parse.$$$ ")")) []
314 >> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs);
317 Outer_Syntax.command "defer_recdef" "defer general recursive functions (TFL)" Keyword.thy_decl
318 (defer_recdef_decl >> Toplevel.theory);
321 Outer_Syntax.local_theory_to_proof' "recdef_tc" "recommence proof of termination condition (TFL)"
323 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.xname --
324 Scan.option (Parse.$$$ "(" |-- Parse.nat --| Parse.$$$ ")")
325 >> (fn ((thm_name, name), i) => recdef_tc thm_name name i));