author | wenzelm |
Mon, 27 Feb 2012 15:48:02 +0100 | |
changeset 46708 | b138dee7bed3 |
parent 46185 | afda84cd4d4b |
child 46947 | b8c7eb0c2f89 |
permissions | -rw-r--r-- |
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31243
diff
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:
31243
diff
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:
35690
diff
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 |
|
29579 | 18 |
val add_recdef: bool -> xstring -> string -> ((binding * string) * Attrib.src list) list -> |
15703 | 19 |
Attrib.src option -> theory -> theory |
35756
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
bulwahn
parents:
35690
diff
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:
35690
diff
changeset
|
22 |
theory -> theory * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list} |
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
24927
diff
changeset
|
23 |
val defer_recdef: xstring -> string list -> (Facts.ref * Attrib.src list) list |
6557 | 24 |
-> theory -> theory * {induct_rules: thm} |
27727
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
wenzelm
parents:
27353
diff
changeset
|
25 |
val defer_recdef_i: xstring -> term list -> thm list -> theory -> theory * {induct_rules: thm} |
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
wenzelm
parents:
27353
diff
changeset
|
26 |
val recdef_tc: bstring * Attrib.src list -> xstring -> int option -> bool -> |
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
wenzelm
parents:
27353
diff
changeset
|
27 |
local_theory -> Proof.state |
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
wenzelm
parents:
27353
diff
changeset
|
28 |
val recdef_tc_i: bstring * Attrib.src list -> string -> int option -> bool -> |
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
wenzelm
parents:
27353
diff
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:
31243
diff
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:
21078
diff
changeset
|
64 |
let |
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
65 |
val (c, thm) = prep_cong raw_thm; |
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
66 |
val _ = if AList.defined (op =) congs c |
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
67 |
then warning ("Overwriting recdef congruence rule for " ^ quote c) |
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
68 |
else (); |
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
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:
21078
diff
changeset
|
74 |
val _ = if AList.defined (op =) congs c |
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
75 |
then () |
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
76 |
else warning ("No recdef congruence rule for " ^ quote c); |
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
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:
35690
diff
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:
22846
diff
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:
33643
diff
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:
22846
diff
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:
22846
diff
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:
22846
diff
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:
22846
diff
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:
22846
diff
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 = |
|
18728 | 146 |
[Args.$$$ recdef_simpN -- Args.colon >> K ((I, simp_add): Method.modifier), |
147 |
Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (I, simp_add), |
|
148 |
Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (I, simp_del), |
|
149 |
Args.$$$ recdef_congN -- Args.colon >> K (I, cong_add), |
|
150 |
Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (I, cong_add), |
|
151 |
Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (I, cong_del), |
|
152 |
Args.$$$ recdef_wfN -- Args.colon >> K (I, wf_add), |
|
153 |
Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (I, wf_add), |
|
154 |
Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (I, wf_del)] @ |
|
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 |
31243 | 167 |
| SOME src => #2 (Method.syntax (Method.sections recdef_modifiers) src ctxt0)); |
21505 | 168 |
val {simps, congs, wfs} = get_hints ctxt; |
42775
a973f82fc885
prefer Proof.context over old-style claset/simpset;
wenzelm
parents:
42361
diff
changeset
|
169 |
val ctxt' = ctxt |
46185
afda84cd4d4b
refer to imp_cong from HOL, not Drule (!) -- cf. 4ed94d92ae19;
wenzelm
parents:
45620
diff
changeset
|
170 |
|> Simplifier.map_simpset (fn ss => ss addsimps simps |> Simplifier.del_cong @{thm imp_cong}); |
42775
a973f82fc885
prefer Proof.context over old-style claset/simpset;
wenzelm
parents:
42361
diff
changeset
|
171 |
in (ctxt', rev (map snd congs), wfs) end; |
9859 | 172 |
|
173 |
fun prepare_hints_i thy () = |
|
15032 | 174 |
let |
42775
a973f82fc885
prefer Proof.context over old-style claset/simpset;
wenzelm
parents:
42361
diff
changeset
|
175 |
val ctxt = Proof_Context.init_global thy; |
15032 | 176 |
val {simps, congs, wfs} = get_global_hints thy; |
42775
a973f82fc885
prefer Proof.context over old-style claset/simpset;
wenzelm
parents:
42361
diff
changeset
|
177 |
val ctxt' = ctxt |
46185
afda84cd4d4b
refer to imp_cong from HOL, not Drule (!) -- cf. 4ed94d92ae19;
wenzelm
parents:
45620
diff
changeset
|
178 |
|> Simplifier.map_simpset (fn ss => ss addsimps simps |> Simplifier.del_cong @{thm imp_cong}); |
42775
a973f82fc885
prefer Proof.context over old-style claset/simpset;
wenzelm
parents:
42361
diff
changeset
|
179 |
in (ctxt', rev (map snd congs), wfs) end; |
9859 | 180 |
|
6439 | 181 |
|
182 |
||
183 |
(** add_recdef(_i) **) |
|
184 |
||
44013
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents:
43278
diff
changeset
|
185 |
fun requires_recdef thy = Theory.requires thy "Old_Recdef" "recursive functions"; |
6557 | 186 |
|
17920 | 187 |
fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy = |
6439 | 188 |
let |
36865 | 189 |
val _ = legacy_feature "Old 'recdef' command -- use 'fun' or 'function' instead"; |
9859 | 190 |
val _ = requires_recdef thy; |
191 |
||
16458 | 192 |
val name = Sign.intern_const thy raw_name; |
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
193 |
val bname = Long_Name.base_name name; |
26478 | 194 |
val _ = writeln ("Defining recursive function " ^ quote name ^ " ..."); |
6429 | 195 |
|
8657 | 196 |
val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs); |
197 |
val eq_atts = map (map (prep_att thy)) raw_eq_atts; |
|
198 |
||
42775
a973f82fc885
prefer Proof.context over old-style claset/simpset;
wenzelm
parents:
42361
diff
changeset
|
199 |
val (ctxt, congs, wfs) = prep_hints thy hints; |
14241
dfae7eb2830c
Prevent recdef from looping when the inductio rule is simplified
paulson
parents:
12876
diff
changeset
|
200 |
(*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
|
201 |
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
|
202 |
give rise to looping conditional rewriting.*) |
42775
a973f82fc885
prefer Proof.context over old-style claset/simpset;
wenzelm
parents:
42361
diff
changeset
|
203 |
val ({lhs, rules = rules_idx, induct, tcs}, thy) = |
a973f82fc885
prefer Proof.context over old-style claset/simpset;
wenzelm
parents:
42361
diff
changeset
|
204 |
tfl_fn not_permissive ctxt congs wfs name R eqs thy; |
21098
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
haftmann
parents:
21078
diff
changeset
|
205 |
val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx); |
33552 | 206 |
val simp_att = |
207 |
if null tcs then [Simplifier.simp_add, Nitpick_Simps.add, Code.add_default_eqn_attribute] |
|
208 |
else []; |
|
18377 | 209 |
val ((simps' :: rules', [induct']), thy) = |
7798
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7262
diff
changeset
|
210 |
thy |
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24624
diff
changeset
|
211 |
|> 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:
36960
diff
changeset
|
212 |
|> Global_Theory.add_thmss |
32952 | 213 |
(((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:
36960
diff
changeset
|
214 |
||>> 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:
35690
diff
changeset
|
215 |
||> 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:
35690
diff
changeset
|
216 |
val result = {lhs = lhs, simps = simps', rules = rules', induct = induct', tcs = tcs}; |
7798
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7262
diff
changeset
|
217 |
val thy = |
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7262
diff
changeset
|
218 |
thy |
6439 | 219 |
|> put_recdef name result |
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24624
diff
changeset
|
220 |
|> Sign.parent_path; |
7798
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7262
diff
changeset
|
221 |
in (thy, result) end; |
6429 | 222 |
|
18728 | 223 |
val add_recdef = gen_add_recdef Tfl.define Attrib.attribute prepare_hints; |
11629 | 224 |
fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w (); |
9859 | 225 |
|
226 |
||
6557 | 227 |
|
228 |
(** defer_recdef(_i) **) |
|
229 |
||
27727
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
wenzelm
parents:
27353
diff
changeset
|
230 |
fun gen_defer_recdef tfl_fn eval_thms raw_name eqs raw_congs thy = |
6557 | 231 |
let |
16458 | 232 |
val name = Sign.intern_const thy raw_name; |
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
233 |
val bname = Long_Name.base_name name; |
6557 | 234 |
|
235 |
val _ = requires_recdef thy; |
|
26478 | 236 |
val _ = writeln ("Deferred recursive function " ^ quote name ^ " ..."); |
6557 | 237 |
|
42361 | 238 |
val congs = eval_thms (Proof_Context.init_global thy) raw_congs; |
42775
a973f82fc885
prefer Proof.context over old-style claset/simpset;
wenzelm
parents:
42361
diff
changeset
|
239 |
val (induct_rules, thy2) = tfl_fn congs name eqs thy; |
18377 | 240 |
val ([induct_rules'], thy3) = |
6557 | 241 |
thy2 |
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24624
diff
changeset
|
242 |
|> 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:
36960
diff
changeset
|
243 |
|> Global_Theory.add_thms [((Binding.name "induct_rules", induct_rules), [])] |
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24624
diff
changeset
|
244 |
||> Sign.parent_path; |
8430 | 245 |
in (thy3, {induct_rules = induct_rules'}) end; |
6557 | 246 |
|
27727
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
wenzelm
parents:
27353
diff
changeset
|
247 |
val defer_recdef = gen_defer_recdef Tfl.defer Attrib.eval_thms; |
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
wenzelm
parents:
27353
diff
changeset
|
248 |
val defer_recdef_i = gen_defer_recdef Tfl.defer_i (K I); |
6557 | 249 |
|
250 |
||
251 |
||
10775 | 252 |
(** recdef_tc(_i) **) |
253 |
||
24457
a33258c78ed2
Adapted to changes in interface of Specification.theorem_i
berghofe
parents:
24039
diff
changeset
|
254 |
fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i int lthy = |
10775 | 255 |
let |
42361 | 256 |
val thy = Proof_Context.theory_of lthy; |
16458 | 257 |
val name = prep_name thy raw_name; |
10775 | 258 |
val atts = map (prep_att thy) raw_atts; |
259 |
val tcs = |
|
260 |
(case get_recdef thy name of |
|
15531 | 261 |
NONE => error ("No recdef definition of constant: " ^ quote name) |
262 |
| SOME {tcs, ...} => tcs); |
|
21351
1fb804b96d7c
recdef_tc(_i): local_theory interface via Specification.theorem_i;
wenzelm
parents:
21098
diff
changeset
|
263 |
val i = the_default 1 opt_i; |
43278 | 264 |
val tc = nth tcs (i - 1) handle General.Subscript => |
10775 | 265 |
error ("No termination condition #" ^ string_of_int i ^ |
266 |
" in recdef definition of " ^ quote name); |
|
21351
1fb804b96d7c
recdef_tc(_i): local_theory interface via Specification.theorem_i;
wenzelm
parents:
21098
diff
changeset
|
267 |
in |
33643
b275f26a638b
eliminated obsolete "internal" kind -- collapsed to unspecific "";
wenzelm
parents:
33552
diff
changeset
|
268 |
Specification.theorem "" NONE (K I) |
33278 | 269 |
(Binding.conceal (Binding.name bname), atts) [] |
270 |
(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:
21098
diff
changeset
|
271 |
end; |
10775 | 272 |
|
21351
1fb804b96d7c
recdef_tc(_i): local_theory interface via Specification.theorem_i;
wenzelm
parents:
21098
diff
changeset
|
273 |
val recdef_tc = gen_recdef_tc Attrib.intern_src Sign.intern_const; |
10775 | 274 |
val recdef_tc_i = gen_recdef_tc (K I) (K I); |
275 |
||
276 |
||
277 |
||
6439 | 278 |
(** package setup **) |
279 |
||
280 |
(* setup theory *) |
|
281 |
||
9859 | 282 |
val setup = |
30528 | 283 |
Attrib.setup @{binding recdef_simp} (Attrib.add_del simp_add simp_del) |
284 |
"declaration of recdef simp rule" #> |
|
285 |
Attrib.setup @{binding recdef_cong} (Attrib.add_del cong_add cong_del) |
|
286 |
"declaration of recdef cong rule" #> |
|
287 |
Attrib.setup @{binding recdef_wf} (Attrib.add_del wf_add wf_del) |
|
288 |
"declaration of recdef wf rule"; |
|
6439 | 289 |
|
290 |
||
6429 | 291 |
(* outer syntax *) |
292 |
||
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
293 |
val _ = List.app Keyword.keyword ["permissive", "congs", "hints"]; |
24867 | 294 |
|
9859 | 295 |
val hints = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
296 |
Parse.$$$ "(" |-- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
297 |
Parse.!!! (Parse.position (Parse.$$$ "hints" -- Args.parse) --| Parse.$$$ ")") >> Args.src; |
9859 | 298 |
|
6429 | 299 |
val recdef_decl = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
300 |
Scan.optional |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
301 |
(Parse.$$$ "(" -- Parse.!!! (Parse.$$$ "permissive" -- Parse.$$$ ")") >> K false) true -- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
302 |
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:
27809
diff
changeset
|
303 |
-- Scan.option hints |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
304 |
>> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map Parse.triple_swap eqs) src); |
6429 | 305 |
|
24867 | 306 |
val _ = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
307 |
Outer_Syntax.command "recdef" "define general recursive functions (TFL)" Keyword.thy_decl |
6429 | 308 |
(recdef_decl >> Toplevel.theory); |
309 |
||
6557 | 310 |
|
311 |
val defer_recdef_decl = |
|
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
312 |
Parse.name -- Scan.repeat1 Parse.prop -- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
313 |
Scan.optional |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
314 |
(Parse.$$$ "(" |-- Parse.$$$ "congs" |-- Parse.!!! (Parse_Spec.xthms1 --| Parse.$$$ ")")) [] |
6557 | 315 |
>> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs); |
316 |
||
24867 | 317 |
val _ = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
318 |
Outer_Syntax.command "defer_recdef" "defer general recursive functions (TFL)" Keyword.thy_decl |
6557 | 319 |
(defer_recdef_decl >> Toplevel.theory); |
320 |
||
24867 | 321 |
val _ = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
322 |
Outer_Syntax.local_theory_to_proof' "recdef_tc" "recommence proof of termination condition (TFL)" |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
323 |
Keyword.thy_goal |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
324 |
((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.xname -- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
325 |
Scan.option (Parse.$$$ "(" |-- Parse.nat --| Parse.$$$ ")") |
26988
742e26213212
more uniform treatment of OuterSyntax.local_theory commands;
wenzelm
parents:
26478
diff
changeset
|
326 |
>> (fn ((thm_name, name), i) => recdef_tc thm_name name i)); |
10775 | 327 |
|
6429 | 328 |
end; |