| author | kleing | 
| Thu, 28 Feb 2002 15:12:09 +0100 | |
| changeset 12973 | 8040cce614e5 | 
| parent 12876 | a70df1e5bf10 | 
| child 14241 | dfae7eb2830c | 
| 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
 | |
| 9859 | 15 | val simp_add_global: theory attribute | 
| 16 | val simp_del_global: theory attribute | |
| 17 | val cong_add_global: theory attribute | |
| 18 | val cong_del_global: theory attribute | |
| 19 | val wf_add_global: theory attribute | |
| 20 | val wf_del_global: theory attribute | |
| 21 | val simp_add_local: Proof.context attribute | |
| 22 | val simp_del_local: Proof.context attribute | |
| 23 | val cong_add_local: Proof.context attribute | |
| 24 | val cong_del_local: Proof.context attribute | |
| 25 | val wf_add_local: Proof.context attribute | |
| 26 | val wf_del_local: Proof.context attribute | |
| 11629 | 27 | val add_recdef: bool -> xstring -> string -> ((bstring * string) * Args.src list) list -> | 
| 9859 | 28 | Args.src option -> theory -> theory | 
| 29 |       * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
 | |
| 11629 | 30 | val add_recdef_i: bool -> xstring -> term -> ((bstring * term) * theory attribute list) list -> | 
| 9859 | 31 |     theory -> theory * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
 | 
| 32 | val add_recdef_old: xstring -> string -> ((bstring * string) * Args.src list) list -> | |
| 33 | simpset * thm list -> theory -> | |
| 34 |     theory * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
 | |
| 6557 | 35 | val defer_recdef: xstring -> string list -> (xstring * Args.src list) list | 
| 36 |     -> theory -> theory * {induct_rules: thm}
 | |
| 12708 | 37 | val defer_recdef_i: xstring -> term list -> (thm list * theory attribute list) list | 
| 6557 | 38 |     -> theory -> theory * {induct_rules: thm}
 | 
| 10775 | 39 | val recdef_tc: bstring * Args.src list -> xstring -> int option | 
| 40 | -> bool -> theory -> ProofHistory.T | |
| 41 | val recdef_tc_i: bstring * theory attribute list -> string -> int option | |
| 42 | -> bool -> theory -> ProofHistory.T | |
| 6439 | 43 | val setup: (theory -> theory) list | 
| 6429 | 44 | end; | 
| 45 | ||
| 46 | structure RecdefPackage: RECDEF_PACKAGE = | |
| 47 | struct | |
| 48 | ||
| 9859 | 49 | |
| 6429 | 50 | val quiet_mode = Tfl.quiet_mode; | 
| 51 | val message = Tfl.message; | |
| 52 | ||
| 53 | ||
| 9859 | 54 | (** recdef hints **) | 
| 6439 | 55 | |
| 9859 | 56 | (* type hints *) | 
| 57 | ||
| 58 | type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
 | |
| 59 | ||
| 60 | fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints;
 | |
| 61 | fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs));
 | |
| 62 | ||
| 63 | fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs)); | |
| 64 | fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs)); | |
| 65 | fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs)); | |
| 66 | ||
| 67 | fun pretty_hints ({simps, congs, wfs}: hints) =
 | |
| 68 | [Pretty.big_list "recdef simp hints:" (map Display.pretty_thm simps), | |
| 69 | Pretty.big_list "recdef cong hints:" (map Display.pretty_thm (map #2 congs)), | |
| 70 | Pretty.big_list "recdef wf hints:" (map Display.pretty_thm wfs)]; | |
| 71 | ||
| 72 | ||
| 73 | (* congruence rules *) | |
| 74 | ||
| 75 | local | |
| 76 | ||
| 77 | val cong_head = | |
| 78 | fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of; | |
| 6439 | 79 | |
| 9859 | 80 | fun prep_cong raw_thm = | 
| 81 | let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end; | |
| 82 | ||
| 83 | in | |
| 84 | ||
| 85 | fun add_cong raw_thm congs = | |
| 86 | let val (c, thm) = prep_cong raw_thm | |
| 87 |   in overwrite_warn (congs, (c, thm)) ("Overwriting recdef congruence rule for " ^ quote c) end;
 | |
| 88 | ||
| 89 | fun del_cong raw_thm congs = | |
| 90 | let | |
| 91 | val (c, thm) = prep_cong raw_thm; | |
| 92 | val (del, rest) = Library.partition (Library.equal c o fst) congs; | |
| 93 |   in if null del then (warning ("No recdef congruence rule for " ^ quote c); congs) else rest end;
 | |
| 94 | ||
| 95 | val add_congs = curry (foldr (uncurry add_cong)); | |
| 96 | ||
| 97 | end; | |
| 98 | ||
| 99 | ||
| 100 | ||
| 101 | (** global and local recdef data **) | |
| 102 | ||
| 103 | (* theory data kind 'HOL/recdef' *) | |
| 6439 | 104 | |
| 8657 | 105 | type recdef_info = {simps: thm list, rules: thm list list, induct: thm, tcs: term list};
 | 
| 6439 | 106 | |
| 9859 | 107 | structure GlobalRecdefArgs = | 
| 6439 | 108 | struct | 
| 109 | val name = "HOL/recdef"; | |
| 9859 | 110 | type T = recdef_info Symtab.table * hints; | 
| 6439 | 111 | |
| 9879 | 112 | val empty = (Symtab.empty, mk_hints ([], [], [])): T; | 
| 6557 | 113 | val copy = I; | 
| 6439 | 114 | val prep_ext = I; | 
| 9859 | 115 | fun merge | 
| 116 |    ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
 | |
| 117 |     (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) =
 | |
| 118 | (Symtab.merge (K true) (tab1, tab2), | |
| 119 | mk_hints (Drule.merge_rules (simps1, simps2), | |
| 120 | Library.merge_alists congs1 congs2, | |
| 121 | Drule.merge_rules (wfs1, wfs2))); | |
| 6439 | 122 | |
| 9859 | 123 | fun print sg (tab, hints) = | 
| 124 |    (Pretty.strs ("recdefs:" :: map #1 (Sign.cond_extern_table sg Sign.constK tab)) ::
 | |
| 125 | pretty_hints hints) |> Pretty.chunks |> Pretty.writeln; | |
| 6439 | 126 | end; | 
| 127 | ||
| 9859 | 128 | structure GlobalRecdefData = TheoryDataFun(GlobalRecdefArgs); | 
| 129 | val print_recdefs = GlobalRecdefData.print; | |
| 6429 | 130 | |
| 6439 | 131 | |
| 9859 | 132 | fun get_recdef thy name = Symtab.lookup (#1 (GlobalRecdefData.get thy), name); | 
| 6439 | 133 | |
| 134 | fun put_recdef name info thy = | |
| 6429 | 135 | let | 
| 9859 | 136 | val (tab, hints) = GlobalRecdefData.get thy; | 
| 137 | val tab' = Symtab.update_new ((name, info), tab) | |
| 6439 | 138 |       handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name);
 | 
| 9859 | 139 | in GlobalRecdefData.put (tab', hints) thy end; | 
| 140 | ||
| 141 | val get_global_hints = #2 o GlobalRecdefData.get; | |
| 142 | val map_global_hints = GlobalRecdefData.map o apsnd; | |
| 143 | ||
| 144 | ||
| 145 | (* proof data kind 'HOL/recdef' *) | |
| 146 | ||
| 147 | structure LocalRecdefArgs = | |
| 148 | struct | |
| 149 | val name = "HOL/recdef"; | |
| 150 | type T = hints; | |
| 151 | val init = get_global_hints; | |
| 152 | fun print _ hints = pretty_hints hints |> Pretty.chunks |> Pretty.writeln; | |
| 153 | end; | |
| 154 | ||
| 155 | structure LocalRecdefData = ProofDataFun(LocalRecdefArgs); | |
| 156 | val get_local_hints = LocalRecdefData.get; | |
| 157 | val map_local_hints = LocalRecdefData.map; | |
| 158 | ||
| 159 | ||
| 160 | (* attributes *) | |
| 161 | ||
| 162 | local | |
| 163 | ||
| 164 | fun global_local f g = | |
| 165 | (fn (thy, thm) => (map_global_hints (f (g thm)) thy, thm), | |
| 166 | fn (ctxt, thm) => (map_local_hints (f (g thm)) ctxt, thm)); | |
| 167 | ||
| 168 | fun mk_attr (add1, add2) (del1, del2) = | |
| 169 | (Attrib.add_del_args add1 del1, Attrib.add_del_args add2 del2); | |
| 170 | ||
| 171 | in | |
| 172 | ||
| 12371 | 173 | val (simp_add_global, simp_add_local) = global_local map_simps Drule.add_rule; | 
| 174 | val (simp_del_global, simp_del_local) = global_local map_simps Drule.del_rule; | |
| 9859 | 175 | val (cong_add_global, cong_add_local) = global_local map_congs add_cong; | 
| 176 | val (cong_del_global, cong_del_local) = global_local map_congs del_cong; | |
| 12371 | 177 | val (wf_add_global, wf_add_local) = global_local map_wfs Drule.add_rule; | 
| 178 | val (wf_del_global, wf_del_local) = global_local map_wfs Drule.del_rule; | |
| 9859 | 179 | |
| 180 | val simp_attr = mk_attr (simp_add_global, simp_add_local) (simp_del_global, simp_del_local); | |
| 181 | val cong_attr = mk_attr (cong_add_global, cong_add_local) (cong_del_global, cong_del_local); | |
| 182 | val wf_attr = mk_attr (wf_add_global, wf_add_local) (wf_del_global, wf_del_local); | |
| 183 | ||
| 184 | end; | |
| 185 | ||
| 186 | ||
| 9949 | 187 | (* modifiers *) | 
| 9859 | 188 | |
| 9949 | 189 | val recdef_simpN = "recdef_simp"; | 
| 190 | val recdef_congN = "recdef_cong"; | |
| 191 | val recdef_wfN = "recdef_wf"; | |
| 9859 | 192 | |
| 193 | val recdef_modifiers = | |
| 9949 | 194 | [Args.$$$ recdef_simpN -- Args.colon >> K ((I, simp_add_local): Method.modifier), | 
| 10032 | 195 | Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (I, simp_add_local), | 
| 196 | Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (I, simp_del_local), | |
| 9949 | 197 | Args.$$$ recdef_congN -- Args.colon >> K (I, cong_add_local), | 
| 10032 | 198 | Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (I, cong_add_local), | 
| 199 | Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (I, cong_del_local), | |
| 9949 | 200 | Args.$$$ recdef_wfN -- Args.colon >> K (I, wf_add_local), | 
| 10032 | 201 | Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (I, wf_add_local), | 
| 202 | Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (I, wf_del_local)] @ | |
| 9949 | 203 | Clasimp.clasimp_modifiers; | 
| 9859 | 204 | |
| 9949 | 205 | |
| 9859 | 206 | |
| 9949 | 207 | (** prepare_hints(_i) **) | 
| 9859 | 208 | |
| 209 | fun prepare_hints thy opt_src = | |
| 210 | let | |
| 211 | val ctxt0 = ProofContext.init thy; | |
| 212 | val ctxt = | |
| 213 | (case opt_src of | |
| 214 | None => ctxt0 | |
| 9949 | 215 | | Some src => Method.only_sectioned_args recdef_modifiers I src ctxt0); | 
| 9859 | 216 |     val {simps, congs, wfs} = get_local_hints ctxt;
 | 
| 217 | val cs = Classical.get_local_claset ctxt; | |
| 218 | val ss = Simplifier.get_local_simpset ctxt addsimps simps; | |
| 219 | in (cs, ss, map #2 congs, wfs) end; | |
| 220 | ||
| 221 | fun prepare_hints_i thy () = | |
| 222 |   let val {simps, congs, wfs} = get_global_hints thy
 | |
| 223 | in (Classical.claset_of thy, Simplifier.simpset_of thy addsimps simps, map #2 congs, wfs) end; | |
| 224 | ||
| 6439 | 225 | |
| 226 | ||
| 227 | (** add_recdef(_i) **) | |
| 228 | ||
| 6557 | 229 | fun requires_recdef thy = Theory.requires thy "Recdef" "recursive functions"; | 
| 230 | ||
| 11629 | 231 | fun gen_add_recdef tfl_fn prep_att prep_hints strict raw_name R eq_srcs hints thy = | 
| 6439 | 232 | let | 
| 9859 | 233 | val _ = requires_recdef thy; | 
| 234 | ||
| 6439 | 235 | val name = Sign.intern_const (Theory.sign_of thy) raw_name; | 
| 236 | val bname = Sign.base_name name; | |
| 6429 | 237 |     val _ = message ("Defining recursive function " ^ quote name ^ " ...");
 | 
| 238 | ||
| 8657 | 239 | val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs); | 
| 240 | val eq_atts = map (map (prep_att thy)) raw_eq_atts; | |
| 241 | ||
| 9859 | 242 | val (cs, ss, congs, wfs) = prep_hints thy hints; | 
| 11629 | 243 |     val (thy, {rules = rules_idx, induct, tcs}) = tfl_fn strict thy cs ss congs wfs name R eqs;
 | 
| 8657 | 244 | val rules = map (map #1) (Library.partition_eq Library.eq_snd rules_idx); | 
| 12448 
473cb9f9e237
Recursive equations to be used for code generation are now registered
 berghofe parents: 
12371diff
changeset | 245 | val simp_att = if null tcs then | 
| 
473cb9f9e237
Recursive equations to be used for code generation are now registered
 berghofe parents: 
12371diff
changeset | 246 | [Simplifier.simp_add_global, RecfunCodegen.add] else []; | 
| 8657 | 247 | |
| 248 | val (thy, (simps' :: rules', [induct'])) = | |
| 7798 
42e94b618f34
return stored thms with proper naming in derivation;
 wenzelm parents: 
7262diff
changeset | 249 | thy | 
| 6439 | 250 | |> Theory.add_path bname | 
| 8657 | 251 |       |> PureThy.add_thmss ((("simps", flat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
 | 
| 12755 
a906a8b364f1
removed case_numbers (already covered by default);
 wenzelm parents: 
12708diff
changeset | 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: 
7262diff
changeset | 254 | val thy = | 
| 
42e94b618f34
return stored thms with proper naming in derivation;
 wenzelm parents: 
7262diff
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: 
7262diff
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 | ||
| 264 | (* add_recdef_old -- legacy interface *) | |
| 265 | ||
| 266 | fun prepare_hints_old thy (ss, thms) = | |
| 267 |   let val {simps, congs, wfs} = get_global_hints thy
 | |
| 268 | in (Classical.claset_of thy, ss addsimps simps, map #2 (add_congs thms congs), wfs) end; | |
| 269 | ||
| 11629 | 270 | val add_recdef_old = gen_add_recdef Tfl.define Attrib.global_attribute prepare_hints_old false; | 
| 6429 | 271 | |
| 272 | ||
| 6557 | 273 | |
| 274 | (** defer_recdef(_i) **) | |
| 275 | ||
| 276 | fun gen_defer_recdef tfl_fn app_thms raw_name eqs raw_congs thy = | |
| 277 | let | |
| 278 | val name = Sign.intern_const (Theory.sign_of thy) raw_name; | |
| 279 | val bname = Sign.base_name name; | |
| 280 | ||
| 281 | val _ = requires_recdef thy; | |
| 282 |     val _ = message ("Deferred recursive function " ^ quote name ^ " ...");
 | |
| 283 | ||
| 284 | val (thy1, congs) = thy |> app_thms raw_congs; | |
| 9859 | 285 | val (thy2, induct_rules) = tfl_fn thy1 congs name eqs; | 
| 8430 | 286 | val (thy3, [induct_rules']) = | 
| 6557 | 287 | thy2 | 
| 288 | |> Theory.add_path bname | |
| 289 |       |> PureThy.add_thms [(("induct_rules", induct_rules), [])]
 | |
| 8430 | 290 | |>> Theory.parent_path; | 
| 291 |   in (thy3, {induct_rules = induct_rules'}) end;
 | |
| 6557 | 292 | |
| 293 | val defer_recdef = gen_defer_recdef Tfl.defer IsarThy.apply_theorems; | |
| 294 | val defer_recdef_i = gen_defer_recdef Tfl.defer_i IsarThy.apply_theorems_i; | |
| 295 | ||
| 296 | ||
| 297 | ||
| 10775 | 298 | (** recdef_tc(_i) **) | 
| 299 | ||
| 300 | fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i int thy = | |
| 301 | let | |
| 302 | val name = prep_name (Theory.sign_of thy) raw_name; | |
| 303 | val atts = map (prep_att thy) raw_atts; | |
| 304 | val tcs = | |
| 305 | (case get_recdef thy name of | |
| 306 |         None => error ("No recdef definition of constant: " ^ quote name)
 | |
| 307 |       | Some {tcs, ...} => tcs);
 | |
| 308 | val i = if_none opt_i 1; | |
| 309 | val tc = Library.nth_elem (i - 1, tcs) handle Library.LIST _ => | |
| 310 |       error ("No termination condition #" ^ string_of_int i ^
 | |
| 311 | " in recdef definition of " ^ quote name); | |
| 312 | in | |
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
12755diff
changeset | 313 | thy | 
| 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
12755diff
changeset | 314 | |> IsarThy.theorem_i Drule.internalK ((bname, atts), (HOLogic.mk_Trueprop tc, ([], []))) int | 
| 10775 | 315 | end; | 
| 316 | ||
| 317 | val recdef_tc = gen_recdef_tc Attrib.global_attribute Sign.intern_const; | |
| 318 | val recdef_tc_i = gen_recdef_tc (K I) (K I); | |
| 319 | ||
| 320 | ||
| 321 | ||
| 6439 | 322 | (** package setup **) | 
| 323 | ||
| 324 | (* setup theory *) | |
| 325 | ||
| 9859 | 326 | val setup = | 
| 327 | [GlobalRecdefData.init, LocalRecdefData.init, | |
| 328 | Attrib.add_attributes | |
| 9949 | 329 | [(recdef_simpN, simp_attr, "declaration of recdef simp rule"), | 
| 330 | (recdef_congN, cong_attr, "declaration of recdef cong rule"), | |
| 331 | (recdef_wfN, wf_attr, "declaration of recdef wf rule")]]; | |
| 6439 | 332 | |
| 333 | ||
| 6429 | 334 | (* outer syntax *) | 
| 335 | ||
| 6723 | 336 | local structure P = OuterParse and K = OuterSyntax.Keyword in | 
| 6429 | 337 | |
| 9859 | 338 | val hints = | 
| 339 |   P.$$$ "(" |-- P.!!! (P.position (P.$$$ "hints" -- P.arguments) --| P.$$$ ")") >> Args.src;
 | |
| 340 | ||
| 6429 | 341 | val recdef_decl = | 
| 11629 | 342 |   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: 
12755diff
changeset | 343 | P.name -- P.term -- Scan.repeat1 (P.opt_thm_name ":" -- P.prop) -- Scan.option hints | 
| 11629 | 344 | >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map P.triple_swap eqs) src); | 
| 6429 | 345 | |
| 346 | val recdefP = | |
| 6723 | 347 | OuterSyntax.command "recdef" "define general recursive functions (TFL)" K.thy_decl | 
| 6429 | 348 | (recdef_decl >> Toplevel.theory); | 
| 349 | ||
| 6557 | 350 | |
| 351 | val defer_recdef_decl = | |
| 8657 | 352 | P.name -- Scan.repeat1 P.prop -- | 
| 6723 | 353 |   Scan.optional (P.$$$ "(" |-- P.$$$ "congs" |-- P.!!! (P.xthms1 --| P.$$$ ")")) []
 | 
| 6557 | 354 | >> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs); | 
| 355 | ||
| 356 | val defer_recdefP = | |
| 6723 | 357 | OuterSyntax.command "defer_recdef" "defer general recursive functions (TFL)" K.thy_decl | 
| 6557 | 358 | (defer_recdef_decl >> Toplevel.theory); | 
| 359 | ||
| 10775 | 360 | |
| 361 | val recdef_tc_decl = | |
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
12755diff
changeset | 362 |   P.opt_thm_name ":" -- P.xname -- Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")");
 | 
| 10775 | 363 | |
| 364 | fun mk_recdef_tc ((thm_name, name), i) = recdef_tc thm_name name i; | |
| 365 | ||
| 366 | val recdef_tcP = | |
| 367 | OuterSyntax.command "recdef_tc" "recommence proof of termination condition (TFL)" K.thy_goal | |
| 368 | (recdef_tc_decl >> (Toplevel.print oo (Toplevel.theory_to_proof o mk_recdef_tc))); | |
| 369 | ||
| 370 | ||
| 11629 | 371 | val _ = OuterSyntax.add_keywords ["permissive", "congs", "hints"]; | 
| 10775 | 372 | val _ = OuterSyntax.add_parsers [recdefP, defer_recdefP, recdef_tcP]; | 
| 6429 | 373 | |
| 374 | end; | |
| 375 | ||
| 376 | ||
| 377 | end; |