author | Fabian Huch <huch@in.tum.de> |
Tue, 07 Jun 2022 17:20:25 +0200 | |
changeset 75521 | 7a289e681454 |
parent 71214 | 5727bcc3c47c |
child 78095 | bc42c074e58f |
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 |
67638 | 9 |
type info = Function_Common.info |
10 |
||
36520 | 11 |
val add_function: (binding * typ option * mixfix) list -> |
63064
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
63019
diff
changeset
|
12 |
Specification.multi_specs -> Function_Common.function_config -> |
67638 | 13 |
(Proof.context -> tactic) -> local_theory -> info * local_theory |
36520 | 14 |
|
15 |
val add_function_cmd: (binding * string option * mixfix) list -> |
|
63064
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
63019
diff
changeset
|
16 |
Specification.multi_specs_cmd -> Function_Common.function_config -> |
67638 | 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 -> |
63064
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
63019
diff
changeset
|
20 |
Specification.multi_specs -> Function_Common.function_config -> |
34232
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 -> |
63064
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
63019
diff
changeset
|
24 |
Specification.multi_specs_cmd -> 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 |
|
53603
59ef06cda7b9
generate elim rules for elimination of function equalities;
Manuel Eberl
parents:
52384
diff
changeset
|
27 |
val prove_termination: term option -> tactic -> local_theory -> |
67638 | 28 |
info * local_theory |
36547
2a9d0ec8c10d
return updated info record after termination proof
krauss
parents:
36522
diff
changeset
|
29 |
val prove_termination_cmd: string option -> tactic -> local_theory -> |
67638 | 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 get_congs : Proof.context -> thm list |
34230 | 36 |
|
67638 | 37 |
val get_info : Proof.context -> term -> info |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
38 |
end |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
39 |
|
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 |
structure Function : FUNCTION = |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
42 |
struct |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
43 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
44 |
open Function_Lib |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
45 |
open Function_Common |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
46 |
|
45592 | 47 |
val simp_attribs = |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
65387
diff
changeset
|
48 |
@{attributes [simp, nitpick_simp]} |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
49 |
|
45592 | 50 |
val psimp_attribs = |
51 |
@{attributes [nitpick_psimp]} |
|
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
52 |
|
63019 | 53 |
fun note_derived (a, atts) (fname, thms) = |
54 |
Local_Theory.note ((derived_name fname a, atts), thms) #> apfst snd |
|
53604 | 55 |
|
63005 | 56 |
fun add_simps fnames post sort extra_qualify label mod_binding moreatts simps lthy = |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
57 |
let |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
58 |
val spec = post simps |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
59 |
|> map (apfst (apsnd (fn ats => moreatts @ ats))) |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
60 |
|> map (apfst (apfst extra_qualify)) |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
61 |
|
70611 | 62 |
val (saved_spec_simps, lthy') = |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
63 |
fold_map Local_Theory.note spec lthy |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
64 |
|
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
65 |
val saved_simps = maps snd saved_spec_simps |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
66 |
val simps_by_f = sort saved_simps |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
67 |
|
63004 | 68 |
fun note fname simps = |
63005 | 69 |
Local_Theory.note ((mod_binding (derived_name fname label), []), simps) #> snd |
70611 | 70 |
in (saved_simps, fold2 note fnames simps_by_f lthy') end |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
71 |
|
62774
cfcb20bbdbd8
reconcile object-logic constraint vs. mixfix constraint;
wenzelm
parents:
61841
diff
changeset
|
72 |
fun prepare_function do_print prep fixspec eqns config lthy = |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
73 |
let |
62996 | 74 |
val ((fixes0, spec0), ctxt') = prep fixspec eqns lthy |
75 |
val fixes = map (apfst (apfst Binding.name_of)) fixes0 |
|
76 |
val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0 |
|
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
77 |
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
|
78 |
|
63004 | 79 |
val fnames = map (fst o fst) fixes0 |
63006 | 80 |
val defname = Binding.conglomerate fnames; |
62996 | 81 |
|
41846
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
krauss
parents:
41417
diff
changeset
|
82 |
val FunctionConfig {partials, default, ...} = config |
41417 | 83 |
val _ = |
44052 | 84 |
if is_some default |
85 |
then legacy_feature "\"function (default)\" -- use 'partial_function' instead" |
|
41417 | 86 |
else () |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
87 |
|
36519
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
krauss
parents:
36518
diff
changeset
|
88 |
val ((goal_state, cont), lthy') = |
63011 | 89 |
Function_Mutual.prepare_function_mutual config defname fixes0 eqs lthy |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
90 |
|
70611 | 91 |
fun afterqed [[proof]] lthy1 = |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
92 |
let |
70611 | 93 |
val result = cont lthy1 (Thm.close_derivation \<^here> proof) |
52384 | 94 |
val FunctionResult {fs, R, dom, psimps, simple_pinducts, |
53603
59ef06cda7b9
generate elim rules for elimination of function equalities;
Manuel Eberl
parents:
52384
diff
changeset
|
95 |
termination, domintros, cases, ...} = result |
59ef06cda7b9
generate elim rules for elimination of function equalities;
Manuel Eberl
parents:
52384
diff
changeset
|
96 |
|
70611 | 97 |
val pelims = Function_Elims.mk_partial_elim_rules lthy1 result |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
98 |
|
59859 | 99 |
val concealed_partial = if partials then I else Binding.concealed |
33394
9c6980f2eb39
conceal "termination" rule, used only by special tools
krauss
parents:
33369
diff
changeset
|
100 |
|
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
101 |
val addsmps = add_simps fnames post sort_cont |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
102 |
|
70611 | 103 |
val (((((psimps', [pinducts']), [termination']), cases'), pelims'), lthy2) = |
104 |
lthy1 |
|
59859 | 105 |
|> addsmps (concealed_partial o Binding.qualify false "partial") |
106 |
"psimps" concealed_partial psimp_attribs psimps |
|
63005 | 107 |
||>> Local_Theory.notes [((concealed_partial (derived_name defname "pinduct"), []), |
50771
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents:
49967
diff
changeset
|
108 |
simple_pinducts |> map (fn th => ([th], |
63019 | 109 |
[Attrib.case_names cnames, Attrib.consumes (1 - Thm.nprems_of th)] @ |
110 |
@{attributes [induct pred]})))] |
|
63004 | 111 |
||>> (apfst snd o |
63005 | 112 |
Local_Theory.note |
113 |
((Binding.concealed (derived_name defname "termination"), []), [termination])) |
|
63019 | 114 |
||>> fold_map (note_derived ("cases", [Attrib.case_names cnames])) |
115 |
(fnames ~~ map single cases) |
|
116 |
||>> fold_map (note_derived ("pelims", [Attrib.consumes 1, Attrib.constraints 1])) |
|
63004 | 117 |
(fnames ~~ pelims) |
53603
59ef06cda7b9
generate elim rules for elimination of function equalities;
Manuel Eberl
parents:
52384
diff
changeset
|
118 |
||> (case domintros of NONE => I | SOME thms => |
63005 | 119 |
Local_Theory.note ((derived_name defname "domintros", []), thms) #> snd) |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
120 |
|
63004 | 121 |
val info = |
122 |
{ add_simps=addsmps, fnames=fnames, case_names=cnames, psimps=psimps', |
|
65387
5dbe02addca5
store totality fact in function info
Lars Hupel <lars.hupel@mytum.de>
parents:
63239
diff
changeset
|
123 |
pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination', totality=NONE, |
53603
59ef06cda7b9
generate elim rules for elimination of function equalities;
Manuel Eberl
parents:
52384
diff
changeset
|
124 |
fs=fs, R=R, dom=dom, defname=defname, is_partial=true, cases=flat cases', |
59ef06cda7b9
generate elim rules for elimination of function equalities;
Manuel Eberl
parents:
52384
diff
changeset
|
125 |
pelims=pelims',elims=NONE} |
34230 | 126 |
|
56932
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56254
diff
changeset
|
127 |
val _ = |
70611 | 128 |
Proof_Display.print_consts do_print (Position.thread_data ()) lthy2 |
56932
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56254
diff
changeset
|
129 |
(K false) (map fst fixes) |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
130 |
in |
44239
47ecd30e018d
less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents:
44192
diff
changeset
|
131 |
(info, |
70611 | 132 |
lthy2 |> Local_Theory.declaration {syntax = false, pervasive = false} |
63004 | 133 |
(fn phi => add_function_data (transform_function_data phi info))) |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
134 |
end |
36520 | 135 |
in |
36519
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
krauss
parents:
36518
diff
changeset
|
136 |
((goal_state, afterqed), lthy') |
36518
a33b986f2e22
function: better separate Isar integration from actual functionality
krauss
parents:
36323
diff
changeset
|
137 |
end |
a33b986f2e22
function: better separate Isar integration from actual functionality
krauss
parents:
36323
diff
changeset
|
138 |
|
62774
cfcb20bbdbd8
reconcile object-logic constraint vs. mixfix constraint;
wenzelm
parents:
61841
diff
changeset
|
139 |
fun gen_add_function do_print prep fixspec eqns config tac lthy = |
36520 | 140 |
let |
141 |
val ((goal_state, afterqed), lthy') = |
|
62774
cfcb20bbdbd8
reconcile object-logic constraint vs. mixfix constraint;
wenzelm
parents:
61841
diff
changeset
|
142 |
prepare_function do_print prep fixspec eqns config lthy |
36520 | 143 |
val pattern_thm = |
144 |
case SINGLE (tac lthy') goal_state of |
|
145 |
NONE => error "pattern completeness and compatibility proof failed" |
|
146 |
| SOME st => Goal.finish lthy' st |
|
147 |
in |
|
148 |
lthy' |
|
149 |
|> afterqed [[pattern_thm]] |
|
150 |
end |
|
151 |
||
63064
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
63019
diff
changeset
|
152 |
val add_function = gen_add_function false Specification.check_multi_specs |
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
63019
diff
changeset
|
153 |
fun add_function_cmd a b c d int = gen_add_function int Specification.read_multi_specs a b c d |
60499 | 154 |
|
62774
cfcb20bbdbd8
reconcile object-logic constraint vs. mixfix constraint;
wenzelm
parents:
61841
diff
changeset
|
155 |
fun gen_function do_print prep fixspec eqns config lthy = |
36518
a33b986f2e22
function: better separate Isar integration from actual functionality
krauss
parents:
36323
diff
changeset
|
156 |
let |
36520 | 157 |
val ((goal_state, afterqed), lthy') = |
62774
cfcb20bbdbd8
reconcile object-logic constraint vs. mixfix constraint;
wenzelm
parents:
61841
diff
changeset
|
158 |
prepare_function do_print prep fixspec eqns config lthy |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
159 |
in |
36518
a33b986f2e22
function: better separate Isar integration from actual functionality
krauss
parents:
36323
diff
changeset
|
160 |
lthy' |
59582 | 161 |
|> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (Thm.concl_of goal_state), [])]] |
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61112
diff
changeset
|
162 |
|> Proof.refine_singleton (Method.primitive_text (K (K goal_state))) |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
163 |
end |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
164 |
|
63064
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
63019
diff
changeset
|
165 |
val function = gen_function false Specification.check_multi_specs |
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
63019
diff
changeset
|
166 |
fun function_cmd a b c int = gen_function int Specification.read_multi_specs a b c |
36520 | 167 |
|
70612
7bf683f3672d
proper positions for 'termination' command (see also 5549e686d6ac);
wenzelm
parents:
70611
diff
changeset
|
168 |
fun prepare_termination_proof prep_binding prep_term raw_term_opt lthy = |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
169 |
let |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
170 |
val term_opt = Option.map (prep_term lthy) raw_term_opt |
48995 | 171 |
val info = |
49967
69774b4f5b8a
recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents:
48995
diff
changeset
|
172 |
(case term_opt of |
69774b4f5b8a
recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents:
48995
diff
changeset
|
173 |
SOME t => |
69774b4f5b8a
recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents:
48995
diff
changeset
|
174 |
(case import_function_data t lthy of |
69774b4f5b8a
recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents:
48995
diff
changeset
|
175 |
SOME info => info |
69774b4f5b8a
recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents:
48995
diff
changeset
|
176 |
| NONE => error ("Not a function: " ^ quote (Syntax.string_of_term lthy t))) |
69774b4f5b8a
recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents:
48995
diff
changeset
|
177 |
| NONE => |
69774b4f5b8a
recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents:
48995
diff
changeset
|
178 |
(case import_last_function lthy of |
69774b4f5b8a
recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents:
48995
diff
changeset
|
179 |
SOME info => info |
69774b4f5b8a
recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents:
48995
diff
changeset
|
180 |
| NONE => error "Not a function")) |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
181 |
|
48995 | 182 |
val { termination, fs, R, add_simps, case_names, psimps, |
53603
59ef06cda7b9
generate elim rules for elimination of function equalities;
Manuel Eberl
parents:
52384
diff
changeset
|
183 |
pinducts, defname, fnames, cases, dom, pelims, ...} = info |
48995 | 184 |
val domT = domain_type (fastype_of R) |
185 |
val goal = HOLogic.mk_Trueprop (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT))) |
|
70611 | 186 |
fun afterqed [[raw_totality]] lthy1 = |
48995 | 187 |
let |
70611 | 188 |
val totality = Thm.close_derivation \<^here> raw_totality |
48995 | 189 |
val remove_domain_condition = |
70611 | 190 |
full_simplify (put_simpset HOL_basic_ss lthy1 |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
50771
diff
changeset
|
191 |
addsimps [totality, @{thm True_implies_equals}]) |
48995 | 192 |
val tsimps = map remove_domain_condition psimps |
193 |
val tinduct = map remove_domain_condition pinducts |
|
53603
59ef06cda7b9
generate elim rules for elimination of function equalities;
Manuel Eberl
parents:
52384
diff
changeset
|
194 |
val telims = map (map remove_domain_condition) pelims |
48995 | 195 |
in |
70611 | 196 |
lthy1 |
70612
7bf683f3672d
proper positions for 'termination' command (see also 5549e686d6ac);
wenzelm
parents:
70611
diff
changeset
|
197 |
|> add_simps prep_binding "simps" prep_binding simp_attribs tsimps |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
65387
diff
changeset
|
198 |
||> Code.declare_default_eqns (map (rpair true) tsimps) |
48995 | 199 |
||>> Local_Theory.note |
70612
7bf683f3672d
proper positions for 'termination' command (see also 5549e686d6ac);
wenzelm
parents:
70611
diff
changeset
|
200 |
((prep_binding (derived_name defname "induct"), [Attrib.case_names case_names]), tinduct) |
63019 | 201 |
||>> fold_map (note_derived ("elims", [Attrib.consumes 1, Attrib.constraints 1])) |
70612
7bf683f3672d
proper positions for 'termination' command (see also 5549e686d6ac);
wenzelm
parents:
70611
diff
changeset
|
202 |
(map prep_binding fnames ~~ telims) |
70611 | 203 |
|-> (fn ((simps,(_,inducts)), elims) => fn lthy2 => |
53603
59ef06cda7b9
generate elim rules for elimination of function equalities;
Manuel Eberl
parents:
52384
diff
changeset
|
204 |
let val info' = { is_partial=false, defname=defname, fnames=fnames, add_simps=add_simps, |
52384 | 205 |
case_names=case_names, fs=fs, R=R, dom=dom, psimps=psimps, pinducts=pinducts, |
65387
5dbe02addca5
store totality fact in function info
Lars Hupel <lars.hupel@mytum.de>
parents:
63239
diff
changeset
|
206 |
simps=SOME simps, inducts=SOME inducts, termination=termination, totality=SOME totality, |
5dbe02addca5
store totality fact in function info
Lars Hupel <lars.hupel@mytum.de>
parents:
63239
diff
changeset
|
207 |
cases=cases, pelims=pelims, elims=SOME elims} |
70612
7bf683f3672d
proper positions for 'termination' command (see also 5549e686d6ac);
wenzelm
parents:
70611
diff
changeset
|
208 |
|> transform_function_data (Morphism.binding_morphism "" prep_binding) |
48995 | 209 |
in |
210 |
(info', |
|
70611 | 211 |
lthy2 |
48995 | 212 |
|> Local_Theory.declaration {syntax = false, pervasive = false} |
63004 | 213 |
(fn phi => add_function_data (transform_function_data phi info')) |
71214
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
wenzelm
parents:
71179
diff
changeset
|
214 |
|> Spec_Rules.add Binding.empty Spec_Rules.equational_recdef fs tsimps) |
48995 | 215 |
end) |
216 |
end |
|
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
217 |
in |
36518
a33b986f2e22
function: better separate Isar integration from actual functionality
krauss
parents:
36323
diff
changeset
|
218 |
(goal, afterqed, termination) |
a33b986f2e22
function: better separate Isar integration from actual functionality
krauss
parents:
36323
diff
changeset
|
219 |
end |
a33b986f2e22
function: better separate Isar integration from actual functionality
krauss
parents:
36323
diff
changeset
|
220 |
|
36520 | 221 |
fun gen_prove_termination prep_term raw_term_opt tac lthy = |
222 |
let |
|
223 |
val (goal, afterqed, termination) = |
|
70609
5549e686d6ac
proper positions for 'termination' command instead of original 'function' command, e.g. relevant for isabelle mmt_import;
wenzelm
parents:
70494
diff
changeset
|
224 |
prepare_termination_proof I prep_term raw_term_opt lthy |
36520 | 225 |
|
226 |
val totality = Goal.prove lthy [] [] goal (K tac) |
|
227 |
in |
|
228 |
afterqed [[totality]] lthy |
|
229 |
end |
|
230 |
||
231 |
val prove_termination = gen_prove_termination Syntax.check_term |
|
232 |
val prove_termination_cmd = gen_prove_termination Syntax.read_term |
|
233 |
||
36519
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
krauss
parents:
36518
diff
changeset
|
234 |
fun gen_termination prep_term raw_term_opt lthy = |
36518
a33b986f2e22
function: better separate Isar integration from actual functionality
krauss
parents:
36323
diff
changeset
|
235 |
let |
70609
5549e686d6ac
proper positions for 'termination' command instead of original 'function' command, e.g. relevant for isabelle mmt_import;
wenzelm
parents:
70494
diff
changeset
|
236 |
val (goal, afterqed, termination) = |
5549e686d6ac
proper positions for 'termination' command instead of original 'function' command, e.g. relevant for isabelle mmt_import;
wenzelm
parents:
70494
diff
changeset
|
237 |
prepare_termination_proof Binding.reset_pos prep_term raw_term_opt lthy |
36518
a33b986f2e22
function: better separate Isar integration from actual functionality
krauss
parents:
36323
diff
changeset
|
238 |
in |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
239 |
lthy |
67713 | 240 |
|> Proof_Context.note_thms "" |
241 |
((Binding.empty, [Context_Rules.rule_del]), [([allI], [])]) |> snd |
|
242 |
|> Proof_Context.note_thms "" |
|
243 |
((Binding.empty, [Context_Rules.intro_bang (SOME 1)]), [([allI], [])]) |> snd |
|
244 |
|> Proof_Context.note_thms "" |
|
245 |
((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]), |
|
246 |
[([Goal.norm_result lthy termination], [])]) |> snd |
|
36547
2a9d0ec8c10d
return updated info record after termination proof
krauss
parents:
36522
diff
changeset
|
247 |
|> Proof.theorem NONE (snd oo afterqed) [[(goal, [])]] |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
248 |
end |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
249 |
|
36519
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
krauss
parents:
36518
diff
changeset
|
250 |
val termination = gen_termination Syntax.check_term |
46bf776a81e0
ML interface uses plain command names, following conventions from typedef
krauss
parents:
36518
diff
changeset
|
251 |
val termination_cmd = gen_termination Syntax.read_term |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
252 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
253 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
254 |
(* Datatype hook to declare datatype congs as "function_congs" *) |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
255 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
256 |
fun add_case_cong n thy = |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
257 |
let |
58112
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents:
57959
diff
changeset
|
258 |
val cong = #case_cong (Old_Datatype_Data.the_info thy n) |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
259 |
|> safe_mk_meta_eq |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
260 |
in |
61112 | 261 |
Context.theory_map (Function_Context_Tree.add_function_cong cong) thy |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
34231
diff
changeset
|
262 |
end |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
263 |
|
58826 | 264 |
val _ = Theory.setup (Old_Datatype_Data.interpretation (K (fold add_case_cong))) |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
265 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
266 |
|
58826 | 267 |
(* get info *) |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
268 |
|
58816 | 269 |
val get_congs = Function_Context_Tree.get_function_congs |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
270 |
|
67634 | 271 |
fun get_info ctxt t = Function_Common.retrieve_function_data ctxt t |
34230 | 272 |
|> the_single |> snd |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
273 |
|
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36547
diff
changeset
|
274 |
|
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
275 |
(* outer syntax *) |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
276 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
277 |
val _ = |
67149 | 278 |
Outer_Syntax.local_theory_to_proof' \<^command_keyword>\<open>function\<close> |
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
45592
diff
changeset
|
279 |
"define general recursive functions" |
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
45592
diff
changeset
|
280 |
(function_parser default_config |
63183 | 281 |
>> (fn (config, (fixes, specs)) => function_cmd fixes specs config)) |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
282 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
283 |
val _ = |
67149 | 284 |
Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>termination\<close> |
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
45592
diff
changeset
|
285 |
"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
|
286 |
(Scan.option Parse.term >> termination_cmd) |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
287 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
288 |
end |