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