| author | haftmann |
| Fri, 03 Nov 2006 14:22:45 +0100 | |
| changeset 21159 | 7f6bdffe3d06 |
| parent 21098 | d0d8a48ae4e6 |
| child 21351 | 1fb804b96d7c |
| 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
|
|
| 20291 | 14 |
val get_hints: Context.generic -> {simps: thm list, congs: (string * thm) list, wfs: thm list}
|
| 18728 | 15 |
val simp_add: attribute |
16 |
val simp_del: attribute |
|
17 |
val cong_add: attribute |
|
18 |
val cong_del: attribute |
|
19 |
val wf_add: attribute |
|
20 |
val wf_del: attribute |
|
| 15703 | 21 |
val add_recdef: bool -> xstring -> string -> ((bstring * string) * Attrib.src list) list -> |
22 |
Attrib.src option -> theory -> theory |
|
| 9859 | 23 |
* {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
|
| 18728 | 24 |
val add_recdef_i: bool -> xstring -> term -> ((bstring * term) * attribute list) list -> |
| 9859 | 25 |
theory -> theory * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
|
| 15703 | 26 |
val defer_recdef: xstring -> string list -> (thmref * Attrib.src list) list |
| 6557 | 27 |
-> theory -> theory * {induct_rules: thm}
|
| 18728 | 28 |
val defer_recdef_i: xstring -> term list -> (thm list * attribute list) list |
| 6557 | 29 |
-> theory -> theory * {induct_rules: thm}
|
| 17336 | 30 |
val recdef_tc: bstring * Attrib.src list -> xstring -> int option -> theory -> Proof.state |
| 18728 | 31 |
val recdef_tc_i: bstring * attribute list -> string -> int option -> theory -> Proof.state |
| 18708 | 32 |
val setup: theory -> theory |
| 6429 | 33 |
end; |
34 |
||
35 |
structure RecdefPackage: RECDEF_PACKAGE = |
|
36 |
struct |
|
37 |
||
38 |
val quiet_mode = Tfl.quiet_mode; |
|
39 |
val message = Tfl.message; |
|
40 |
||
41 |
||
| 9859 | 42 |
(** recdef hints **) |
| 6439 | 43 |
|
| 9859 | 44 |
(* type hints *) |
45 |
||
46 |
type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
|
|
47 |
||
48 |
fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints;
|
|
49 |
fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs));
|
|
50 |
||
51 |
fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs)); |
|
52 |
fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs)); |
|
53 |
fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs)); |
|
54 |
||
55 |
fun pretty_hints ({simps, congs, wfs}: hints) =
|
|
56 |
[Pretty.big_list "recdef simp hints:" (map Display.pretty_thm simps), |
|
|
21098
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
57 |
Pretty.big_list "recdef cong hints:" (map Display.pretty_thm (map snd congs)), |
| 9859 | 58 |
Pretty.big_list "recdef wf hints:" (map Display.pretty_thm wfs)]; |
59 |
||
60 |
||
61 |
(* congruence rules *) |
|
62 |
||
63 |
local |
|
64 |
||
65 |
val cong_head = |
|
66 |
fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of; |
|
| 6439 | 67 |
|
| 9859 | 68 |
fun prep_cong raw_thm = |
69 |
let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end; |
|
70 |
||
71 |
in |
|
72 |
||
73 |
fun add_cong raw_thm congs = |
|
|
21098
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
74 |
let |
|
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
75 |
val (c, thm) = prep_cong raw_thm; |
|
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
76 |
val _ = if AList.defined (op =) congs c |
|
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
77 |
then warning ("Overwriting recdef congruence rule for " ^ quote c)
|
|
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
78 |
else (); |
|
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
79 |
in AList.update (op =) (c, thm) congs end; |
| 9859 | 80 |
|
81 |
fun del_cong raw_thm congs = |
|
82 |
let |
|
83 |
val (c, thm) = prep_cong raw_thm; |
|
|
21098
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
84 |
val _ = if AList.defined (op =) congs c |
|
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
85 |
then () |
|
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
86 |
else warning ("No recdef congruence rule for " ^ quote c);
|
|
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
87 |
in AList.delete (op =) c congs end; |
| 9859 | 88 |
|
89 |
end; |
|
90 |
||
91 |
||
92 |
||
93 |
(** global and local recdef data **) |
|
94 |
||
| 17920 | 95 |
(* theory data *) |
| 6439 | 96 |
|
| 8657 | 97 |
type recdef_info = {simps: thm list, rules: thm list list, induct: thm, tcs: term list};
|
| 6439 | 98 |
|
| 16458 | 99 |
structure GlobalRecdefData = TheoryDataFun |
100 |
(struct |
|
| 6439 | 101 |
val name = "HOL/recdef"; |
| 9859 | 102 |
type T = recdef_info Symtab.table * hints; |
| 6439 | 103 |
|
| 9879 | 104 |
val empty = (Symtab.empty, mk_hints ([], [], [])): T; |
| 6557 | 105 |
val copy = I; |
| 16458 | 106 |
val extend = I; |
107 |
fun merge _ |
|
| 9859 | 108 |
((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
|
| 16458 | 109 |
(tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T =
|
| 9859 | 110 |
(Symtab.merge (K true) (tab1, tab2), |
111 |
mk_hints (Drule.merge_rules (simps1, simps2), |
|
|
21098
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
112 |
AList.merge (op =) eq_thm (congs1, congs2), |
| 9859 | 113 |
Drule.merge_rules (wfs1, wfs2))); |
| 6439 | 114 |
|
| 16458 | 115 |
fun print thy (tab, hints) = |
116 |
(Pretty.strs ("recdefs:" :: map #1 (NameSpace.extern_table (Sign.const_space thy, tab))) ::
|
|
| 16364 | 117 |
pretty_hints hints) |> Pretty.chunks |> Pretty.writeln; |
| 16458 | 118 |
end); |
| 6439 | 119 |
|
| 9859 | 120 |
val print_recdefs = GlobalRecdefData.print; |
| 6429 | 121 |
|
| 6439 | 122 |
|
| 17412 | 123 |
val get_recdef = Symtab.lookup o #1 o GlobalRecdefData.get; |
| 6439 | 124 |
|
125 |
fun put_recdef name info thy = |
|
| 6429 | 126 |
let |
| 9859 | 127 |
val (tab, hints) = GlobalRecdefData.get thy; |
| 17412 | 128 |
val tab' = Symtab.update_new (name, info) tab |
| 6439 | 129 |
handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name);
|
| 9859 | 130 |
in GlobalRecdefData.put (tab', hints) thy end; |
131 |
||
132 |
val get_global_hints = #2 o GlobalRecdefData.get; |
|
133 |
val map_global_hints = GlobalRecdefData.map o apsnd; |
|
134 |
||
135 |
||
| 17920 | 136 |
(* proof data *) |
| 9859 | 137 |
|
| 16458 | 138 |
structure LocalRecdefData = ProofDataFun |
139 |
(struct |
|
| 9859 | 140 |
val name = "HOL/recdef"; |
141 |
type T = hints; |
|
142 |
val init = get_global_hints; |
|
143 |
fun print _ hints = pretty_hints hints |> Pretty.chunks |> Pretty.writeln; |
|
| 16458 | 144 |
end); |
| 9859 | 145 |
|
146 |
val get_local_hints = LocalRecdefData.get; |
|
147 |
val map_local_hints = LocalRecdefData.map; |
|
148 |
||
149 |
||
| 20291 | 150 |
(* generic data *) |
151 |
||
152 |
fun get_hints (Context.Theory thy) = get_global_hints thy |
|
153 |
| get_hints (Context.Proof ctxt) = get_local_hints ctxt; |
|
| 9859 | 154 |
|
| 18688 | 155 |
fun map_hints f (Context.Theory thy) = Context.Theory (map_global_hints f thy) |
156 |
| map_hints f (Context.Proof ctxt) = Context.Proof (map_local_hints f ctxt); |
|
| 9859 | 157 |
|
| 20291 | 158 |
|
159 |
(* attributes *) |
|
160 |
||
| 18728 | 161 |
fun attrib f = Thm.declaration_attribute (map_hints o f); |
| 9859 | 162 |
|
| 18688 | 163 |
val simp_add = attrib (map_simps o Drule.add_rule); |
164 |
val simp_del = attrib (map_simps o Drule.del_rule); |
|
165 |
val cong_add = attrib (map_congs o add_cong); |
|
166 |
val cong_del = attrib (map_congs o del_cong); |
|
167 |
val wf_add = attrib (map_wfs o Drule.add_rule); |
|
168 |
val wf_del = attrib (map_wfs o Drule.del_rule); |
|
| 9859 | 169 |
|
170 |
||
| 9949 | 171 |
(* modifiers *) |
| 9859 | 172 |
|
| 9949 | 173 |
val recdef_simpN = "recdef_simp"; |
174 |
val recdef_congN = "recdef_cong"; |
|
175 |
val recdef_wfN = "recdef_wf"; |
|
| 9859 | 176 |
|
177 |
val recdef_modifiers = |
|
| 18728 | 178 |
[Args.$$$ recdef_simpN -- Args.colon >> K ((I, simp_add): Method.modifier), |
179 |
Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (I, simp_add), |
|
180 |
Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (I, simp_del), |
|
181 |
Args.$$$ recdef_congN -- Args.colon >> K (I, cong_add), |
|
182 |
Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (I, cong_add), |
|
183 |
Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (I, cong_del), |
|
184 |
Args.$$$ recdef_wfN -- Args.colon >> K (I, wf_add), |
|
185 |
Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (I, wf_add), |
|
186 |
Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (I, wf_del)] @ |
|
| 9949 | 187 |
Clasimp.clasimp_modifiers; |
| 9859 | 188 |
|
| 9949 | 189 |
|
| 9859 | 190 |
|
| 9949 | 191 |
(** prepare_hints(_i) **) |
| 9859 | 192 |
|
193 |
fun prepare_hints thy opt_src = |
|
194 |
let |
|
195 |
val ctxt0 = ProofContext.init thy; |
|
196 |
val ctxt = |
|
197 |
(case opt_src of |
|
| 15531 | 198 |
NONE => ctxt0 |
199 |
| SOME src => Method.only_sectioned_args recdef_modifiers I src ctxt0); |
|
| 9859 | 200 |
val {simps, congs, wfs} = get_local_hints ctxt;
|
| 15032 | 201 |
val cs = local_claset_of ctxt; |
202 |
val ss = local_simpset_of ctxt addsimps simps; |
|
|
21098
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
203 |
in (cs, ss, rev (map snd congs), wfs) end; |
| 9859 | 204 |
|
205 |
fun prepare_hints_i thy () = |
|
| 15032 | 206 |
let |
207 |
val ctxt0 = ProofContext.init thy; |
|
208 |
val {simps, congs, wfs} = get_global_hints thy;
|
|
|
21098
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
209 |
in (local_claset_of ctxt0, local_simpset_of ctxt0 addsimps simps, rev (map snd congs), wfs) end; |
| 9859 | 210 |
|
| 6439 | 211 |
|
212 |
||
213 |
(** add_recdef(_i) **) |
|
214 |
||
| 6557 | 215 |
fun requires_recdef thy = Theory.requires thy "Recdef" "recursive functions"; |
216 |
||
| 17920 | 217 |
fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy = |
| 6439 | 218 |
let |
| 9859 | 219 |
val _ = requires_recdef thy; |
220 |
||
| 16458 | 221 |
val name = Sign.intern_const thy raw_name; |
| 6439 | 222 |
val bname = Sign.base_name name; |
| 6429 | 223 |
val _ = message ("Defining recursive function " ^ quote name ^ " ...");
|
224 |
||
| 8657 | 225 |
val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs); |
226 |
val eq_atts = map (map (prep_att thy)) raw_eq_atts; |
|
227 |
||
| 9859 | 228 |
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
|
229 |
(*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
|
230 |
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
|
231 |
give rise to looping conditional rewriting.*) |
|
dfae7eb2830c
Prevent recdef from looping when the inductio rule is simplified
paulson
parents:
12876
diff
changeset
|
232 |
val (thy, {rules = rules_idx, induct, tcs}) =
|
| 17920 | 233 |
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
|
234 |
congs wfs name R eqs; |
|
21098
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
235 |
val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx); |
| 18728 | 236 |
val simp_att = if null tcs then [Simplifier.simp_add, RecfunCodegen.add NONE] else []; |
| 8657 | 237 |
|
| 18377 | 238 |
val ((simps' :: rules', [induct']), thy) = |
|
7798
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7262
diff
changeset
|
239 |
thy |
| 6439 | 240 |
|> Theory.add_path bname |
| 18688 | 241 |
|> PureThy.add_thmss |
242 |
((("simps", List.concat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
|
|
| 18377 | 243 |
||>> PureThy.add_thms [(("induct", induct), [])];
|
| 8657 | 244 |
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
|
245 |
val thy = |
|
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7262
diff
changeset
|
246 |
thy |
| 6439 | 247 |
|> put_recdef name result |
| 6429 | 248 |
|> Theory.parent_path; |
|
7798
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7262
diff
changeset
|
249 |
in (thy, result) end; |
| 6429 | 250 |
|
| 18728 | 251 |
val add_recdef = gen_add_recdef Tfl.define Attrib.attribute prepare_hints; |
| 11629 | 252 |
fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w (); |
| 9859 | 253 |
|
254 |
||
| 6557 | 255 |
|
256 |
(** defer_recdef(_i) **) |
|
257 |
||
258 |
fun gen_defer_recdef tfl_fn app_thms raw_name eqs raw_congs thy = |
|
259 |
let |
|
| 16458 | 260 |
val name = Sign.intern_const thy raw_name; |
| 6557 | 261 |
val bname = Sign.base_name name; |
262 |
||
263 |
val _ = requires_recdef thy; |
|
264 |
val _ = message ("Deferred recursive function " ^ quote name ^ " ...");
|
|
265 |
||
|
18418
bf448d999b7e
re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents:
18377
diff
changeset
|
266 |
val (congs, thy1) = thy |> app_thms raw_congs; |
| 9859 | 267 |
val (thy2, induct_rules) = tfl_fn thy1 congs name eqs; |
| 18377 | 268 |
val ([induct_rules'], thy3) = |
| 6557 | 269 |
thy2 |
270 |
|> Theory.add_path bname |
|
271 |
|> PureThy.add_thms [(("induct_rules", induct_rules), [])]
|
|
| 18377 | 272 |
||> Theory.parent_path; |
| 8430 | 273 |
in (thy3, {induct_rules = induct_rules'}) end;
|
| 6557 | 274 |
|
275 |
val defer_recdef = gen_defer_recdef Tfl.defer IsarThy.apply_theorems; |
|
276 |
val defer_recdef_i = gen_defer_recdef Tfl.defer_i IsarThy.apply_theorems_i; |
|
277 |
||
278 |
||
279 |
||
| 10775 | 280 |
(** recdef_tc(_i) **) |
281 |
||
| 17336 | 282 |
fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i thy = |
| 10775 | 283 |
let |
| 16458 | 284 |
val name = prep_name thy raw_name; |
| 10775 | 285 |
val atts = map (prep_att thy) raw_atts; |
286 |
val tcs = |
|
287 |
(case get_recdef thy name of |
|
| 15531 | 288 |
NONE => error ("No recdef definition of constant: " ^ quote name)
|
289 |
| SOME {tcs, ...} => tcs);
|
|
| 15570 | 290 |
val i = getOpt (opt_i, 1); |
291 |
val tc = List.nth (tcs, i - 1) handle Subscript => |
|
| 10775 | 292 |
error ("No termination condition #" ^ string_of_int i ^
|
293 |
" in recdef definition of " ^ quote name); |
|
| 19585 | 294 |
in IsarThy.theorem_i PureThy.internalK (bname, atts) (HOLogic.mk_Trueprop tc, []) thy end; |
| 10775 | 295 |
|
| 18728 | 296 |
val recdef_tc = gen_recdef_tc Attrib.attribute Sign.intern_const; |
| 10775 | 297 |
val recdef_tc_i = gen_recdef_tc (K I) (K I); |
298 |
||
299 |
||
300 |
||
| 6439 | 301 |
(** package setup **) |
302 |
||
303 |
(* setup theory *) |
|
304 |
||
| 9859 | 305 |
val setup = |
| 18708 | 306 |
GlobalRecdefData.init #> |
307 |
LocalRecdefData.init #> |
|
| 9859 | 308 |
Attrib.add_attributes |
| 18728 | 309 |
[(recdef_simpN, Attrib.add_del_args simp_add simp_del, "declaration of recdef simp rule"), |
|
19770
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
19686
diff
changeset
|
310 |
(recdef_congN, Attrib.add_del_args cong_add cong_del, "declaration of recdef cong rule"), |
| 18728 | 311 |
(recdef_wfN, Attrib.add_del_args wf_add wf_del, "declaration of recdef wf rule")]; |
| 6439 | 312 |
|
313 |
||
| 6429 | 314 |
(* outer syntax *) |
315 |
||
| 17057 | 316 |
local structure P = OuterParse and K = OuterKeyword in |
| 6429 | 317 |
|
| 9859 | 318 |
val hints = |
319 |
P.$$$ "(" |-- P.!!! (P.position (P.$$$ "hints" -- P.arguments) --| P.$$$ ")") >> Args.src;
|
|
320 |
||
| 6429 | 321 |
val recdef_decl = |
| 11629 | 322 |
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
|
323 |
P.name -- P.term -- Scan.repeat1 (P.opt_thm_name ":" -- P.prop) -- Scan.option hints |
| 11629 | 324 |
>> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map P.triple_swap eqs) src); |
| 6429 | 325 |
|
326 |
val recdefP = |
|
| 6723 | 327 |
OuterSyntax.command "recdef" "define general recursive functions (TFL)" K.thy_decl |
| 6429 | 328 |
(recdef_decl >> Toplevel.theory); |
329 |
||
| 6557 | 330 |
|
331 |
val defer_recdef_decl = |
|
| 8657 | 332 |
P.name -- Scan.repeat1 P.prop -- |
| 6723 | 333 |
Scan.optional (P.$$$ "(" |-- P.$$$ "congs" |-- P.!!! (P.xthms1 --| P.$$$ ")")) []
|
| 6557 | 334 |
>> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs); |
335 |
||
336 |
val defer_recdefP = |
|
| 6723 | 337 |
OuterSyntax.command "defer_recdef" "defer general recursive functions (TFL)" K.thy_decl |
| 6557 | 338 |
(defer_recdef_decl >> Toplevel.theory); |
339 |
||
| 10775 | 340 |
val recdef_tcP = |
341 |
OuterSyntax.command "recdef_tc" "recommence proof of termination condition (TFL)" K.thy_goal |
|
| 17336 | 342 |
(P.opt_thm_name ":" -- P.xname -- Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")")
|
343 |
>> (fn ((thm_name, name), i) => |
|
344 |
Toplevel.print o Toplevel.theory_to_proof (recdef_tc thm_name name i))); |
|
| 10775 | 345 |
|
346 |
||
| 11629 | 347 |
val _ = OuterSyntax.add_keywords ["permissive", "congs", "hints"]; |
| 10775 | 348 |
val _ = OuterSyntax.add_parsers [recdefP, defer_recdefP, recdef_tcP]; |
| 6429 | 349 |
|
350 |
end; |
|
351 |
||
352 |
end; |