| author | wenzelm | 
| Mon, 16 Nov 2009 13:53:02 +0100 | |
| changeset 33712 | cffc97238102 | 
| parent 33699 | f33b036ef318 | 
| child 34952 | bd7e347eb768 | 
| 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  | 
10  | 
    -> {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  | 
| 9859 | 20  | 
      * {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 ->  | 
| 9859 | 22  | 
    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: 
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  | 
|
| 8657 | 87  | 
type recdef_info = {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  | 
|
163  | 
val ctxt0 = ProofContext.init thy;  | 
|
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;
 | 
| 
32149
 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
169  | 
val cs = claset_of ctxt;  | 
| 
 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
170  | 
val ss = simpset_of ctxt addsimps simps;  | 
| 
21098
 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 
haftmann 
parents: 
21078 
diff
changeset
 | 
171  | 
in (cs, ss, rev (map snd congs), wfs) end;  | 
| 9859 | 172  | 
|
173  | 
fun prepare_hints_i thy () =  | 
|
| 15032 | 174  | 
let  | 
175  | 
val ctxt0 = ProofContext.init thy;  | 
|
176  | 
    val {simps, congs, wfs} = get_global_hints thy;
 | 
|
| 
32149
 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
177  | 
in (claset_of ctxt0, simpset_of ctxt0 addsimps simps, rev (map snd congs), wfs) end;  | 
| 9859 | 178  | 
|
| 6439 | 179  | 
|
180  | 
||
181  | 
(** add_recdef(_i) **)  | 
|
182  | 
||
| 6557 | 183  | 
fun requires_recdef thy = Theory.requires thy "Recdef" "recursive functions";  | 
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  | 
| 9859 | 187  | 
val _ = requires_recdef thy;  | 
188  | 
||
| 16458 | 189  | 
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
 | 
190  | 
val bname = Long_Name.base_name name;  | 
| 26478 | 191  | 
    val _ = writeln ("Defining recursive function " ^ quote name ^ " ...");
 | 
| 6429 | 192  | 
|
| 8657 | 193  | 
val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs);  | 
194  | 
val eq_atts = map (map (prep_att thy)) raw_eq_atts;  | 
|
195  | 
||
| 9859 | 196  | 
val (cs, ss, congs, wfs) = prep_hints thy hints;  | 
| 
14241
 
dfae7eb2830c
Prevent recdef from looping when the inductio rule is simplified
 
paulson 
parents: 
12876 
diff
changeset
 | 
197  | 
(*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
 | 
198  | 
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
 | 
199  | 
give rise to looping conditional rewriting.*)  | 
| 
 
dfae7eb2830c
Prevent recdef from looping when the inductio rule is simplified
 
paulson 
parents: 
12876 
diff
changeset
 | 
200  | 
    val (thy, {rules = rules_idx, induct, tcs}) =
 | 
| 17920 | 201  | 
tfl_fn not_permissive thy cs (ss delcongs [imp_cong])  | 
| 
14241
 
dfae7eb2830c
Prevent recdef from looping when the inductio rule is simplified
 
paulson 
parents: 
12876 
diff
changeset
 | 
202  | 
congs wfs name R eqs;  | 
| 
21098
 
d0d8a48ae4e6
switched merge_alists'' to AList.merge'' whenever appropriate
 
haftmann 
parents: 
21078 
diff
changeset
 | 
203  | 
val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx);  | 
| 33552 | 204  | 
val simp_att =  | 
205  | 
if null tcs then [Simplifier.simp_add, Nitpick_Simps.add, Code.add_default_eqn_attribute]  | 
|
206  | 
else [];  | 
|
| 8657 | 207  | 
|
| 18377 | 208  | 
val ((simps' :: rules', [induct']), thy) =  | 
| 
7798
 
42e94b618f34
return stored thms with proper naming in derivation;
 
wenzelm 
parents: 
7262 
diff
changeset
 | 
209  | 
thy  | 
| 
24712
 
64ed05609568
proper Sign operations instead of Theory aliases;
 
wenzelm 
parents: 
24624 
diff
changeset
 | 
210  | 
|> Sign.add_path bname  | 
| 18688 | 211  | 
|> PureThy.add_thmss  | 
| 32952 | 212  | 
(((Binding.name "simps", flat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))  | 
| 29579 | 213  | 
||>> PureThy.add_thms [((Binding.name "induct", induct), [])];  | 
| 8657 | 214  | 
    val result = {simps = simps', rules = rules', induct = induct', tcs = tcs};
 | 
| 
7798
 
42e94b618f34
return stored thms with proper naming in derivation;
 
wenzelm 
parents: 
7262 
diff
changeset
 | 
215  | 
val thy =  | 
| 
 
42e94b618f34
return stored thms with proper naming in derivation;
 
wenzelm 
parents: 
7262 
diff
changeset
 | 
216  | 
thy  | 
| 6439 | 217  | 
|> put_recdef name result  | 
| 
24712
 
64ed05609568
proper Sign operations instead of Theory aliases;
 
wenzelm 
parents: 
24624 
diff
changeset
 | 
218  | 
|> Sign.parent_path;  | 
| 
7798
 
42e94b618f34
return stored thms with proper naming in derivation;
 
wenzelm 
parents: 
7262 
diff
changeset
 | 
219  | 
in (thy, result) end;  | 
| 6429 | 220  | 
|
| 18728 | 221  | 
val add_recdef = gen_add_recdef Tfl.define Attrib.attribute prepare_hints;  | 
| 11629 | 222  | 
fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w ();  | 
| 9859 | 223  | 
|
224  | 
||
| 6557 | 225  | 
|
226  | 
(** defer_recdef(_i) **)  | 
|
227  | 
||
| 
27727
 
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
 
wenzelm 
parents: 
27353 
diff
changeset
 | 
228  | 
fun gen_defer_recdef tfl_fn eval_thms raw_name eqs raw_congs thy =  | 
| 6557 | 229  | 
let  | 
| 16458 | 230  | 
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
 | 
231  | 
val bname = Long_Name.base_name name;  | 
| 6557 | 232  | 
|
233  | 
val _ = requires_recdef thy;  | 
|
| 26478 | 234  | 
    val _ = writeln ("Deferred recursive function " ^ quote name ^ " ...");
 | 
| 6557 | 235  | 
|
| 
27727
 
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
 
wenzelm 
parents: 
27353 
diff
changeset
 | 
236  | 
val congs = eval_thms (ProofContext.init thy) raw_congs;  | 
| 
 
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
 
wenzelm 
parents: 
27353 
diff
changeset
 | 
237  | 
val (thy2, induct_rules) = tfl_fn thy congs name eqs;  | 
| 18377 | 238  | 
val ([induct_rules'], thy3) =  | 
| 6557 | 239  | 
thy2  | 
| 
24712
 
64ed05609568
proper Sign operations instead of Theory aliases;
 
wenzelm 
parents: 
24624 
diff
changeset
 | 
240  | 
|> Sign.add_path bname  | 
| 29579 | 241  | 
|> PureThy.add_thms [((Binding.name "induct_rules", induct_rules), [])]  | 
| 
24712
 
64ed05609568
proper Sign operations instead of Theory aliases;
 
wenzelm 
parents: 
24624 
diff
changeset
 | 
242  | 
||> Sign.parent_path;  | 
| 8430 | 243  | 
  in (thy3, {induct_rules = induct_rules'}) end;
 | 
| 6557 | 244  | 
|
| 
27727
 
2397e310b2cc
simplified defer_recdef(_i): plain facts via Attrib.eval_thms;
 
wenzelm 
parents: 
27353 
diff
changeset
 | 
245  | 
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
 | 
246  | 
val defer_recdef_i = gen_defer_recdef Tfl.defer_i (K I);  | 
| 6557 | 247  | 
|
248  | 
||
249  | 
||
| 10775 | 250  | 
(** recdef_tc(_i) **)  | 
251  | 
||
| 
24457
 
a33258c78ed2
Adapted to changes in interface of Specification.theorem_i
 
berghofe 
parents: 
24039 
diff
changeset
 | 
252  | 
fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i int lthy =  | 
| 10775 | 253  | 
let  | 
| 
21351
 
1fb804b96d7c
recdef_tc(_i): local_theory interface via Specification.theorem_i;
 
wenzelm 
parents: 
21098 
diff
changeset
 | 
254  | 
val thy = ProofContext.theory_of lthy;  | 
| 16458 | 255  | 
val name = prep_name thy raw_name;  | 
| 10775 | 256  | 
val atts = map (prep_att thy) raw_atts;  | 
257  | 
val tcs =  | 
|
258  | 
(case get_recdef thy name of  | 
|
| 15531 | 259  | 
        NONE => error ("No recdef definition of constant: " ^ quote name)
 | 
260  | 
      | SOME {tcs, ...} => tcs);
 | 
|
| 
21351
 
1fb804b96d7c
recdef_tc(_i): local_theory interface via Specification.theorem_i;
 
wenzelm 
parents: 
21098 
diff
changeset
 | 
261  | 
val i = the_default 1 opt_i;  | 
| 
 
1fb804b96d7c
recdef_tc(_i): local_theory interface via Specification.theorem_i;
 
wenzelm 
parents: 
21098 
diff
changeset
 | 
262  | 
val tc = nth tcs (i - 1) handle Subscript =>  | 
| 10775 | 263  | 
      error ("No termination condition #" ^ string_of_int i ^
 | 
264  | 
" in recdef definition of " ^ quote name);  | 
|
| 
21351
 
1fb804b96d7c
recdef_tc(_i): local_theory interface via Specification.theorem_i;
 
wenzelm 
parents: 
21098 
diff
changeset
 | 
265  | 
in  | 
| 
33643
 
b275f26a638b
eliminated obsolete "internal" kind -- collapsed to unspecific "";
 
wenzelm 
parents: 
33552 
diff
changeset
 | 
266  | 
Specification.theorem "" NONE (K I)  | 
| 33278 | 267  | 
(Binding.conceal (Binding.name bname), atts) []  | 
268  | 
(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
 | 
269  | 
end;  | 
| 10775 | 270  | 
|
| 
21351
 
1fb804b96d7c
recdef_tc(_i): local_theory interface via Specification.theorem_i;
 
wenzelm 
parents: 
21098 
diff
changeset
 | 
271  | 
val recdef_tc = gen_recdef_tc Attrib.intern_src Sign.intern_const;  | 
| 10775 | 272  | 
val recdef_tc_i = gen_recdef_tc (K I) (K I);  | 
273  | 
||
274  | 
||
275  | 
||
| 6439 | 276  | 
(** package setup **)  | 
277  | 
||
278  | 
(* setup theory *)  | 
|
279  | 
||
| 9859 | 280  | 
val setup =  | 
| 30528 | 281  | 
  Attrib.setup @{binding recdef_simp} (Attrib.add_del simp_add simp_del)
 | 
282  | 
"declaration of recdef simp rule" #>  | 
|
283  | 
  Attrib.setup @{binding recdef_cong} (Attrib.add_del cong_add cong_del)
 | 
|
284  | 
"declaration of recdef cong rule" #>  | 
|
285  | 
  Attrib.setup @{binding recdef_wf} (Attrib.add_del wf_add wf_del)
 | 
|
286  | 
"declaration of recdef wf rule";  | 
|
| 6439 | 287  | 
|
288  | 
||
| 6429 | 289  | 
(* outer syntax *)  | 
290  | 
||
| 17057 | 291  | 
local structure P = OuterParse and K = OuterKeyword in  | 
| 6429 | 292  | 
|
| 
27353
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
26988 
diff
changeset
 | 
293  | 
val _ = List.app OuterKeyword.keyword ["permissive", "congs", "hints"];  | 
| 24867 | 294  | 
|
| 9859 | 295  | 
val hints =  | 
| 
27809
 
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27727 
diff
changeset
 | 
296  | 
  P.$$$ "(" |-- P.!!! (P.position (P.$$$ "hints" -- Args.parse) --| P.$$$ ")") >> Args.src;
 | 
| 9859 | 297  | 
|
| 6429 | 298  | 
val recdef_decl =  | 
| 11629 | 299  | 
  Scan.optional (P.$$$ "(" -- P.!!! (P.$$$ "permissive" -- P.$$$ ")") >> K false) true --
 | 
| 29579 | 300  | 
P.name -- P.term -- Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop)  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27809 
diff
changeset
 | 
301  | 
-- 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: 
26478 
diff
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: 
26478 
diff
changeset
 | 
320  | 
K.thy_goal  | 
| 
30223
 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 
wenzelm 
parents: 
29871 
diff
changeset
 | 
321  | 
((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.xname --  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27809 
diff
changeset
 | 
322  | 
        Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")")
 | 
| 
26988
 
742e26213212
more uniform treatment of OuterSyntax.local_theory commands;
 
wenzelm 
parents: 
26478 
diff
changeset
 | 
323  | 
>> (fn ((thm_name, name), i) => recdef_tc thm_name name i));  | 
| 10775 | 324  | 
|
| 6429 | 325  | 
end;  | 
326  | 
||
327  | 
end;  |