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