krauss@19564: (* Title: HOL/Tools/function_package/fundef_package.ML krauss@19564: ID: $Id$ krauss@19564: Author: Alexander Krauss, TU Muenchen krauss@19564: wenzelm@20363: A package for general recursive function definitions. krauss@19564: Isar commands. krauss@19564: krauss@19564: *) krauss@19564: wenzelm@20363: signature FUNDEF_PACKAGE = krauss@19564: sig wenzelm@20877: val add_fundef : (string * string option * mixfix) list krauss@20523: -> ((bstring * Attrib.src list) * string list) list list wenzelm@20877: -> FundefCommon.fundef_config wenzelm@20877: -> local_theory krauss@21211: -> string * Proof.state krauss@20523: wenzelm@20877: val add_fundef_i: (string * typ option * mixfix) list krauss@20523: -> ((bstring * Attrib.src list) * term list) list list wenzelm@20877: -> bool wenzelm@20877: -> local_theory krauss@21211: -> string * Proof.state krauss@21211: krauss@21211: val setup_termination_proof : string option -> local_theory -> Proof.state krauss@19564: krauss@19564: val cong_add: attribute krauss@19564: val cong_del: attribute wenzelm@20363: krauss@19564: val setup : theory -> theory krauss@21235: val setup_case_cong_hook : theory -> theory krauss@19770: val get_congs : theory -> thm list krauss@19564: end krauss@19564: krauss@19564: krauss@20523: structure FundefPackage = krauss@19564: struct krauss@19564: krauss@21051: open FundefLib krauss@19564: open FundefCommon krauss@19564: krauss@20654: (*fn (ctxt,s) => Variable.importT_terms (fst (ProofContext.read_termTs ctxt (K false) (K NONE) (K NONE) [] [s])) ctxt;*) krauss@19564: krauss@20523: fun burrow_snd f ps = (* ('a list -> 'b list) -> ('c * 'a) list -> ('c * 'b) list *) krauss@20523: let val (xs, ys) = split_list ps krauss@20523: in xs ~~ f ys end krauss@20270: krauss@20523: fun restore_spec_structure reps spec = krauss@20523: (burrow o burrow_snd o burrow o K) reps spec wenzelm@20363: krauss@20523: fun add_simps label moreatts mutual_info fixes psimps spec lthy = krauss@19564: let krauss@20523: val fnames = map (fst o fst) fixes krauss@20523: val psimps_by_f = FundefMutual.sort_by_function mutual_info fnames psimps krauss@19564: krauss@20523: fun add_for_f fname psimps = wenzelm@20877: LocalTheory.note ((NameSpace.append fname label, Attrib.internal Simplifier.simp_add :: moreatts), psimps) #> snd krauss@19564: in krauss@20523: lthy krauss@20523: |> fold_map (fold_map LocalTheory.note) (restore_spec_structure psimps spec) |> snd krauss@20523: |> fold2 add_for_f fnames psimps_by_f krauss@19564: end krauss@19564: krauss@19564: krauss@20523: fun fundef_afterqed fixes spec mutual_info defname data [[result]] lthy = krauss@20523: let krauss@20523: val fundef_data = FundefMutual.mk_partial_rules_mutual lthy mutual_info data result krauss@20523: val FundefMResult {psimps, subset_pinducts, simple_pinducts, termination, domintros, cases, ...} = fundef_data wenzelm@20877: val qualify = NameSpace.append defname; krauss@20523: in krauss@20523: lthy krauss@20523: |> add_simps "psimps" [] mutual_info fixes psimps spec wenzelm@20877: |> LocalTheory.note ((qualify "domintros", []), domintros) |> snd wenzelm@20877: |> LocalTheory.note ((qualify "termination", []), [termination]) |> snd wenzelm@20877: |> LocalTheory.note ((qualify "cases", []), [cases]) |> snd wenzelm@20877: |> LocalTheory.note ((qualify "pinduct", [Attrib.internal (InductAttrib.induct_set "")]), simple_pinducts) |> snd wenzelm@20877: |> LocalTheory.declaration (add_fundef_data defname (fundef_data, mutual_info, (fixes,spec))) krauss@20523: end (* FIXME: Add cases for induct and cases thm *) krauss@20523: krauss@20523: krauss@20523: krauss@20523: fun prep_with_flags prep fixspec eqnss_flags global_flag lthy = krauss@19564: let krauss@20523: val eqnss = map (map (apsnd (map fst))) eqnss_flags krauss@20523: val flags = map (map (map (fn (_, f) => global_flag orelse f) o snd)) eqnss_flags krauss@20523: krauss@20523: val ((fixes, _), ctxt') = prep fixspec [] lthy krauss@20523: val spec = map (fn eqns => snd (fst (prep [] eqns ctxt'))) eqnss krauss@20523: |> map (map (apsnd (map (fn t => fold_rev (mk_forall o Free) (frees_in_term ctxt' t) t)))) (* Add quantifiers *) krauss@20523: |> map2 (map2 (fn fs => fn (r, thms) => (r, fs ~~ thms))) flags wenzelm@20877: |> (burrow o burrow_snd o burrow) krauss@20523: (FundefSplit.split_some_equations lthy) krauss@20523: |> map (map (apsnd flat)) krauss@20523: in krauss@20523: ((fixes, spec), ctxt') krauss@20523: end krauss@20523: krauss@19564: krauss@20654: fun gen_add_fundef prep_spec fixspec eqnss_flags config lthy = krauss@20523: let krauss@20654: val FundefConfig {sequential, default, ...} = config krauss@20654: krauss@20654: val ((fixes, spec), ctxt') = prep_with_flags prep_spec fixspec eqnss_flags sequential lthy krauss@20523: val t_eqns = spec krauss@20523: |> flat |> map snd |> flat (* flatten external structure *) krauss@20523: wenzelm@20877: val ((mutual_info, name, prep_result as Prep {goal, goalI, ...}), lthy) = krauss@20654: FundefMutual.prepare_fundef_mutual fixes t_eqns default lthy krauss@20523: krauss@20523: val afterqed = fundef_afterqed fixes spec mutual_info name prep_result krauss@20523: in krauss@21211: (name, lthy krauss@21211: |> Proof.theorem_i PureThy.internalK NONE afterqed NONE ("", []) [(("", []), [(goal, [])])] krauss@21211: |> Proof.refine (Method.primitive_text (fn _ => goalI)) |> Seq.hd) krauss@20523: end krauss@20523: krauss@20523: krauss@20523: fun total_termination_afterqed defname data [[totality]] lthy = krauss@20523: let krauss@20523: val (FundefMResult {psimps, simple_pinducts, ... }, mutual_info, (fixes, stmts)) = data krauss@19564: wenzelm@20363: val remove_domain_condition = full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals]) krauss@19564: krauss@20523: val tsimps = map remove_domain_condition psimps wenzelm@20363: val tinduct = map remove_domain_condition simple_pinducts krauss@19770: wenzelm@20877: (* FIXME: How to generate code from (possibly) local contexts krauss@20523: val has_guards = exists ((fn (Const ("Trueprop", _) $ _) => false | _ => true) o prop_of) tsimps krauss@20523: val allatts = if has_guards then [] else [Attrib.internal (RecfunCodegen.add NONE)] krauss@20523: *) wenzelm@20877: val qualify = NameSpace.append defname; krauss@19564: in krauss@20523: lthy krauss@20523: |> add_simps "simps" [] mutual_info fixes tsimps stmts wenzelm@20877: |> LocalTheory.note ((qualify "induct", []), tinduct) |> snd wenzelm@20363: end krauss@19564: krauss@19564: krauss@21211: fun setup_termination_proof name_opt lthy = krauss@20523: let krauss@20523: val name = the_default (get_last_fundef (Context.Proof lthy)) name_opt krauss@20523: val data = the (get_fundef_data name (Context.Proof lthy)) krauss@20523: handle Option.Option => raise ERROR ("No such function definition: " ^ name) krauss@20523: krauss@20523: val (res as FundefMResult {termination, ...}, _, _) = data krauss@20523: val goal = FundefTermination.mk_total_termination_goal data krauss@20523: in wenzelm@20877: lthy krauss@20523: |> ProofContext.note_thmss_i [(("termination", krauss@21237: [ContextRules.intro_query NONE]), [(ProofContext.standard lthy [termination], [])])] |> snd krauss@20523: |> Proof.theorem_i PureThy.internalK NONE krauss@20523: (total_termination_afterqed name data) NONE ("", []) krauss@20523: [(("", []), [(goal, [])])] krauss@20523: end krauss@20523: krauss@20523: krauss@20523: val add_fundef = gen_add_fundef Specification.read_specification krauss@20523: val add_fundef_i = gen_add_fundef Specification.cert_specification krauss@19611: krauss@19564: krauss@19564: krauss@19564: (* congruence rules *) krauss@19564: wenzelm@19617: val cong_add = Thm.declaration_attribute (map_fundef_congs o Drule.add_rule o safe_mk_meta_eq); wenzelm@19617: val cong_del = Thm.declaration_attribute (map_fundef_congs o Drule.del_rule o safe_mk_meta_eq); krauss@19564: krauss@21235: (* Datatype hook to declare datatype congs as "fundef_congs" *) krauss@21235: krauss@21235: krauss@21235: fun add_case_cong n thy = krauss@21235: Context.theory_map (map_fundef_congs (Drule.add_rule krauss@21235: (DatatypePackage.get_datatype thy n |> the krauss@21235: |> #case_cong krauss@21235: |> safe_mk_meta_eq))) krauss@21235: thy krauss@21235: krauss@21235: val case_cong_hook = fold add_case_cong krauss@21235: krauss@21235: val setup_case_cong_hook = krauss@21235: DatatypeHooks.add case_cong_hook krauss@21235: #> (fn thy => case_cong_hook (Symtab.keys (DatatypePackage.get_datatypes thy)) thy) krauss@19564: krauss@19564: (* setup *) krauss@19564: krauss@21235: val setup = krauss@21235: FundefData.init krauss@21235: #> FundefCongs.init krauss@21235: #> Attrib.add_attributes krauss@21235: [("fundef_cong", Attrib.add_del_args cong_add cong_del, "declaration of congruence rule for function definitions")] krauss@21235: #> setup_case_cong_hook krauss@19564: krauss@19770: val get_congs = FundefCommon.get_fundef_congs o Context.Theory krauss@19770: krauss@19770: krauss@21235: krauss@19564: (* outer syntax *) krauss@19564: krauss@19564: local structure P = OuterParse and K = OuterKeyword in krauss@19564: krauss@20270: krauss@20270: krauss@20523: fun or_list1 s = P.enum1 "|" s krauss@20523: krauss@20523: val opt_sequential = Scan.optional ((P.$$$ "(" |-- P.$$$ "sequential" --| P.$$$ ")") >> K true) false krauss@20523: krauss@20523: val otherwise = P.$$$ "(" |-- P.$$$ "otherwise" --| P.$$$ ")" krauss@20523: val statement_ow = P.and_list1 (P.opt_thm_name ":" -- Scan.repeat1 (P.prop -- Scan.optional (otherwise >> K true) false)) krauss@20523: val statements_ow = or_list1 statement_ow krauss@20270: krauss@20270: krauss@19564: val functionP = krauss@19564: OuterSyntax.command "function" "define general recursive functions" K.thy_goal krauss@21051: ((config_parser default_config -- P.opt_locale_target -- P.fixes --| P.$$$ "where" -- statements_ow) krauss@20654: >> (fn (((config, target), fixes), statements) => krauss@21211: Toplevel.local_theory_to_proof target (add_fundef fixes statements config #> snd) krauss@21211: #> Toplevel.print)); krauss@20523: krauss@19564: krauss@21051: krauss@19564: val terminationP = krauss@19564: OuterSyntax.command "termination" "prove termination of a recursive function" K.thy_goal wenzelm@21000: (P.opt_locale_target -- Scan.option P.name wenzelm@21000: >> (fn (target, name) => wenzelm@21000: Toplevel.print o krauss@21211: Toplevel.local_theory_to_proof target (setup_termination_proof name))); krauss@19564: krauss@20338: krauss@19564: val _ = OuterSyntax.add_parsers [functionP]; krauss@19564: val _ = OuterSyntax.add_parsers [terminationP]; krauss@20654: val _ = OuterSyntax.add_keywords ["sequential", "otherwise"]; krauss@19564: krauss@19564: end; krauss@19564: krauss@19564: wenzelm@19585: end