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@23203: -> ((bstring * Attrib.src list) * string) list wenzelm@20877: -> FundefCommon.fundef_config krauss@23203: -> bool list wenzelm@20877: -> local_theory krauss@22733: -> Proof.state krauss@20523: wenzelm@20877: val add_fundef_i: (string * typ option * mixfix) list krauss@23203: -> ((bstring * Attrib.src list) * term) list krauss@22170: -> FundefCommon.fundef_config krauss@23203: -> bool list wenzelm@20877: -> local_theory krauss@22733: -> Proof.state krauss@21211: krauss@21211: val setup_termination_proof : string option -> local_theory -> Proof.state krauss@19564: 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@23203: structure FundefPackage : FUNDEF_PACKAGE = krauss@19564: struct krauss@19564: krauss@21051: open FundefLib krauss@19564: open FundefCommon krauss@19564: wenzelm@21435: val note_theorem = LocalTheory.note Thm.theoremK wenzelm@21435: krauss@22166: fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_" krauss@22166: krauss@23819: fun add_simps fnames post sort label moreatts simps lthy = krauss@19564: let krauss@23819: val atts = Attrib.internal (K Simplifier.simp_add) :: moreatts krauss@23819: val spec = post simps krauss@23819: |> map (apfst (apsnd (append atts))) krauss@21255: krauss@22166: val (saved_spec_simps, lthy) = krauss@23819: fold_map note_theorem spec lthy krauss@23819: krauss@22498: val saved_simps = flat (map snd saved_spec_simps) krauss@22166: val simps_by_f = sort saved_simps krauss@19564: krauss@22166: fun add_for_f fname simps = krauss@23819: note_theorem ((NameSpace.qualified fname label, []), simps) #> snd krauss@19564: in krauss@22166: (saved_simps, krauss@22166: fold2 add_for_f fnames simps_by_f lthy) krauss@19564: end krauss@19564: krauss@23203: fun fundef_afterqed config fixes post defname cont sort_cont [[proof]] lthy = krauss@20523: let krauss@22733: val FundefResult {fs, R, psimps, trsimps, subset_pinducts, simple_pinducts, termination, domintros, cases, ...} = krauss@22166: cont (Goal.close_result proof) krauss@22166: krauss@23819: val fnames = map (fst o fst) fixes wenzelm@21391: val qualify = NameSpace.qualified defname krauss@23819: val addsmps = add_simps fnames post sort_cont wenzelm@21602: wenzelm@21602: val (((psimps', pinducts'), (_, [termination'])), lthy) = krauss@21255: lthy krauss@22170: |> addsmps "psimps" [] psimps krauss@22170: ||> fold_option (snd oo addsmps "simps" []) trsimps krauss@22166: ||>> note_theorem ((qualify "pinduct", krauss@22166: [Attrib.internal (K (InductAttrib.induct_set ""))]), simple_pinducts) krauss@22166: ||>> note_theorem ((qualify "termination", []), [termination]) krauss@22166: ||> (snd o note_theorem ((qualify "cases", []), [cases])) krauss@22166: ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros krauss@21255: krauss@22170: val cdata = FundefCtxData { add_simps=addsmps, psimps=psimps', krauss@22733: pinducts=snd pinducts', termination=termination', fs=fs, R=R, defname=defname } wenzelm@22668: val cdata' = cdata |> morph_fundef_data (LocalTheory.target_morphism lthy); (* FIXME !? *) krauss@20523: in krauss@22623: lthy krauss@22733: |> LocalTheory.declaration (fn phi => add_fundef_data (morph_fundef_data phi cdata)) (* save in target *) krauss@22733: |> Context.proof_map (add_fundef_data cdata') (* also save in local context *) krauss@20523: end (* FIXME: Add cases for induct and cases thm *) krauss@20523: krauss@20523: krauss@23203: fun prepare_spec prep fixspec eqnss lthy = (* FIXME: obsolete when "read_specification" etc. is there *) krauss@19564: let krauss@23203: val eqns = map (apsnd single) eqnss krauss@22498: krauss@23203: val ((fixes, _), ctxt') = prep fixspec [] lthy krauss@23203: fun prep_eqn e = the_single (snd (fst (prep [] [e] ctxt'))) krauss@23189: krauss@23203: val spec = map prep_eqn eqns krauss@23203: |> map (apsnd (map (fn t => fold_rev (mk_forall o Free) (frees_in_term ctxt' t) t))) (* Add quantifiers *) krauss@20523: in krauss@20523: ((fixes, spec), ctxt') krauss@20523: end krauss@20523: krauss@23203: fun gen_add_fundef prep fixspec eqnss config flags lthy = krauss@20523: let krauss@23203: val ((fixes, spec), ctxt') = prepare_spec prep fixspec eqnss lthy krauss@23819: val (eqs, post, sort_cont) = FundefCommon.get_preproc lthy config flags ctxt' fixes spec krauss@22166: krauss@22166: val defname = mk_defname fixes krauss@22166: krauss@23819: val ((goalstate, cont), lthy) = krauss@23203: FundefMutual.prepare_fundef_mutual config defname fixes eqs lthy krauss@20523: krauss@23203: val afterqed = fundef_afterqed config fixes post defname cont sort_cont krauss@20523: in krauss@22733: lthy krauss@22733: |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]] krauss@22733: |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd krauss@20523: end krauss@20523: krauss@22733: fun total_termination_afterqed data [[totality]] lthy = krauss@20523: let krauss@22733: val FundefCtxData { add_simps, psimps, pinducts, defname, ... } = data krauss@21255: krauss@22166: val totality = Goal.close_result totality krauss@19564: krauss@21255: val remove_domain_condition = full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals]) krauss@19564: krauss@22166: val tsimps = map remove_domain_condition psimps krauss@22166: val tinduct = map remove_domain_condition pinducts krauss@19770: krauss@21511: val has_guards = exists ((fn (Const ("Trueprop", _) $ _) => false | _ => true) o prop_of) tsimps haftmann@24624: val allatts = if has_guards then [] else [Attrib.internal (K RecfunCodegen.add_default)] wenzelm@21602: krauss@21511: val qualify = NameSpace.qualified defname; krauss@19564: in krauss@21511: lthy krauss@22166: |> add_simps "simps" allatts tsimps |> snd krauss@22166: |> note_theorem ((qualify "induct", []), tinduct) |> snd wenzelm@20363: end krauss@19564: krauss@19564: krauss@22733: fun setup_termination_proof term_opt lthy = krauss@20523: let krauss@22733: val data = the (case term_opt of wenzelm@24508: SOME t => import_fundef_data (Syntax.read_term lthy t) (Context.Proof lthy) krauss@22733: | NONE => import_last_fundef (Context.Proof lthy)) krauss@22733: handle Option.Option => raise ERROR ("Not a function: " ^ quote (the_default "" term_opt)) krauss@20523: krauss@22166: val FundefCtxData {termination, R, ...} = data krauss@22325: val domT = domain_type (fastype_of R) krauss@22325: val goal = HOLogic.mk_Trueprop (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT))) krauss@20523: in wenzelm@20877: lthy krauss@22325: |> ProofContext.note_thmss_i "" [(("", [ContextRules.rule_del]), [([allI], [])])] |> snd krauss@22325: |> ProofContext.note_thmss_i "" [(("", [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd wenzelm@21602: |> ProofContext.note_thmss_i "" krauss@22325: [(("termination", [ContextRules.intro_bang (SOME 0)]), wenzelm@21602: [([Goal.norm_result termination], [])])] |> snd krauss@22733: |> Proof.theorem_i NONE (total_termination_afterqed data) [[(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@21235: (* Datatype hook to declare datatype congs as "fundef_congs" *) krauss@21235: krauss@21235: wenzelm@21602: fun add_case_cong n thy = krauss@24168: Context.theory_map (FundefCtxTree.map_fundef_congs (Thm.add_thm krauss@21235: (DatatypePackage.get_datatype thy n |> the krauss@21235: |> #case_cong wenzelm@21602: |> safe_mk_meta_eq))) krauss@21235: thy krauss@21235: krauss@21235: val case_cong_hook = fold add_case_cong krauss@21235: wenzelm@21602: 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@23819: krauss@23473: (* ad-hoc method to convert elimination-style goals to existential statements *) krauss@23473: krauss@23473: fun insert_int_goal thy subg st = krauss@23473: let krauss@23473: val goal = hd (prems_of st) krauss@23473: val (ps, imp) = dest_all_all goal krauss@23473: val cps = map (cterm_of thy) ps krauss@23473: krauss@23473: val imp_subg = fold (fn p => fn t => betapply (t,p)) ps subg krauss@23473: val new_subg = implies $ imp_subg $ imp krauss@23473: |> fold_rev mk_forall ps krauss@23473: |> cterm_of thy krauss@23473: |> assume krauss@23473: krauss@23473: val sg2 = imp_subg krauss@23473: |> fold_rev mk_forall ps krauss@23473: |> cterm_of thy krauss@23473: |> assume krauss@23473: krauss@23473: val t' = new_subg krauss@23473: |> fold forall_elim cps krauss@23473: |> flip implies_elim (fold forall_elim cps sg2) krauss@23473: |> fold_rev forall_intr cps krauss@23473: krauss@23473: val st' = implies_elim st t' krauss@23473: |> implies_intr (cprop_of sg2) krauss@23473: |> implies_intr (cprop_of new_subg) krauss@23473: in krauss@23473: Seq.single st' krauss@23473: end krauss@23473: krauss@23473: fun mk_cases_statement thy t = krauss@23473: let krauss@23473: fun mk_clause t = krauss@23473: let krauss@23473: val (qs, imp) = dest_all_all t krauss@23473: in krauss@23473: Logic.strip_imp_prems imp krauss@23473: |> map (ObjectLogic.atomize_term thy) krauss@23473: |> foldr1 HOLogic.mk_conj krauss@23473: |> fold_rev (fn Free (v,T) => fn t => HOLogic.mk_exists (v,T,t)) qs krauss@23473: end krauss@23473: krauss@23473: val (ps, imp) = dest_all_all t krauss@23473: in krauss@23473: Logic.strip_imp_prems imp krauss@23473: |> map mk_clause krauss@23473: |> foldr1 HOLogic.mk_disj krauss@23473: |> HOLogic.mk_Trueprop krauss@23473: |> fold_rev lambda ps krauss@23473: end krauss@23473: krauss@23473: fun elim_to_cases1 ctxt st = krauss@23473: let krauss@23473: val thy = theory_of_thm st krauss@23473: val [subg] = prems_of st krauss@23473: val cex = mk_cases_statement thy subg krauss@23473: in krauss@23473: (insert_int_goal thy cex krauss@23473: THEN REPEAT_ALL_NEW (Tactic.ematch_tac [disjE, exE, conjE]) 1 krauss@23473: THEN REPEAT (Goal.assume_rule_tac ctxt 1) krauss@23473: (* THEN REPEAT (etac thin_rl 1)*)) st krauss@23473: end krauss@23473: krauss@23473: fun elim_to_cases_tac ctxt = SELECT_GOAL (elim_to_cases1 ctxt) krauss@23473: krauss@23473: val elim_to_cases_setup = Method.add_methods krauss@23473: [("elim_to_cases", Method.ctxt_args (Method.SIMPLE_METHOD' o elim_to_cases_tac), krauss@23473: "convert elimination-style goal to a disjunction of existentials")] krauss@23473: krauss@19564: (* setup *) krauss@19564: wenzelm@21602: val setup = wenzelm@22846: Attrib.add_attributes krauss@24168: [("fundef_cong", Attrib.add_del_args FundefCtxTree.cong_add FundefCtxTree.cong_del, wenzelm@22846: "declaration of congruence rule for function definitions")] wenzelm@22846: #> setup_case_cong_hook wenzelm@22846: #> FundefRelation.setup krauss@23473: #> elim_to_cases_setup krauss@19564: krauss@24168: val get_congs = FundefCtxTree.get_fundef_congs o Context.Theory krauss@19770: krauss@21235: krauss@19564: (* outer syntax *) krauss@19564: krauss@19564: local structure P = OuterParse and K = OuterKeyword in krauss@19564: krauss@19564: val functionP = krauss@19564: OuterSyntax.command "function" "define general recursive functions" K.thy_goal krauss@22498: (fundef_parser default_config krauss@23203: >> (fn ((config, fixes), (flags, statements)) => krauss@23203: Toplevel.local_theory_to_proof (target_of config) (add_fundef fixes statements config flags) krauss@21211: #> Toplevel.print)); krauss@20523: krauss@19564: val terminationP = krauss@19564: OuterSyntax.command "termination" "prove termination of a recursive function" K.thy_goal krauss@22733: (P.opt_target -- Scan.option P.term 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@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