1 (* Title: HOL/Tools/function_package/fundef_package.ML
3 Author: Alexander Krauss, TU Muenchen
5 A package for general recursive function definitions.
10 signature FUNDEF_PACKAGE =
12 val add_fundef : (string * string option * mixfix) list
13 -> ((bstring * Attrib.src list) * string list) list list
14 -> FundefCommon.fundef_config
16 -> string * Proof.state
18 val add_fundef_i: (string * typ option * mixfix) list
19 -> ((bstring * Attrib.src list) * term list) list list
22 -> string * Proof.state
24 val setup_termination_proof : string option -> local_theory -> Proof.state
26 val cong_add: attribute
27 val cong_del: attribute
29 val setup : theory -> theory
30 val setup_case_cong_hook : theory -> theory
31 val get_congs : theory -> thm list
35 structure FundefPackage =
41 val note_theorem = LocalTheory.note Thm.theoremK
43 fun burrow_snd f ps = (* ('a list -> 'b list) -> ('c * 'a) list -> ('c * 'b) list *)
44 let val (xs, ys) = split_list ps
47 fun restore_spec_structure reps spec =
48 (burrow o burrow_snd o burrow o K) reps spec
50 fun add_simps label moreatts mutual_info fixes psimps spec lthy =
52 val fnames = map (fst o fst) fixes
54 val (saved_spec_psimps, lthy) =
55 fold_map (fold_map note_theorem) (restore_spec_structure psimps spec) lthy
56 val saved_psimps = flat (map snd (flat saved_spec_psimps))
58 val psimps_by_f = FundefMutual.sort_by_function mutual_info fnames saved_psimps
60 fun add_for_f fname psimps =
62 ((NameSpace.qualified fname label, Attrib.internal (K Simplifier.simp_add) :: moreatts),
66 fold2 add_for_f fnames psimps_by_f lthy)
70 fun fundef_afterqed config fixes spec mutual_info defname data [[result]] lthy =
72 val FundefConfig { domintros, ...} = config
73 val Prep {f, R, ...} = data
74 val fundef_data = FundefMutual.mk_partial_rules_mutual lthy mutual_info data result
75 val FundefMResult {psimps, subset_pinducts, simple_pinducts, termination, domintros, cases, ...} = fundef_data
76 val qualify = NameSpace.qualified defname
78 val (((psimps', pinducts'), (_, [termination'])), lthy) =
80 |> PROFILE "adding_psimps" (add_simps "psimps" [] mutual_info fixes psimps spec)
81 ||>> PROFILE "adding pinduct"
82 (note_theorem ((qualify "pinduct",
83 [Attrib.internal (K (InductAttrib.induct_set ""))]), simple_pinducts))
84 ||>> PROFILE "adding termination"
85 (note_theorem ((qualify "termination", []), [termination]))
86 ||> (snd o PROFILE "adding cases" (note_theorem ((qualify "cases", []), [cases])))
87 ||> (snd o PROFILE "adding domintros" (note_theorem ((qualify "domintros", []), domintros)))
89 val cdata = FundefCtxData { fixes=fixes, spec=spec, mutual=mutual_info, psimps=psimps',
90 pinducts=snd pinducts', termination=termination', f=f, R=R }
94 lthy (* FIXME proper handling of morphism *)
95 |> LocalTheory.declaration (fn phi => add_fundef_data defname cdata) (* save in target *)
96 |> Context.proof_map (add_fundef_data defname cdata) (* also save in local context *)
97 end (* FIXME: Add cases for induct and cases thm *)
101 fun prep_with_flags prep fixspec eqnss_flags global_flag lthy =
103 val eqnss = map (map (apsnd (map fst))) eqnss_flags
104 val flags = map (map (map (fn (_, f) => global_flag orelse f) o snd)) eqnss_flags
106 val ((fixes, _), ctxt') = prep fixspec [] lthy
107 val spec = map (fn eqns => snd (fst (prep [] eqns ctxt'))) eqnss
108 |> map (map (apsnd (map (fn t => fold_rev (mk_forall o Free) (frees_in_term ctxt' t) t)))) (* Add quantifiers *)
109 |> map2 (map2 (fn fs => fn (r, thms) => (r, fs ~~ thms))) flags
110 |> (burrow o burrow_snd o burrow)
111 (FundefSplit.split_some_equations lthy)
112 |> map (map (apsnd flat))
114 ((fixes, spec), ctxt')
118 fun gen_add_fundef prep_spec fixspec eqnss_flags config lthy =
120 val FundefConfig {sequential, default, ...} = config
122 val ((fixes, spec), ctxt') = prep_with_flags prep_spec fixspec eqnss_flags sequential lthy
124 |> flat |> map snd |> flat (* flatten external structure *)
126 val ((mutual_info, name, prep_result as Prep {goal, goalI, ...}), lthy) =
127 FundefMutual.prepare_fundef_mutual fixes t_eqns default lthy
129 val afterqed = fundef_afterqed config fixes spec mutual_info name prep_result
132 |> Proof.theorem_i NONE afterqed [[(goal, [])]]
133 |> Proof.refine (Method.primitive_text (fn _ => goalI)) |> Seq.hd)
137 fun total_termination_afterqed defname data [[totality]] lthy =
139 val FundefCtxData { fixes, spec, mutual, psimps, pinducts, ... } = data
141 val totality = PROFILE "closing" Goal.close_result totality
143 val remove_domain_condition = full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
145 val tsimps = PROFILE "making tsimps" (map remove_domain_condition) psimps
146 val tinduct = PROFILE "making tinduct" (map remove_domain_condition) pinducts
148 (* FIXME: How to generate code from (possibly) local contexts*)
149 val has_guards = exists ((fn (Const ("Trueprop", _) $ _) => false | _ => true) o prop_of) tsimps
150 val allatts = if has_guards then [] else [Attrib.internal (K (RecfunCodegen.add NONE))]
152 val qualify = NameSpace.qualified defname;
155 |> PROFILE "adding tsimps" (add_simps "simps" allatts mutual fixes tsimps spec) |> snd
156 |> PROFILE "adding tinduct" (note_theorem ((qualify "induct", []), tinduct)) |> snd
160 fun setup_termination_proof name_opt lthy =
162 val name = the_default (get_last_fundef (Context.Proof lthy)) name_opt
163 val data = the (get_fundef_data name (Context.Proof lthy))
164 handle Option.Option => raise ERROR ("No such function definition: " ^ name)
166 val FundefCtxData {termination, f, R, ...} = data
167 val goal = FundefTermination.mk_total_termination_goal f R
170 |> ProofContext.note_thmss_i ""
171 [(("termination", [ContextRules.intro_query NONE]),
172 [([Goal.norm_result termination], [])])] |> snd
173 |> set_termination_rule termination
174 |> Proof.theorem_i NONE (total_termination_afterqed name data) [[(goal, [])]]
178 val add_fundef = gen_add_fundef Specification.read_specification
179 val add_fundef_i = gen_add_fundef Specification.cert_specification
183 (* congruence rules *)
185 val cong_add = Thm.declaration_attribute (map_fundef_congs o Drule.add_rule o safe_mk_meta_eq);
186 val cong_del = Thm.declaration_attribute (map_fundef_congs o Drule.del_rule o safe_mk_meta_eq);
188 (* Datatype hook to declare datatype congs as "fundef_congs" *)
191 fun add_case_cong n thy =
192 Context.theory_map (map_fundef_congs (Drule.add_rule
193 (DatatypePackage.get_datatype thy n |> the
195 |> safe_mk_meta_eq)))
198 val case_cong_hook = fold add_case_cong
200 val setup_case_cong_hook =
201 DatatypeHooks.add case_cong_hook
202 #> (fn thy => case_cong_hook (Symtab.keys (DatatypePackage.get_datatypes thy)) thy)
209 #> TerminationRule.init
210 #> Attrib.add_attributes
211 [("fundef_cong", Attrib.add_del_args cong_add cong_del, "declaration of congruence rule for function definitions")]
212 #> setup_case_cong_hook
213 #> FundefRelation.setup
215 val get_congs = FundefCommon.get_fundef_congs o Context.Theory
221 local structure P = OuterParse and K = OuterKeyword in
225 fun or_list1 s = P.enum1 "|" s
227 val opt_sequential = Scan.optional ((P.$$$ "(" |-- P.$$$ "sequential" --| P.$$$ ")") >> K true) false
229 val otherwise = P.$$$ "(" |-- P.$$$ "otherwise" --| P.$$$ ")"
230 val statement_ow = P.and_list1 (SpecParse.opt_thm_name ":" -- Scan.repeat1 (P.prop -- Scan.optional (otherwise >> K true) false))
231 val statements_ow = or_list1 statement_ow
235 OuterSyntax.command "function" "define general recursive functions" K.thy_goal
236 ((config_parser default_config -- P.opt_target -- P.fixes --| P.$$$ "where" -- statements_ow)
237 >> (fn (((config, target), fixes), statements) =>
238 Toplevel.local_theory_to_proof target (add_fundef fixes statements config #> snd)
244 OuterSyntax.command "termination" "prove termination of a recursive function" K.thy_goal
245 (P.opt_target -- Scan.option P.name
246 >> (fn (target, name) =>
248 Toplevel.local_theory_to_proof target (setup_termination_proof name)));
251 val _ = OuterSyntax.add_parsers [functionP];
252 val _ = OuterSyntax.add_parsers [terminationP];
253 val _ = OuterSyntax.add_keywords ["sequential", "otherwise"];