| author | blanchet | 
| Wed, 05 Sep 2012 11:16:34 +0200 | |
| changeset 49151 | ff86a2253f05 | 
| parent 48995 | 0e1cab4a334e | 
| child 49967 | 69774b4f5b8a | 
| permissions | -rw-r--r-- | 
| 37744 | 1  | 
(* Title: HOL/Tools/Function/function.ML  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
2  | 
Author: Alexander Krauss, TU Muenchen  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
3  | 
|
| 41114 | 4  | 
Main entry points to the function package.  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
7  | 
signature FUNCTION =  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
9  | 
include FUNCTION_DATA  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
10  | 
|
| 36520 | 11  | 
val add_function: (binding * typ option * mixfix) list ->  | 
12  | 
(Attrib.binding * term) list -> Function_Common.function_config ->  | 
|
| 
36522
 
e80a95279ef6
return info record (relative to auxiliary context!)
 
krauss 
parents: 
36520 
diff
changeset
 | 
13  | 
(Proof.context -> tactic) -> local_theory -> info * local_theory  | 
| 36520 | 14  | 
|
15  | 
val add_function_cmd: (binding * string option * mixfix) list ->  | 
|
16  | 
(Attrib.binding * string) list -> Function_Common.function_config ->  | 
|
| 
44239
 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
17  | 
(Proof.context -> tactic) -> bool -> local_theory -> info * local_theory  | 
| 36520 | 18  | 
|
| 
36519
 
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
 
krauss 
parents: 
36518 
diff
changeset
 | 
19  | 
val function: (binding * typ option * mixfix) list ->  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
20  | 
(Attrib.binding * term) list -> Function_Common.function_config ->  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
21  | 
local_theory -> Proof.state  | 
| 34230 | 22  | 
|
| 
36519
 
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
 
krauss 
parents: 
36518 
diff
changeset
 | 
23  | 
val function_cmd: (binding * string option * mixfix) list ->  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
24  | 
(Attrib.binding * string) list -> Function_Common.function_config ->  | 
| 
44239
 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
25  | 
bool -> local_theory -> Proof.state  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
26  | 
|
| 
36547
 
2a9d0ec8c10d
return updated info record after termination proof
 
krauss 
parents: 
36522 
diff
changeset
 | 
27  | 
val prove_termination: term option -> tactic -> local_theory ->  | 
| 
 
2a9d0ec8c10d
return updated info record after termination proof
 
krauss 
parents: 
36522 
diff
changeset
 | 
28  | 
info * local_theory  | 
| 
 
2a9d0ec8c10d
return updated info record after termination proof
 
krauss 
parents: 
36522 
diff
changeset
 | 
29  | 
val prove_termination_cmd: string option -> tactic -> local_theory ->  | 
| 
 
2a9d0ec8c10d
return updated info record after termination proof
 
krauss 
parents: 
36522 
diff
changeset
 | 
30  | 
info * local_theory  | 
| 36520 | 31  | 
|
| 
36519
 
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
 
krauss 
parents: 
36518 
diff
changeset
 | 
32  | 
val termination : term option -> local_theory -> Proof.state  | 
| 
 
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
 
krauss 
parents: 
36518 
diff
changeset
 | 
33  | 
val termination_cmd : string option -> local_theory -> Proof.state  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
34  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
35  | 
val setup : theory -> theory  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
36  | 
val get_congs : Proof.context -> thm list  | 
| 34230 | 37  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
38  | 
val get_info : Proof.context -> term -> info  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
39  | 
end  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
40  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
42  | 
structure Function : FUNCTION =  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
43  | 
struct  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
44  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
45  | 
open Function_Lib  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
46  | 
open Function_Common  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
47  | 
|
| 45592 | 48  | 
val simp_attribs =  | 
49  | 
  @{attributes [simp, nitpick_simp]} @ [Attrib.internal (K Code.add_default_eqn_attribute)]
 | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
50  | 
|
| 45592 | 51  | 
val psimp_attribs =  | 
52  | 
  @{attributes [nitpick_psimp]}
 | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
53  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
54  | 
fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
55  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
56  | 
fun add_simps fnames post sort extra_qualify label mod_binding moreatts  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
57  | 
simps lthy =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
58  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
59  | 
val spec = post simps  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
60  | 
|> map (apfst (apsnd (fn ats => moreatts @ ats)))  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
61  | 
|> map (apfst (apfst extra_qualify))  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
62  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
63  | 
val (saved_spec_simps, lthy) =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
64  | 
fold_map Local_Theory.note spec lthy  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
65  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
66  | 
val saved_simps = maps snd saved_spec_simps  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
67  | 
val simps_by_f = sort saved_simps  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
68  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
69  | 
fun add_for_f fname simps =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
70  | 
Local_Theory.note  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
71  | 
((mod_binding (Binding.qualify true fname (Binding.name label)), []), simps)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
72  | 
#> snd  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
73  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
74  | 
(saved_simps, fold2 add_for_f fnames simps_by_f lthy)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
75  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
76  | 
|
| 
44239
 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
77  | 
fun prepare_function do_print prep default_constraint fixspec eqns config lthy =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
78  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
79  | 
val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
80  | 
val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
81  | 
val fixes = map (apfst (apfst Binding.name_of)) fixes0;  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
82  | 
val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
83  | 
val (eqs, post, sort_cont, cnames) = get_preproc lthy config ctxt' fixes spec  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
84  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
85  | 
val defname = mk_defname fixes  | 
| 
41846
 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 
krauss 
parents: 
41417 
diff
changeset
 | 
86  | 
    val FunctionConfig {partials, default, ...} = config
 | 
| 41417 | 87  | 
val _ =  | 
| 44052 | 88  | 
if is_some default  | 
89  | 
then legacy_feature "\"function (default)\" -- use 'partial_function' instead"  | 
|
| 41417 | 90  | 
else ()  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
91  | 
|
| 
36519
 
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
 
krauss 
parents: 
36518 
diff
changeset
 | 
92  | 
val ((goal_state, cont), lthy') =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
93  | 
Function_Mutual.prepare_function_mutual config defname fixes eqs lthy  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
94  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
95  | 
fun afterqed [[proof]] lthy =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
96  | 
let  | 
| 
41846
 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 
krauss 
parents: 
41417 
diff
changeset
 | 
97  | 
        val FunctionResult {fs, R, psimps, simple_pinducts,
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
98  | 
termination, domintros, cases, ...} =  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
99  | 
cont (Thm.close_derivation proof)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
100  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
101  | 
val fnames = map (fst o fst) fixes  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
102  | 
fun qualify n = Binding.name n  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
103  | 
|> Binding.qualify true defname  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
104  | 
val conceal_partial = if partials then I else Binding.conceal  | 
| 
33394
 
9c6980f2eb39
conceal "termination" rule, used only by special tools
 
krauss 
parents: 
33369 
diff
changeset
 | 
105  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
106  | 
val addsmps = add_simps fnames post sort_cont  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
107  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
108  | 
val (((psimps', pinducts'), (_, [termination'])), lthy) =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
109  | 
lthy  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
110  | 
|> addsmps (conceal_partial o Binding.qualify false "partial")  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
111  | 
"psimps" conceal_partial psimp_attribs psimps  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
112  | 
||>> Local_Theory.note ((conceal_partial (qualify "pinduct"),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
113  | 
[Attrib.internal (K (Rule_Cases.case_names cnames)),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
114  | 
Attrib.internal (K (Rule_Cases.consumes 1)),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
115  | 
Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
116  | 
||>> Local_Theory.note ((Binding.conceal (qualify "termination"), []), [termination])  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
117  | 
||> (snd o Local_Theory.note ((qualify "cases",  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
118  | 
[Attrib.internal (K (Rule_Cases.case_names cnames))]), [cases]))  | 
| 40076 | 119  | 
||> (case domintros of NONE => I | SOME thms =>  | 
120  | 
Local_Theory.note ((qualify "domintros", []), thms) #> snd)  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
121  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
122  | 
        val info = { add_simps=addsmps, case_names=cnames, psimps=psimps',
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
123  | 
pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination',  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
124  | 
fs=fs, R=R, defname=defname, is_partial=true }  | 
| 34230 | 125  | 
|
| 
44239
 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
126  | 
val _ = Proof_Display.print_consts do_print lthy (K false) (map fst fixes)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
127  | 
in  | 
| 
44239
 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
128  | 
(info,  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
45290 
diff
changeset
 | 
129  | 
         lthy |> Local_Theory.declaration {syntax = false, pervasive = false}
 | 
| 
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
45290 
diff
changeset
 | 
130  | 
(add_function_data o transform_function_data info))  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
131  | 
end  | 
| 36520 | 132  | 
in  | 
| 
36519
 
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
 
krauss 
parents: 
36518 
diff
changeset
 | 
133  | 
((goal_state, afterqed), lthy')  | 
| 
36518
 
a33b986f2e22
function: better separate Isar integration from actual functionality
 
krauss 
parents: 
36323 
diff
changeset
 | 
134  | 
end  | 
| 
 
a33b986f2e22
function: better separate Isar integration from actual functionality
 
krauss 
parents: 
36323 
diff
changeset
 | 
135  | 
|
| 
44239
 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
136  | 
fun gen_add_function do_print prep default_constraint fixspec eqns config tac lthy =  | 
| 36520 | 137  | 
let  | 
138  | 
val ((goal_state, afterqed), lthy') =  | 
|
| 
44239
 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
139  | 
prepare_function do_print prep default_constraint fixspec eqns config lthy  | 
| 36520 | 140  | 
val pattern_thm =  | 
141  | 
case SINGLE (tac lthy') goal_state of  | 
|
142  | 
NONE => error "pattern completeness and compatibility proof failed"  | 
|
143  | 
| SOME st => Goal.finish lthy' st  | 
|
144  | 
in  | 
|
145  | 
lthy'  | 
|
146  | 
|> afterqed [[pattern_thm]]  | 
|
147  | 
end  | 
|
148  | 
||
149  | 
val add_function =  | 
|
| 
37145
 
01aa36932739
renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
 
wenzelm 
parents: 
36960 
diff
changeset
 | 
150  | 
gen_add_function false Specification.check_spec (Type_Infer.anyT HOLogic.typeS)  | 
| 
44239
 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
151  | 
fun add_function_cmd a b c d int = gen_add_function int Specification.read_spec "_::type" a b c d  | 
| 36520 | 152  | 
|
| 
44239
 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
153  | 
fun gen_function do_print prep default_constraint fixspec eqns config lthy =  | 
| 
36518
 
a33b986f2e22
function: better separate Isar integration from actual functionality
 
krauss 
parents: 
36323 
diff
changeset
 | 
154  | 
let  | 
| 36520 | 155  | 
val ((goal_state, afterqed), lthy') =  | 
| 
44239
 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
156  | 
prepare_function do_print prep default_constraint fixspec eqns config lthy  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
157  | 
in  | 
| 
36518
 
a33b986f2e22
function: better separate Isar integration from actual functionality
 
krauss 
parents: 
36323 
diff
changeset
 | 
158  | 
lthy'  | 
| 
36522
 
e80a95279ef6
return info record (relative to auxiliary context!)
 
krauss 
parents: 
36520 
diff
changeset
 | 
159  | 
|> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (concl_of goal_state), [])]]  | 
| 
36519
 
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
 
krauss 
parents: 
36518 
diff
changeset
 | 
160  | 
|> Proof.refine (Method.primitive_text (K goal_state)) |> Seq.hd  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
161  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
162  | 
|
| 
36519
 
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
 
krauss 
parents: 
36518 
diff
changeset
 | 
163  | 
val function =  | 
| 
37145
 
01aa36932739
renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
 
wenzelm 
parents: 
36960 
diff
changeset
 | 
164  | 
gen_function false Specification.check_spec (Type_Infer.anyT HOLogic.typeS)  | 
| 
44239
 
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
165  | 
fun function_cmd a b c int = gen_function int Specification.read_spec "_::type" a b c  | 
| 36520 | 166  | 
|
| 
36518
 
a33b986f2e22
function: better separate Isar integration from actual functionality
 
krauss 
parents: 
36323 
diff
changeset
 | 
167  | 
fun prepare_termination_proof prep_term raw_term_opt lthy =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
168  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
169  | 
val term_opt = Option.map (prep_term lthy) raw_term_opt  | 
| 48995 | 170  | 
val info =  | 
171  | 
the (case term_opt of  | 
|
172  | 
SOME t => (import_function_data t lthy  | 
|
173  | 
handle Option.Option =>  | 
|
174  | 
                 error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
 | 
|
175  | 
| NONE => (import_last_function lthy handle Option.Option => error "Not a function"))  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
176  | 
|
| 48995 | 177  | 
    val { termination, fs, R, add_simps, case_names, psimps,
 | 
178  | 
pinducts, defname, ...} = info  | 
|
179  | 
val domT = domain_type (fastype_of R)  | 
|
180  | 
    val goal = HOLogic.mk_Trueprop (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
 | 
|
181  | 
fun afterqed [[totality]] lthy =  | 
|
182  | 
let  | 
|
183  | 
val totality = Thm.close_derivation totality  | 
|
184  | 
val remove_domain_condition =  | 
|
185  | 
          full_simplify (HOL_basic_ss addsimps [totality, @{thm True_implies_equals}])
 | 
|
186  | 
val tsimps = map remove_domain_condition psimps  | 
|
187  | 
val tinduct = map remove_domain_condition pinducts  | 
|
| 34230 | 188  | 
|
| 48995 | 189  | 
fun qualify n = Binding.name n  | 
190  | 
|> Binding.qualify true defname  | 
|
191  | 
in  | 
|
192  | 
lthy  | 
|
193  | 
|> add_simps I "simps" I simp_attribs tsimps  | 
|
194  | 
||>> Local_Theory.note  | 
|
195  | 
((qualify "induct",  | 
|
196  | 
[Attrib.internal (K (Rule_Cases.case_names case_names))]),  | 
|
197  | 
tinduct)  | 
|
198  | 
|-> (fn (simps, (_, inducts)) => fn lthy =>  | 
|
199  | 
          let val info' = { is_partial=false, defname=defname, add_simps=add_simps,
 | 
|
200  | 
case_names=case_names, fs=fs, R=R, psimps=psimps, pinducts=pinducts,  | 
|
201  | 
simps=SOME simps, inducts=SOME inducts, termination=termination }  | 
|
202  | 
in  | 
|
203  | 
(info',  | 
|
204  | 
lthy  | 
|
205  | 
             |> Local_Theory.declaration {syntax = false, pervasive = false}
 | 
|
206  | 
(add_function_data o transform_function_data info')  | 
|
207  | 
|> Spec_Rules.add Spec_Rules.Equational (fs, tsimps))  | 
|
208  | 
end)  | 
|
209  | 
end  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
210  | 
in  | 
| 
36518
 
a33b986f2e22
function: better separate Isar integration from actual functionality
 
krauss 
parents: 
36323 
diff
changeset
 | 
211  | 
(goal, afterqed, termination)  | 
| 
 
a33b986f2e22
function: better separate Isar integration from actual functionality
 
krauss 
parents: 
36323 
diff
changeset
 | 
212  | 
end  | 
| 
 
a33b986f2e22
function: better separate Isar integration from actual functionality
 
krauss 
parents: 
36323 
diff
changeset
 | 
213  | 
|
| 36520 | 214  | 
fun gen_prove_termination prep_term raw_term_opt tac lthy =  | 
215  | 
let  | 
|
216  | 
val (goal, afterqed, termination) =  | 
|
217  | 
prepare_termination_proof prep_term raw_term_opt lthy  | 
|
218  | 
||
219  | 
val totality = Goal.prove lthy [] [] goal (K tac)  | 
|
220  | 
in  | 
|
221  | 
afterqed [[totality]] lthy  | 
|
222  | 
end  | 
|
223  | 
||
224  | 
val prove_termination = gen_prove_termination Syntax.check_term  | 
|
225  | 
val prove_termination_cmd = gen_prove_termination Syntax.read_term  | 
|
226  | 
||
| 
36519
 
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
 
krauss 
parents: 
36518 
diff
changeset
 | 
227  | 
fun gen_termination prep_term raw_term_opt lthy =  | 
| 
36518
 
a33b986f2e22
function: better separate Isar integration from actual functionality
 
krauss 
parents: 
36323 
diff
changeset
 | 
228  | 
let  | 
| 
 
a33b986f2e22
function: better separate Isar integration from actual functionality
 
krauss 
parents: 
36323 
diff
changeset
 | 
229  | 
val (goal, afterqed, termination) = prepare_termination_proof prep_term raw_term_opt lthy  | 
| 
 
a33b986f2e22
function: better separate Isar integration from actual functionality
 
krauss 
parents: 
36323 
diff
changeset
 | 
230  | 
in  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
231  | 
lthy  | 
| 42361 | 232  | 
|> Proof_Context.note_thmss ""  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
233  | 
[((Binding.empty, [Context_Rules.rule_del]), [([allI], [])])] |> snd  | 
| 42361 | 234  | 
|> Proof_Context.note_thmss ""  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
235  | 
[((Binding.empty, [Context_Rules.intro_bang (SOME 1)]), [([allI], [])])] |> snd  | 
| 42361 | 236  | 
|> Proof_Context.note_thmss ""  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
237  | 
[((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
238  | 
[([Goal.norm_result termination], [])])] |> snd  | 
| 
36547
 
2a9d0ec8c10d
return updated info record after termination proof
 
krauss 
parents: 
36522 
diff
changeset
 | 
239  | 
|> Proof.theorem NONE (snd oo afterqed) [[(goal, [])]]  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
240  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
241  | 
|
| 
36519
 
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
 
krauss 
parents: 
36518 
diff
changeset
 | 
242  | 
val termination = gen_termination Syntax.check_term  | 
| 
 
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
 
krauss 
parents: 
36518 
diff
changeset
 | 
243  | 
val termination_cmd = gen_termination Syntax.read_term  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
244  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
245  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
246  | 
(* Datatype hook to declare datatype congs as "function_congs" *)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
247  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
248  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
249  | 
fun add_case_cong n thy =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
250  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
251  | 
val cong = #case_cong (Datatype.the_info thy n)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
252  | 
|> safe_mk_meta_eq  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
253  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
254  | 
Context.theory_map  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
255  | 
(Function_Ctx_Tree.map_function_congs (Thm.add_thm cong)) thy  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34231 
diff
changeset
 | 
256  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
257  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
258  | 
val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
259  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
260  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
261  | 
(* setup *)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
262  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
263  | 
val setup =  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
264  | 
  Attrib.setup @{binding fundef_cong}
 | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
265  | 
(Attrib.add_del Function_Ctx_Tree.cong_add Function_Ctx_Tree.cong_del)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
266  | 
"declaration of congruence rule for function definitions"  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
267  | 
#> setup_case_cong  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
268  | 
#> Function_Common.Termination_Simps.setup  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
269  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
270  | 
val get_congs = Function_Ctx_Tree.get_function_congs  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
271  | 
|
| 34230 | 272  | 
fun get_info ctxt t = Item_Net.retrieve (get_function ctxt) t  | 
273  | 
|> the_single |> snd  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
274  | 
|
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36547 
diff
changeset
 | 
275  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
276  | 
(* outer syntax *)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
277  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
278  | 
val _ =  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
45592 
diff
changeset
 | 
279  | 
  Outer_Syntax.local_theory_to_proof' @{command_spec "function"}
 | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
45592 
diff
changeset
 | 
280  | 
"define general recursive functions"  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
45592 
diff
changeset
 | 
281  | 
(function_parser default_config  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
45592 
diff
changeset
 | 
282  | 
>> (fn ((config, fixes), statements) => function_cmd fixes statements config))  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
283  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
284  | 
val _ =  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
45592 
diff
changeset
 | 
285  | 
  Outer_Syntax.local_theory_to_proof @{command_spec "termination"}
 | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
45592 
diff
changeset
 | 
286  | 
"prove termination of a recursive function"  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
45592 
diff
changeset
 | 
287  | 
(Scan.option Parse.term >> termination_cmd)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
288  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
289  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
290  | 
end  |