author | nipkow |
Thu, 22 Dec 2005 17:57:09 +0100 | |
changeset 18493 | 343da052b961 |
parent 18418 | bf448d999b7e |
child 18688 | abf0f018b5ec |
permissions | -rw-r--r-- |
6429 | 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 |
|
6439 | 11 |
val print_recdefs: theory -> unit |
8657 | 12 |
val get_recdef: theory -> string |
13 |
-> {simps: thm list, rules: thm list list, induct: thm, tcs: term list} option |
|
9859 | 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 |
|
15703 | 26 |
val add_recdef: bool -> xstring -> string -> ((bstring * string) * Attrib.src list) list -> |
27 |
Attrib.src option -> theory -> theory |
|
9859 | 28 |
* {simps: thm list, rules: thm list list, induct: thm, tcs: term list} |
11629 | 29 |
val add_recdef_i: bool -> xstring -> term -> ((bstring * term) * theory attribute list) list -> |
9859 | 30 |
theory -> theory * {simps: thm list, rules: thm list list, induct: thm, tcs: term list} |
15703 | 31 |
val defer_recdef: xstring -> string list -> (thmref * Attrib.src list) list |
6557 | 32 |
-> theory -> theory * {induct_rules: thm} |
12708 | 33 |
val defer_recdef_i: xstring -> term list -> (thm list * theory attribute list) list |
6557 | 34 |
-> theory -> theory * {induct_rules: thm} |
17336 | 35 |
val recdef_tc: bstring * Attrib.src list -> xstring -> int option -> theory -> Proof.state |
10775 | 36 |
val recdef_tc_i: bstring * theory attribute list -> string -> int option |
17336 | 37 |
-> theory -> Proof.state |
6439 | 38 |
val setup: (theory -> theory) list |
6429 | 39 |
end; |
40 |
||
41 |
structure RecdefPackage: RECDEF_PACKAGE = |
|
42 |
struct |
|
43 |
||
9859 | 44 |
|
6429 | 45 |
val quiet_mode = Tfl.quiet_mode; |
46 |
val message = Tfl.message; |
|
47 |
||
48 |
||
9859 | 49 |
(** recdef hints **) |
6439 | 50 |
|
9859 | 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; |
|
6439 | 74 |
|
9859 | 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 |
|
17593 | 82 |
in overwrite_warn (congs, (c, thm)) ("Overwriting recdef congruence rule for " ^ quote c) end; |
9859 | 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 |
||
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
90 |
val add_congs = foldr (uncurry add_cong); |
9859 | 91 |
|
92 |
end; |
|
93 |
||
94 |
||
95 |
||
96 |
(** global and local recdef data **) |
|
97 |
||
17920 | 98 |
(* theory data *) |
6439 | 99 |
|
8657 | 100 |
type recdef_info = {simps: thm list, rules: thm list list, induct: thm, tcs: term list}; |
6439 | 101 |
|
16458 | 102 |
structure GlobalRecdefData = TheoryDataFun |
103 |
(struct |
|
6439 | 104 |
val name = "HOL/recdef"; |
9859 | 105 |
type T = recdef_info Symtab.table * hints; |
6439 | 106 |
|
9879 | 107 |
val empty = (Symtab.empty, mk_hints ([], [], [])): T; |
6557 | 108 |
val copy = I; |
16458 | 109 |
val extend = I; |
110 |
fun merge _ |
|
9859 | 111 |
((tab1, {simps = simps1, congs = congs1, wfs = wfs1}), |
16458 | 112 |
(tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T = |
9859 | 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))); |
|
6439 | 117 |
|
16458 | 118 |
fun print thy (tab, hints) = |
119 |
(Pretty.strs ("recdefs:" :: map #1 (NameSpace.extern_table (Sign.const_space thy, tab))) :: |
|
16364 | 120 |
pretty_hints hints) |> Pretty.chunks |> Pretty.writeln; |
16458 | 121 |
end); |
6439 | 122 |
|
9859 | 123 |
val print_recdefs = GlobalRecdefData.print; |
6429 | 124 |
|
6439 | 125 |
|
17412 | 126 |
val get_recdef = Symtab.lookup o #1 o GlobalRecdefData.get; |
6439 | 127 |
|
128 |
fun put_recdef name info thy = |
|
6429 | 129 |
let |
9859 | 130 |
val (tab, hints) = GlobalRecdefData.get thy; |
17412 | 131 |
val tab' = Symtab.update_new (name, info) tab |
6439 | 132 |
handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name); |
9859 | 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 |
||
17920 | 139 |
(* proof data *) |
9859 | 140 |
|
16458 | 141 |
structure LocalRecdefData = ProofDataFun |
142 |
(struct |
|
9859 | 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; |
|
16458 | 147 |
end); |
9859 | 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 |
||
12371 | 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; |
|
9859 | 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; |
|
12371 | 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; |
|
9859 | 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 |
||
9949 | 180 |
(* modifiers *) |
9859 | 181 |
|
9949 | 182 |
val recdef_simpN = "recdef_simp"; |
183 |
val recdef_congN = "recdef_cong"; |
|
184 |
val recdef_wfN = "recdef_wf"; |
|
9859 | 185 |
|
186 |
val recdef_modifiers = |
|
9949 | 187 |
[Args.$$$ recdef_simpN -- Args.colon >> K ((I, simp_add_local): Method.modifier), |
10032 | 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), |
|
9949 | 190 |
Args.$$$ recdef_congN -- Args.colon >> K (I, cong_add_local), |
10032 | 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), |
|
9949 | 193 |
Args.$$$ recdef_wfN -- Args.colon >> K (I, wf_add_local), |
10032 | 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)] @ |
|
9949 | 196 |
Clasimp.clasimp_modifiers; |
9859 | 197 |
|
9949 | 198 |
|
9859 | 199 |
|
9949 | 200 |
(** prepare_hints(_i) **) |
9859 | 201 |
|
202 |
fun prepare_hints thy opt_src = |
|
203 |
let |
|
204 |
val ctxt0 = ProofContext.init thy; |
|
205 |
val ctxt = |
|
206 |
(case opt_src of |
|
15531 | 207 |
NONE => ctxt0 |
208 |
| SOME src => Method.only_sectioned_args recdef_modifiers I src ctxt0); |
|
9859 | 209 |
val {simps, congs, wfs} = get_local_hints ctxt; |
15032 | 210 |
val cs = local_claset_of ctxt; |
211 |
val ss = local_simpset_of ctxt addsimps simps; |
|
9859 | 212 |
in (cs, ss, map #2 congs, wfs) end; |
213 |
||
214 |
fun prepare_hints_i thy () = |
|
15032 | 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; |
|
9859 | 219 |
|
6439 | 220 |
|
221 |
||
222 |
(** add_recdef(_i) **) |
|
223 |
||
6557 | 224 |
fun requires_recdef thy = Theory.requires thy "Recdef" "recursive functions"; |
225 |
||
17920 | 226 |
fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy = |
6439 | 227 |
let |
9859 | 228 |
val _ = requires_recdef thy; |
229 |
||
16458 | 230 |
val name = Sign.intern_const thy raw_name; |
6439 | 231 |
val bname = Sign.base_name name; |
6429 | 232 |
val _ = message ("Defining recursive function " ^ quote name ^ " ..."); |
233 |
||
8657 | 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 |
||
9859 | 237 |
val (cs, ss, congs, wfs) = prep_hints thy hints; |
14241
dfae7eb2830c
Prevent recdef from looping when the inductio rule is simplified
paulson
parents:
12876
diff
changeset
|
238 |
(*We must remove imp_cong to prevent looping when the induction rule |
dfae7eb2830c
Prevent recdef from looping when the inductio rule is simplified
paulson
parents:
12876
diff
changeset
|
239 |
is simplified. Many induction rules have nested implications that would |
dfae7eb2830c
Prevent recdef from looping when the inductio rule is simplified
paulson
parents:
12876
diff
changeset
|
240 |
give rise to looping conditional rewriting.*) |
dfae7eb2830c
Prevent recdef from looping when the inductio rule is simplified
paulson
parents:
12876
diff
changeset
|
241 |
val (thy, {rules = rules_idx, induct, tcs}) = |
17920 | 242 |
tfl_fn not_permissive thy cs (ss delcongs [imp_cong]) |
14241
dfae7eb2830c
Prevent recdef from looping when the inductio rule is simplified
paulson
parents:
12876
diff
changeset
|
243 |
congs wfs name R eqs; |
17496 | 244 |
val rules = map (map #1) (Library.partition_eq (Library.eq_snd (op =)) rules_idx); |
12448
473cb9f9e237
Recursive equations to be used for code generation are now registered
berghofe
parents:
12371
diff
changeset
|
245 |
val simp_att = if null tcs then |
16646 | 246 |
[Simplifier.simp_add_global, RecfunCodegen.add NONE] else []; |
8657 | 247 |
|
18377 | 248 |
val ((simps' :: rules', [induct']), thy) = |
7798
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7262
diff
changeset
|
249 |
thy |
6439 | 250 |
|> Theory.add_path bname |
15570 | 251 |
|> PureThy.add_thmss ((("simps", List.concat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts)) |
18377 | 252 |
||>> PureThy.add_thms [(("induct", induct), [])]; |
8657 | 253 |
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
|
254 |
val thy = |
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7262
diff
changeset
|
255 |
thy |
6439 | 256 |
|> put_recdef name result |
6429 | 257 |
|> Theory.parent_path; |
7798
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7262
diff
changeset
|
258 |
in (thy, result) end; |
6429 | 259 |
|
9859 | 260 |
val add_recdef = gen_add_recdef Tfl.define Attrib.global_attribute prepare_hints; |
11629 | 261 |
fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w (); |
9859 | 262 |
|
263 |
||
6557 | 264 |
|
265 |
(** defer_recdef(_i) **) |
|
266 |
||
267 |
fun gen_defer_recdef tfl_fn app_thms raw_name eqs raw_congs thy = |
|
268 |
let |
|
16458 | 269 |
val name = Sign.intern_const thy raw_name; |
6557 | 270 |
val bname = Sign.base_name name; |
271 |
||
272 |
val _ = requires_recdef thy; |
|
273 |
val _ = message ("Deferred recursive function " ^ quote name ^ " ..."); |
|
274 |
||
18418
bf448d999b7e
re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents:
18377
diff
changeset
|
275 |
val (congs, thy1) = thy |> app_thms raw_congs; |
9859 | 276 |
val (thy2, induct_rules) = tfl_fn thy1 congs name eqs; |
18377 | 277 |
val ([induct_rules'], thy3) = |
6557 | 278 |
thy2 |
279 |
|> Theory.add_path bname |
|
280 |
|> PureThy.add_thms [(("induct_rules", induct_rules), [])] |
|
18377 | 281 |
||> Theory.parent_path; |
8430 | 282 |
in (thy3, {induct_rules = induct_rules'}) end; |
6557 | 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 |
||
10775 | 289 |
(** recdef_tc(_i) **) |
290 |
||
17336 | 291 |
fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i thy = |
10775 | 292 |
let |
16458 | 293 |
val name = prep_name thy raw_name; |
10775 | 294 |
val atts = map (prep_att thy) raw_atts; |
295 |
val tcs = |
|
296 |
(case get_recdef thy name of |
|
15531 | 297 |
NONE => error ("No recdef definition of constant: " ^ quote name) |
298 |
| SOME {tcs, ...} => tcs); |
|
15570 | 299 |
val i = getOpt (opt_i, 1); |
300 |
val tc = List.nth (tcs, i - 1) handle Subscript => |
|
10775 | 301 |
error ("No termination condition #" ^ string_of_int i ^ |
302 |
" in recdef definition of " ^ quote name); |
|
17336 | 303 |
in IsarThy.theorem_i Drule.internalK (bname, atts) (HOLogic.mk_Trueprop tc, ([], [])) thy end; |
10775 | 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 |
||
6439 | 310 |
(** package setup **) |
311 |
||
312 |
(* setup theory *) |
|
313 |
||
9859 | 314 |
val setup = |
315 |
[GlobalRecdefData.init, LocalRecdefData.init, |
|
316 |
Attrib.add_attributes |
|
9949 | 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")]]; |
|
6439 | 320 |
|
321 |
||
6429 | 322 |
(* outer syntax *) |
323 |
||
17057 | 324 |
local structure P = OuterParse and K = OuterKeyword in |
6429 | 325 |
|
9859 | 326 |
val hints = |
327 |
P.$$$ "(" |-- P.!!! (P.position (P.$$$ "hints" -- P.arguments) --| P.$$$ ")") >> Args.src; |
|
328 |
||
6429 | 329 |
val recdef_decl = |
11629 | 330 |
Scan.optional (P.$$$ "(" -- P.!!! (P.$$$ "permissive" -- P.$$$ ")") >> K false) true -- |
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
12755
diff
changeset
|
331 |
P.name -- P.term -- Scan.repeat1 (P.opt_thm_name ":" -- P.prop) -- Scan.option hints |
11629 | 332 |
>> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map P.triple_swap eqs) src); |
6429 | 333 |
|
334 |
val recdefP = |
|
6723 | 335 |
OuterSyntax.command "recdef" "define general recursive functions (TFL)" K.thy_decl |
6429 | 336 |
(recdef_decl >> Toplevel.theory); |
337 |
||
6557 | 338 |
|
339 |
val defer_recdef_decl = |
|
8657 | 340 |
P.name -- Scan.repeat1 P.prop -- |
6723 | 341 |
Scan.optional (P.$$$ "(" |-- P.$$$ "congs" |-- P.!!! (P.xthms1 --| P.$$$ ")")) [] |
6557 | 342 |
>> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs); |
343 |
||
344 |
val defer_recdefP = |
|
6723 | 345 |
OuterSyntax.command "defer_recdef" "defer general recursive functions (TFL)" K.thy_decl |
6557 | 346 |
(defer_recdef_decl >> Toplevel.theory); |
347 |
||
10775 | 348 |
val recdef_tcP = |
349 |
OuterSyntax.command "recdef_tc" "recommence proof of termination condition (TFL)" K.thy_goal |
|
17336 | 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))); |
|
10775 | 353 |
|
354 |
||
11629 | 355 |
val _ = OuterSyntax.add_keywords ["permissive", "congs", "hints"]; |
10775 | 356 |
val _ = OuterSyntax.add_parsers [recdefP, defer_recdefP, recdef_tcP]; |
6429 | 357 |
|
358 |
end; |
|
359 |
||
360 |
end; |