| author | wenzelm | 
| Mon, 22 Sep 2014 16:28:24 +0200 | |
| changeset 58419 | 593917a7ad02 | 
| parent 58048 | aa6296d09e0e | 
| child 58825 | 2065f49da190 | 
| permissions | -rw-r--r-- | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31243diff
changeset | 1 | (* Title: HOL/Tools/recdef.ML | 
| 6429 | 2 | Author: Markus Wenzel, TU Muenchen | 
| 3 | ||
| 4 | Wrapper module for Konrad Slind's TFL package. | |
| 5 | *) | |
| 6 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31243diff
changeset | 7 | signature RECDEF = | 
| 6429 | 8 | sig | 
| 8657 | 9 | val get_recdef: theory -> string | 
| 35756 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35690diff
changeset | 10 |     -> {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list} option
 | 
| 21505 | 11 |   val get_hints: Proof.context -> {simps: thm list, congs: (string * thm) list, wfs: thm list}
 | 
| 18728 | 12 | val simp_add: attribute | 
| 13 | val simp_del: attribute | |
| 14 | val cong_add: attribute | |
| 15 | val cong_del: attribute | |
| 16 | val wf_add: attribute | |
| 17 | val wf_del: attribute | |
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57964diff
changeset | 18 | val add_recdef: bool -> xstring -> string -> ((binding * string) * Token.src list) list -> | 
| 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57964diff
changeset | 19 | Token.src option -> theory -> theory | 
| 35756 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35690diff
changeset | 20 |       * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
 | 
| 29579 | 21 | val add_recdef_i: bool -> xstring -> term -> ((binding * term) * attribute list) list -> | 
| 35756 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35690diff
changeset | 22 |     theory -> theory * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
 | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57964diff
changeset | 23 | val defer_recdef: xstring -> string list -> (Facts.ref * Token.src list) list | 
| 6557 | 24 |     -> theory -> theory * {induct_rules: thm}
 | 
| 27727 
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
 wenzelm parents: 
27353diff
changeset | 25 |   val defer_recdef_i: xstring -> term list -> thm list -> theory -> theory * {induct_rules: thm}
 | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57964diff
changeset | 26 | val recdef_tc: bstring * Token.src list -> xstring -> int option -> bool -> | 
| 27727 
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
 wenzelm parents: 
27353diff
changeset | 27 | local_theory -> Proof.state | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57964diff
changeset | 28 | val recdef_tc_i: bstring * Token.src list -> string -> int option -> bool -> | 
| 27727 
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
 wenzelm parents: 
27353diff
changeset | 29 | local_theory -> Proof.state | 
| 18708 | 30 | val setup: theory -> theory | 
| 6429 | 31 | end; | 
| 32 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31243diff
changeset | 33 | structure Recdef: RECDEF = | 
| 6429 | 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 | ||
| 51 | (* congruence rules *) | |
| 52 | ||
| 53 | local | |
| 54 | ||
| 55 | val cong_head = | |
| 56 | fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of; | |
| 6439 | 57 | |
| 9859 | 58 | fun prep_cong raw_thm = | 
| 59 | let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end; | |
| 60 | ||
| 61 | in | |
| 62 | ||
| 63 | fun add_cong raw_thm congs = | |
| 21098 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 haftmann parents: 
21078diff
changeset | 64 | let | 
| 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 haftmann parents: 
21078diff
changeset | 65 | val (c, thm) = prep_cong raw_thm; | 
| 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 haftmann parents: 
21078diff
changeset | 66 | val _ = if AList.defined (op =) congs c | 
| 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 haftmann parents: 
21078diff
changeset | 67 |       then warning ("Overwriting recdef congruence rule for " ^ quote c)
 | 
| 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 haftmann parents: 
21078diff
changeset | 68 | else (); | 
| 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 haftmann parents: 
21078diff
changeset | 69 | in AList.update (op =) (c, thm) congs end; | 
| 9859 | 70 | |
| 71 | fun del_cong raw_thm congs = | |
| 72 | let | |
| 73 | val (c, thm) = prep_cong raw_thm; | |
| 21098 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 haftmann parents: 
21078diff
changeset | 74 | val _ = if AList.defined (op =) congs c | 
| 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 haftmann parents: 
21078diff
changeset | 75 | then () | 
| 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 haftmann parents: 
21078diff
changeset | 76 |       else warning ("No recdef congruence rule for " ^ quote c);
 | 
| 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 haftmann parents: 
21078diff
changeset | 77 | in AList.delete (op =) c congs end; | 
| 9859 | 78 | |
| 79 | end; | |
| 80 | ||
| 81 | ||
| 82 | ||
| 83 | (** global and local recdef data **) | |
| 84 | ||
| 17920 | 85 | (* theory data *) | 
| 6439 | 86 | |
| 35756 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35690diff
changeset | 87 | type recdef_info = {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list};
 | 
| 6439 | 88 | |
| 33522 | 89 | structure GlobalRecdefData = Theory_Data | 
| 22846 | 90 | ( | 
| 9859 | 91 | type T = recdef_info Symtab.table * hints; | 
| 9879 | 92 | val empty = (Symtab.empty, mk_hints ([], [], [])): T; | 
| 16458 | 93 | val extend = I; | 
| 33522 | 94 | fun merge | 
| 9859 | 95 |    ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
 | 
| 16458 | 96 |     (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T =
 | 
| 9859 | 97 | (Symtab.merge (K true) (tab1, tab2), | 
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 98 | mk_hints (Thm.merge_thms (simps1, simps2), | 
| 33699 
f33b036ef318
permissive AList.merge -- most likely setup for theory data (beware of spurious AList.DUP);
 wenzelm parents: 
33643diff
changeset | 99 | AList.merge (op =) (K true) (congs1, congs2), | 
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 100 | Thm.merge_thms (wfs1, wfs2))); | 
| 22846 | 101 | ); | 
| 6439 | 102 | |
| 17412 | 103 | val get_recdef = Symtab.lookup o #1 o GlobalRecdefData.get; | 
| 6439 | 104 | |
| 105 | fun put_recdef name info thy = | |
| 6429 | 106 | let | 
| 9859 | 107 | val (tab, hints) = GlobalRecdefData.get thy; | 
| 17412 | 108 | val tab' = Symtab.update_new (name, info) tab | 
| 6439 | 109 |       handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name);
 | 
| 9859 | 110 | in GlobalRecdefData.put (tab', hints) thy end; | 
| 111 | ||
| 112 | val get_global_hints = #2 o GlobalRecdefData.get; | |
| 113 | ||
| 114 | ||
| 17920 | 115 | (* proof data *) | 
| 9859 | 116 | |
| 33519 | 117 | structure LocalRecdefData = Proof_Data | 
| 22846 | 118 | ( | 
| 9859 | 119 | type T = hints; | 
| 120 | val init = get_global_hints; | |
| 22846 | 121 | ); | 
| 9859 | 122 | |
| 21505 | 123 | val get_hints = LocalRecdefData.get; | 
| 124 | fun map_hints f = Context.mapping (GlobalRecdefData.map (apsnd f)) (LocalRecdefData.map f); | |
| 9859 | 125 | |
| 20291 | 126 | |
| 127 | (* attributes *) | |
| 128 | ||
| 18728 | 129 | fun attrib f = Thm.declaration_attribute (map_hints o f); | 
| 9859 | 130 | |
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 131 | 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 | 132 | val simp_del = attrib (map_simps o Thm.del_thm); | 
| 18688 | 133 | val cong_add = attrib (map_congs o add_cong); | 
| 134 | 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 | 135 | 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 | 136 | val wf_del = attrib (map_wfs o Thm.del_thm); | 
| 9859 | 137 | |
| 138 | ||
| 9949 | 139 | (* modifiers *) | 
| 9859 | 140 | |
| 9949 | 141 | val recdef_simpN = "recdef_simp"; | 
| 142 | val recdef_congN = "recdef_cong"; | |
| 143 | val recdef_wfN = "recdef_wf"; | |
| 9859 | 144 | |
| 145 | val recdef_modifiers = | |
| 58048 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58028diff
changeset | 146 |  [Args.$$$ recdef_simpN -- Args.colon >> K (Method.modifier simp_add @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58028diff
changeset | 147 |   Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58028diff
changeset | 148 |   Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58028diff
changeset | 149 |   Args.$$$ recdef_congN -- Args.colon >> K (Method.modifier cong_add @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58028diff
changeset | 150 |   Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (Method.modifier cong_add @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58028diff
changeset | 151 |   Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (Method.modifier cong_del @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58028diff
changeset | 152 |   Args.$$$ recdef_wfN -- Args.colon >> K (Method.modifier wf_add @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58028diff
changeset | 153 |   Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (Method.modifier wf_add @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58028diff
changeset | 154 |   Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (Method.modifier wf_del @{here})] @
 | 
| 9949 | 155 | Clasimp.clasimp_modifiers; | 
| 9859 | 156 | |
| 9949 | 157 | |
| 9859 | 158 | |
| 9949 | 159 | (** prepare_hints(_i) **) | 
| 9859 | 160 | |
| 161 | fun prepare_hints thy opt_src = | |
| 162 | let | |
| 42361 | 163 | val ctxt0 = Proof_Context.init_global thy; | 
| 9859 | 164 | val ctxt = | 
| 165 | (case opt_src of | |
| 15531 | 166 | NONE => ctxt0 | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57964diff
changeset | 167 | | SOME src => #2 (Token.syntax (Method.sections recdef_modifiers) src ctxt0)); | 
| 21505 | 168 |     val {simps, congs, wfs} = get_hints ctxt;
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50214diff
changeset | 169 |     val ctxt' = ctxt addsimps simps |> Simplifier.del_cong @{thm imp_cong};
 | 
| 42775 
a973f82fc885
prefer Proof.context over old-style claset/simpset;
 wenzelm parents: 
42361diff
changeset | 170 | in (ctxt', rev (map snd congs), wfs) end; | 
| 9859 | 171 | |
| 172 | fun prepare_hints_i thy () = | |
| 15032 | 173 | let | 
| 42775 
a973f82fc885
prefer Proof.context over old-style claset/simpset;
 wenzelm parents: 
42361diff
changeset | 174 | val ctxt = Proof_Context.init_global thy; | 
| 15032 | 175 |     val {simps, congs, wfs} = get_global_hints thy;
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50214diff
changeset | 176 |     val ctxt' = ctxt addsimps simps |> Simplifier.del_cong @{thm imp_cong};
 | 
| 42775 
a973f82fc885
prefer Proof.context over old-style claset/simpset;
 wenzelm parents: 
42361diff
changeset | 177 | in (ctxt', rev (map snd congs), wfs) end; | 
| 9859 | 178 | |
| 6439 | 179 | |
| 180 | ||
| 181 | (** add_recdef(_i) **) | |
| 182 | ||
| 56249 | 183 | fun requires_recdef thy = Theory.requires thy (Context.theory_name @{theory}) "recursive functions";
 | 
| 6557 | 184 | |
| 17920 | 185 | fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy = | 
| 6439 | 186 | let | 
| 36865 | 187 | val _ = legacy_feature "Old 'recdef' command -- use 'fun' or 'function' instead"; | 
| 9859 | 188 | val _ = requires_recdef thy; | 
| 189 | ||
| 16458 | 190 | val name = Sign.intern_const thy raw_name; | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 191 | val bname = Long_Name.base_name name; | 
| 26478 | 192 |     val _ = writeln ("Defining recursive function " ^ quote name ^ " ...");
 | 
| 6429 | 193 | |
| 8657 | 194 | val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs); | 
| 195 | val eq_atts = map (map (prep_att thy)) raw_eq_atts; | |
| 196 | ||
| 42775 
a973f82fc885
prefer Proof.context over old-style claset/simpset;
 wenzelm parents: 
42361diff
changeset | 197 | val (ctxt, congs, wfs) = prep_hints thy hints; | 
| 14241 
dfae7eb2830c
Prevent recdef from looping when the inductio rule is simplified
 paulson parents: 
12876diff
changeset | 198 | (*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 | 199 | 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 | 200 | give rise to looping conditional rewriting.*) | 
| 42775 
a973f82fc885
prefer Proof.context over old-style claset/simpset;
 wenzelm parents: 
42361diff
changeset | 201 |     val ({lhs, rules = rules_idx, induct, tcs}, thy) =
 | 
| 
a973f82fc885
prefer Proof.context over old-style claset/simpset;
 wenzelm parents: 
42361diff
changeset | 202 | tfl_fn not_permissive ctxt congs wfs name R eqs thy; | 
| 21098 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 haftmann parents: 
21078diff
changeset | 203 | val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx); | 
| 33552 | 204 | val simp_att = | 
| 57964 | 205 | if null tcs then [Simplifier.simp_add, | 
| 206 |         Named_Theorems.add @{named_theorems nitpick_simp}, Code.add_default_eqn_attribute]
 | |
| 33552 | 207 | else []; | 
| 18377 | 208 | val ((simps' :: rules', [induct']), thy) = | 
| 7798 
42e94b618f34
return stored thms with proper naming in derivation;
 wenzelm parents: 
7262diff
changeset | 209 | thy | 
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24624diff
changeset | 210 | |> Sign.add_path bname | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
36960diff
changeset | 211 | |> Global_Theory.add_thmss | 
| 32952 | 212 | (((Binding.name "simps", flat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts)) | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
36960diff
changeset | 213 | ||>> Global_Theory.add_thms [((Binding.name "induct", induct), [])] | 
| 35756 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35690diff
changeset | 214 | ||> Spec_Rules.add_global Spec_Rules.Equational ([lhs], flat rules); | 
| 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35690diff
changeset | 215 |     val result = {lhs = lhs, simps = simps', rules = rules', induct = induct', tcs = tcs};
 | 
| 7798 
42e94b618f34
return stored thms with proper naming in derivation;
 wenzelm parents: 
7262diff
changeset | 216 | val thy = | 
| 
42e94b618f34
return stored thms with proper naming in derivation;
 wenzelm parents: 
7262diff
changeset | 217 | thy | 
| 6439 | 218 | |> put_recdef name result | 
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24624diff
changeset | 219 | |> Sign.parent_path; | 
| 7798 
42e94b618f34
return stored thms with proper naming in derivation;
 wenzelm parents: 
7262diff
changeset | 220 | in (thy, result) end; | 
| 6429 | 221 | |
| 47815 | 222 | val add_recdef = gen_add_recdef Tfl.define Attrib.attribute_cmd_global prepare_hints; | 
| 11629 | 223 | fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w (); | 
| 9859 | 224 | |
| 225 | ||
| 6557 | 226 | |
| 227 | (** defer_recdef(_i) **) | |
| 228 | ||
| 27727 
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
 wenzelm parents: 
27353diff
changeset | 229 | fun gen_defer_recdef tfl_fn eval_thms raw_name eqs raw_congs thy = | 
| 6557 | 230 | let | 
| 16458 | 231 | val name = Sign.intern_const thy raw_name; | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 232 | val bname = Long_Name.base_name name; | 
| 6557 | 233 | |
| 234 | val _ = requires_recdef thy; | |
| 26478 | 235 |     val _ = writeln ("Deferred recursive function " ^ quote name ^ " ...");
 | 
| 6557 | 236 | |
| 42361 | 237 | val congs = eval_thms (Proof_Context.init_global thy) raw_congs; | 
| 42775 
a973f82fc885
prefer Proof.context over old-style claset/simpset;
 wenzelm parents: 
42361diff
changeset | 238 | val (induct_rules, thy2) = tfl_fn congs name eqs thy; | 
| 18377 | 239 | val ([induct_rules'], thy3) = | 
| 6557 | 240 | thy2 | 
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24624diff
changeset | 241 | |> Sign.add_path bname | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
36960diff
changeset | 242 | |> Global_Theory.add_thms [((Binding.name "induct_rules", induct_rules), [])] | 
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24624diff
changeset | 243 | ||> Sign.parent_path; | 
| 8430 | 244 |   in (thy3, {induct_rules = induct_rules'}) end;
 | 
| 6557 | 245 | |
| 27727 
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
 wenzelm parents: 
27353diff
changeset | 246 | val defer_recdef = gen_defer_recdef Tfl.defer Attrib.eval_thms; | 
| 
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
 wenzelm parents: 
27353diff
changeset | 247 | val defer_recdef_i = gen_defer_recdef Tfl.defer_i (K I); | 
| 6557 | 248 | |
| 249 | ||
| 250 | ||
| 10775 | 251 | (** recdef_tc(_i) **) | 
| 252 | ||
| 24457 
a33258c78ed2
Adapted to changes in interface of Specification.theorem_i
 berghofe parents: 
24039diff
changeset | 253 | fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i int lthy = | 
| 10775 | 254 | let | 
| 42361 | 255 | val thy = Proof_Context.theory_of lthy; | 
| 16458 | 256 | val name = prep_name thy raw_name; | 
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
51717diff
changeset | 257 | val atts = map (prep_att lthy) raw_atts; | 
| 10775 | 258 | val tcs = | 
| 259 | (case get_recdef thy name of | |
| 15531 | 260 |         NONE => error ("No recdef definition of constant: " ^ quote name)
 | 
| 261 |       | SOME {tcs, ...} => tcs);
 | |
| 21351 
1fb804b96d7c
recdef_tc(_i): local_theory interface via Specification.theorem_i;
 wenzelm parents: 
21098diff
changeset | 262 | val i = the_default 1 opt_i; | 
| 43278 | 263 | val tc = nth tcs (i - 1) handle General.Subscript => | 
| 10775 | 264 |       error ("No termination condition #" ^ string_of_int i ^
 | 
| 265 | " in recdef definition of " ^ quote name); | |
| 21351 
1fb804b96d7c
recdef_tc(_i): local_theory interface via Specification.theorem_i;
 wenzelm parents: 
21098diff
changeset | 266 | in | 
| 33643 
b275f26a638b
eliminated obsolete "internal" kind -- collapsed to unspecific "";
 wenzelm parents: 
33552diff
changeset | 267 | Specification.theorem "" NONE (K I) | 
| 47067 
4ef29b0c568f
optional 'includes' element for long theorem statements;
 wenzelm parents: 
46961diff
changeset | 268 | (Binding.conceal (Binding.name bname), atts) [] [] | 
| 33278 | 269 | (Element.Shows [(Attrib.empty_binding, [(HOLogic.mk_Trueprop tc, [])])]) int lthy | 
| 21351 
1fb804b96d7c
recdef_tc(_i): local_theory interface via Specification.theorem_i;
 wenzelm parents: 
21098diff
changeset | 270 | end; | 
| 10775 | 271 | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
51717diff
changeset | 272 | val recdef_tc = gen_recdef_tc Attrib.check_src Sign.intern_const; | 
| 10775 | 273 | val recdef_tc_i = gen_recdef_tc (K I) (K I); | 
| 274 | ||
| 275 | ||
| 276 | ||
| 6439 | 277 | (** package setup **) | 
| 278 | ||
| 279 | (* setup theory *) | |
| 280 | ||
| 9859 | 281 | val setup = | 
| 30528 | 282 |   Attrib.setup @{binding recdef_simp} (Attrib.add_del simp_add simp_del)
 | 
| 283 | "declaration of recdef simp rule" #> | |
| 284 |   Attrib.setup @{binding recdef_cong} (Attrib.add_del cong_add cong_del)
 | |
| 285 | "declaration of recdef cong rule" #> | |
| 286 |   Attrib.setup @{binding recdef_wf} (Attrib.add_del wf_add wf_del)
 | |
| 287 | "declaration of recdef wf rule"; | |
| 6439 | 288 | |
| 289 | ||
| 6429 | 290 | (* outer syntax *) | 
| 291 | ||
| 9859 | 292 | val hints = | 
| 46949 | 293 |   @{keyword "("} |--
 | 
| 56201 | 294 |     Parse.!!! (Parse.position @{keyword "hints"} -- Parse.args --| @{keyword ")"})
 | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57964diff
changeset | 295 | >> uncurry Token.src; | 
| 9859 | 296 | |
| 6429 | 297 | val recdef_decl = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 298 | Scan.optional | 
| 46949 | 299 |     (@{keyword "("} -- Parse.!!! (@{keyword "permissive"} -- @{keyword ")"}) >> K false) true --
 | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 300 | Parse.name -- Parse.term -- Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
27809diff
changeset | 301 | -- Scan.option hints | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 302 | >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map Parse.triple_swap eqs) src); | 
| 6429 | 303 | |
| 24867 | 304 | val _ = | 
| 50214 | 305 |   Outer_Syntax.command @{command_spec "recdef"} "define general recursive functions (obsolete TFL)"
 | 
| 6429 | 306 | (recdef_decl >> Toplevel.theory); | 
| 307 | ||
| 6557 | 308 | |
| 309 | val defer_recdef_decl = | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 310 | Parse.name -- Scan.repeat1 Parse.prop -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 311 | Scan.optional | 
| 58028 
e4250d370657
tuned signature -- define some elementary operations earlier;
 wenzelm parents: 
58011diff
changeset | 312 |     (@{keyword "("} |-- @{keyword "congs"} |-- Parse.!!! (Parse.xthms1 --| @{keyword ")"})) []
 | 
| 6557 | 313 | >> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs); | 
| 314 | ||
| 24867 | 315 | val _ = | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 316 |   Outer_Syntax.command @{command_spec "defer_recdef"}
 | 
| 50214 | 317 | "defer general recursive functions (obsolete TFL)" | 
| 6557 | 318 | (defer_recdef_decl >> Toplevel.theory); | 
| 319 | ||
| 24867 | 320 | val _ = | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 321 |   Outer_Syntax.local_theory_to_proof' @{command_spec "recdef_tc"}
 | 
| 50214 | 322 | "recommence proof of termination condition (obsolete TFL)" | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 323 | ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.xname -- | 
| 46949 | 324 |         Scan.option (@{keyword "("} |-- Parse.nat --| @{keyword ")"})
 | 
| 26988 
742e26213212
more uniform treatment of OuterSyntax.local_theory commands;
 wenzelm parents: 
26478diff
changeset | 325 | >> (fn ((thm_name, name), i) => recdef_tc thm_name name i)); | 
| 10775 | 326 | |
| 6429 | 327 | end; |