krauss@19564
|
1 |
(* Title: HOL/Tools/function_package/fundef_package.ML
|
krauss@19564
|
2 |
ID: $Id$
|
krauss@19564
|
3 |
Author: Alexander Krauss, TU Muenchen
|
krauss@19564
|
4 |
|
wenzelm@20363
|
5 |
A package for general recursive function definitions.
|
krauss@19564
|
6 |
Isar commands.
|
krauss@19564
|
7 |
|
krauss@19564
|
8 |
*)
|
krauss@19564
|
9 |
|
wenzelm@20363
|
10 |
signature FUNDEF_PACKAGE =
|
krauss@19564
|
11 |
sig
|
wenzelm@20877
|
12 |
val add_fundef : (string * string option * mixfix) list
|
krauss@22498
|
13 |
-> ((bstring * Attrib.src list) * (string * bool)) list
|
wenzelm@20877
|
14 |
-> FundefCommon.fundef_config
|
wenzelm@20877
|
15 |
-> local_theory
|
krauss@21211
|
16 |
-> string * Proof.state
|
krauss@20523
|
17 |
|
wenzelm@20877
|
18 |
val add_fundef_i: (string * typ option * mixfix) list
|
krauss@22498
|
19 |
-> ((bstring * Attrib.src list) * (term * bool)) list
|
krauss@22170
|
20 |
-> FundefCommon.fundef_config
|
wenzelm@20877
|
21 |
-> local_theory
|
krauss@21211
|
22 |
-> string * Proof.state
|
krauss@21211
|
23 |
|
krauss@21211
|
24 |
val setup_termination_proof : string option -> local_theory -> Proof.state
|
krauss@19564
|
25 |
|
krauss@19564
|
26 |
val cong_add: attribute
|
krauss@19564
|
27 |
val cong_del: attribute
|
wenzelm@20363
|
28 |
|
krauss@19564
|
29 |
val setup : theory -> theory
|
krauss@21235
|
30 |
val setup_case_cong_hook : theory -> theory
|
krauss@19770
|
31 |
val get_congs : theory -> thm list
|
krauss@19564
|
32 |
end
|
krauss@19564
|
33 |
|
krauss@19564
|
34 |
|
krauss@22170
|
35 |
structure FundefPackage : FUNDEF_PACKAGE =
|
krauss@19564
|
36 |
struct
|
krauss@19564
|
37 |
|
krauss@21051
|
38 |
open FundefLib
|
krauss@19564
|
39 |
open FundefCommon
|
krauss@19564
|
40 |
|
wenzelm@21435
|
41 |
val note_theorem = LocalTheory.note Thm.theoremK
|
wenzelm@21435
|
42 |
|
krauss@22166
|
43 |
fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
|
krauss@22166
|
44 |
|
krauss@20523
|
45 |
fun burrow_snd f ps = (* ('a list -> 'b list) -> ('c * 'a) list -> ('c * 'b) list *)
|
krauss@20523
|
46 |
let val (xs, ys) = split_list ps
|
krauss@20523
|
47 |
in xs ~~ f ys end
|
krauss@20270
|
48 |
|
krauss@20523
|
49 |
fun restore_spec_structure reps spec =
|
krauss@22498
|
50 |
(burrow_snd o burrow o K) reps spec
|
wenzelm@20363
|
51 |
|
krauss@22166
|
52 |
fun add_simps fixes spec sort label moreatts simps lthy =
|
krauss@19564
|
53 |
let
|
krauss@20523
|
54 |
val fnames = map (fst o fst) fixes
|
krauss@21255
|
55 |
|
krauss@22166
|
56 |
val (saved_spec_simps, lthy) =
|
krauss@22498
|
57 |
fold_map note_theorem (restore_spec_structure simps spec) lthy
|
krauss@22498
|
58 |
val saved_simps = flat (map snd saved_spec_simps)
|
krauss@21255
|
59 |
|
krauss@22166
|
60 |
val simps_by_f = sort saved_simps
|
krauss@19564
|
61 |
|
krauss@22166
|
62 |
fun add_for_f fname simps =
|
wenzelm@21435
|
63 |
note_theorem
|
wenzelm@21658
|
64 |
((NameSpace.qualified fname label, Attrib.internal (K Simplifier.simp_add) :: moreatts),
|
krauss@22166
|
65 |
simps) #> snd
|
krauss@19564
|
66 |
in
|
krauss@22166
|
67 |
(saved_simps,
|
krauss@22166
|
68 |
fold2 add_for_f fnames simps_by_f lthy)
|
krauss@19564
|
69 |
end
|
krauss@19564
|
70 |
|
krauss@22166
|
71 |
fun fundef_afterqed config fixes spec defname cont sort_cont [[proof]] lthy =
|
krauss@20523
|
72 |
let
|
krauss@22166
|
73 |
val FundefResult {f, R, psimps, trsimps, subset_pinducts, simple_pinducts, termination, domintros, cases, ...} =
|
krauss@22166
|
74 |
cont (Goal.close_result proof)
|
krauss@22166
|
75 |
|
wenzelm@21391
|
76 |
val qualify = NameSpace.qualified defname
|
krauss@22170
|
77 |
val addsmps = add_simps fixes spec sort_cont
|
wenzelm@21602
|
78 |
|
wenzelm@21602
|
79 |
val (((psimps', pinducts'), (_, [termination'])), lthy) =
|
krauss@21255
|
80 |
lthy
|
krauss@22170
|
81 |
|> addsmps "psimps" [] psimps
|
krauss@22170
|
82 |
||> fold_option (snd oo addsmps "simps" []) trsimps
|
krauss@22166
|
83 |
||>> note_theorem ((qualify "pinduct",
|
krauss@22166
|
84 |
[Attrib.internal (K (InductAttrib.induct_set ""))]), simple_pinducts)
|
krauss@22166
|
85 |
||>> note_theorem ((qualify "termination", []), [termination])
|
krauss@22166
|
86 |
||> (snd o note_theorem ((qualify "cases", []), [cases]))
|
krauss@22166
|
87 |
||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
|
krauss@21255
|
88 |
|
krauss@22170
|
89 |
val cdata = FundefCtxData { add_simps=addsmps, psimps=psimps',
|
krauss@21255
|
90 |
pinducts=snd pinducts', termination=termination', f=f, R=R }
|
wenzelm@22668
|
91 |
val cdata' = cdata |> morph_fundef_data (LocalTheory.target_morphism lthy); (* FIXME !? *)
|
krauss@20523
|
92 |
in
|
krauss@22623
|
93 |
lthy
|
krauss@22635
|
94 |
|> LocalTheory.declaration (fn phi => add_fundef_data defname (morph_fundef_data phi cdata)) (* save in target *)
|
wenzelm@22668
|
95 |
|> Context.proof_map (add_fundef_data defname cdata') (* also save in local context *)
|
krauss@20523
|
96 |
end (* FIXME: Add cases for induct and cases thm *)
|
krauss@20523
|
97 |
|
krauss@20523
|
98 |
|
krauss@20523
|
99 |
|
krauss@20523
|
100 |
fun prep_with_flags prep fixspec eqnss_flags global_flag lthy =
|
krauss@19564
|
101 |
let
|
krauss@22498
|
102 |
val flags = map (fn x => global_flag orelse (snd (snd x))) eqnss_flags
|
krauss@22498
|
103 |
val eqns = map (apsnd (single o fst)) eqnss_flags
|
krauss@20523
|
104 |
|
krauss@20523
|
105 |
val ((fixes, _), ctxt') = prep fixspec [] lthy
|
krauss@22498
|
106 |
|
krauss@22498
|
107 |
fun prep_eqn e = the_single (snd (fst (prep [] [e] ctxt')))
|
krauss@22498
|
108 |
|> apsnd the_single
|
krauss@22498
|
109 |
|
krauss@22498
|
110 |
val spec = map prep_eqn eqns
|
krauss@22498
|
111 |
|> map (apsnd (fn t => fold_rev (mk_forall o Free) (frees_in_term ctxt' t) t)) (* Add quantifiers *)
|
krauss@22498
|
112 |
|> burrow_snd (fn ts => FundefSplit.split_some_equations ctxt' (flags ~~ ts))
|
krauss@20523
|
113 |
in
|
krauss@20523
|
114 |
((fixes, spec), ctxt')
|
krauss@20523
|
115 |
end
|
krauss@20523
|
116 |
|
krauss@20654
|
117 |
fun gen_add_fundef prep_spec fixspec eqnss_flags config lthy =
|
krauss@20523
|
118 |
let
|
krauss@22166
|
119 |
val FundefConfig {sequential, default, tailrec, ...} = config
|
krauss@20654
|
120 |
|
krauss@20654
|
121 |
val ((fixes, spec), ctxt') = prep_with_flags prep_spec fixspec eqnss_flags sequential lthy
|
krauss@22166
|
122 |
|
krauss@22166
|
123 |
val defname = mk_defname fixes
|
krauss@22166
|
124 |
|
krauss@22498
|
125 |
val t_eqns = spec |> map snd |> flat (* flatten external structure *)
|
krauss@20523
|
126 |
|
krauss@22166
|
127 |
val ((goalstate, cont, sort_cont), lthy) =
|
krauss@22166
|
128 |
FundefMutual.prepare_fundef_mutual config defname fixes t_eqns default lthy
|
krauss@20523
|
129 |
|
krauss@22166
|
130 |
val afterqed = fundef_afterqed config fixes spec defname cont sort_cont
|
krauss@20523
|
131 |
in
|
krauss@22166
|
132 |
(defname, lthy
|
krauss@22166
|
133 |
|> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
|
krauss@22166
|
134 |
|> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd)
|
krauss@20523
|
135 |
end
|
krauss@20523
|
136 |
|
krauss@20523
|
137 |
|
krauss@20523
|
138 |
fun total_termination_afterqed defname data [[totality]] lthy =
|
krauss@20523
|
139 |
let
|
krauss@22166
|
140 |
val FundefCtxData { add_simps, psimps, pinducts, ... } = data
|
krauss@21255
|
141 |
|
krauss@22166
|
142 |
val totality = Goal.close_result totality
|
krauss@22725
|
143 |
|> Thm.varifyT (* FIXME ! *)
|
krauss@19564
|
144 |
|
krauss@21255
|
145 |
val remove_domain_condition = full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
|
krauss@19564
|
146 |
|
krauss@22166
|
147 |
val tsimps = map remove_domain_condition psimps
|
krauss@22166
|
148 |
val tinduct = map remove_domain_condition pinducts
|
krauss@19770
|
149 |
|
krauss@21511
|
150 |
(* FIXME: How to generate code from (possibly) local contexts*)
|
krauss@21511
|
151 |
val has_guards = exists ((fn (Const ("Trueprop", _) $ _) => false | _ => true) o prop_of) tsimps
|
wenzelm@21658
|
152 |
val allatts = if has_guards then [] else [Attrib.internal (K (RecfunCodegen.add NONE))]
|
wenzelm@21602
|
153 |
|
krauss@21511
|
154 |
val qualify = NameSpace.qualified defname;
|
krauss@19564
|
155 |
in
|
krauss@21511
|
156 |
lthy
|
krauss@22166
|
157 |
|> add_simps "simps" allatts tsimps |> snd
|
krauss@22166
|
158 |
|> note_theorem ((qualify "induct", []), tinduct) |> snd
|
wenzelm@20363
|
159 |
end
|
krauss@19564
|
160 |
|
krauss@19564
|
161 |
|
krauss@21211
|
162 |
fun setup_termination_proof name_opt lthy =
|
krauss@20523
|
163 |
let
|
krauss@22166
|
164 |
val defname = the_default (get_last_fundef (Context.Proof lthy)) name_opt
|
krauss@22166
|
165 |
val data = the (get_fundef_data defname (Context.Proof lthy))
|
krauss@22166
|
166 |
handle Option.Option => raise ERROR ("No such function definition: " ^ defname)
|
krauss@20523
|
167 |
|
krauss@22166
|
168 |
val FundefCtxData {termination, R, ...} = data
|
krauss@22325
|
169 |
val domT = domain_type (fastype_of R)
|
krauss@22325
|
170 |
val goal = HOLogic.mk_Trueprop (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
|
krauss@22725
|
171 |
|> Type.freeze (* FIXME ! *)
|
krauss@20523
|
172 |
in
|
wenzelm@20877
|
173 |
lthy
|
krauss@22325
|
174 |
|> ProofContext.note_thmss_i "" [(("", [ContextRules.rule_del]), [([allI], [])])] |> snd
|
krauss@22325
|
175 |
|> ProofContext.note_thmss_i "" [(("", [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
|
wenzelm@21602
|
176 |
|> ProofContext.note_thmss_i ""
|
krauss@22325
|
177 |
[(("termination", [ContextRules.intro_bang (SOME 0)]),
|
wenzelm@21602
|
178 |
[([Goal.norm_result termination], [])])] |> snd
|
krauss@21319
|
179 |
|> set_termination_rule termination
|
krauss@22166
|
180 |
|> Proof.theorem_i NONE (total_termination_afterqed defname data) [[(goal, [])]]
|
krauss@20523
|
181 |
end
|
krauss@20523
|
182 |
|
krauss@20523
|
183 |
|
krauss@20523
|
184 |
val add_fundef = gen_add_fundef Specification.read_specification
|
krauss@20523
|
185 |
val add_fundef_i = gen_add_fundef Specification.cert_specification
|
krauss@19611
|
186 |
|
krauss@19564
|
187 |
|
krauss@19564
|
188 |
|
krauss@19564
|
189 |
(* congruence rules *)
|
krauss@19564
|
190 |
|
wenzelm@19617
|
191 |
val cong_add = Thm.declaration_attribute (map_fundef_congs o Drule.add_rule o safe_mk_meta_eq);
|
wenzelm@19617
|
192 |
val cong_del = Thm.declaration_attribute (map_fundef_congs o Drule.del_rule o safe_mk_meta_eq);
|
krauss@19564
|
193 |
|
krauss@21235
|
194 |
(* Datatype hook to declare datatype congs as "fundef_congs" *)
|
krauss@21235
|
195 |
|
krauss@21235
|
196 |
|
wenzelm@21602
|
197 |
fun add_case_cong n thy =
|
wenzelm@21602
|
198 |
Context.theory_map (map_fundef_congs (Drule.add_rule
|
krauss@21235
|
199 |
(DatatypePackage.get_datatype thy n |> the
|
krauss@21235
|
200 |
|> #case_cong
|
wenzelm@21602
|
201 |
|> safe_mk_meta_eq)))
|
krauss@21235
|
202 |
thy
|
krauss@21235
|
203 |
|
krauss@21235
|
204 |
val case_cong_hook = fold add_case_cong
|
krauss@21235
|
205 |
|
wenzelm@21602
|
206 |
val setup_case_cong_hook =
|
krauss@21235
|
207 |
DatatypeHooks.add case_cong_hook
|
krauss@21235
|
208 |
#> (fn thy => case_cong_hook (Symtab.keys (DatatypePackage.get_datatypes thy)) thy)
|
krauss@19564
|
209 |
|
krauss@19564
|
210 |
(* setup *)
|
krauss@19564
|
211 |
|
wenzelm@21602
|
212 |
val setup =
|
wenzelm@21602
|
213 |
FundefData.init
|
krauss@21235
|
214 |
#> FundefCongs.init
|
krauss@21319
|
215 |
#> TerminationRule.init
|
krauss@21235
|
216 |
#> Attrib.add_attributes
|
krauss@21235
|
217 |
[("fundef_cong", Attrib.add_del_args cong_add cong_del, "declaration of congruence rule for function definitions")]
|
krauss@21235
|
218 |
#> setup_case_cong_hook
|
krauss@21319
|
219 |
#> FundefRelation.setup
|
krauss@19564
|
220 |
|
krauss@19770
|
221 |
val get_congs = FundefCommon.get_fundef_congs o Context.Theory
|
krauss@19770
|
222 |
|
krauss@19770
|
223 |
|
krauss@21235
|
224 |
|
krauss@19564
|
225 |
(* outer syntax *)
|
krauss@19564
|
226 |
|
krauss@19564
|
227 |
local structure P = OuterParse and K = OuterKeyword in
|
krauss@19564
|
228 |
|
krauss@19564
|
229 |
val functionP =
|
krauss@19564
|
230 |
OuterSyntax.command "function" "define general recursive functions" K.thy_goal
|
krauss@22498
|
231 |
(fundef_parser default_config
|
krauss@22498
|
232 |
>> (fn ((config, fixes), statements) =>
|
krauss@22498
|
233 |
Toplevel.local_theory_to_proof (target_of config) (add_fundef fixes statements config #> snd)
|
krauss@21211
|
234 |
#> Toplevel.print));
|
krauss@20523
|
235 |
|
krauss@19564
|
236 |
val terminationP =
|
krauss@19564
|
237 |
OuterSyntax.command "termination" "prove termination of a recursive function" K.thy_goal
|
wenzelm@22102
|
238 |
(P.opt_target -- Scan.option P.name
|
wenzelm@21000
|
239 |
>> (fn (target, name) =>
|
wenzelm@21000
|
240 |
Toplevel.print o
|
krauss@21211
|
241 |
Toplevel.local_theory_to_proof target (setup_termination_proof name)));
|
krauss@19564
|
242 |
|
krauss@19564
|
243 |
val _ = OuterSyntax.add_parsers [functionP];
|
krauss@19564
|
244 |
val _ = OuterSyntax.add_parsers [terminationP];
|
krauss@20654
|
245 |
val _ = OuterSyntax.add_keywords ["sequential", "otherwise"];
|
krauss@19564
|
246 |
|
krauss@19564
|
247 |
end;
|
krauss@19564
|
248 |
|
krauss@19564
|
249 |
|
wenzelm@19585
|
250 |
end
|