src/HOL/Tools/function_package/fundef_package.ML
author wenzelm
Sat, 06 Oct 2007 16:50:04 +0200
changeset 24867 e5b55d7be9bb
parent 24861 cc669ca5f382
child 24977 9f98751c9628
permissions -rw-r--r--
simplified interfaces for outer syntax;
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@23203
    13
                      -> ((bstring * Attrib.src list) * string) list 
wenzelm@20877
    14
                      -> FundefCommon.fundef_config
krauss@23203
    15
                      -> bool list
wenzelm@20877
    16
                      -> local_theory
krauss@22733
    17
                      -> Proof.state
krauss@20523
    18
wenzelm@20877
    19
    val add_fundef_i:  (string * typ option * mixfix) list
krauss@23203
    20
                       -> ((bstring * Attrib.src list) * term) list
krauss@22170
    21
                       -> FundefCommon.fundef_config
krauss@23203
    22
                       -> bool list
wenzelm@20877
    23
                       -> local_theory
krauss@22733
    24
                       -> Proof.state
krauss@21211
    25
krauss@21211
    26
    val setup_termination_proof : string option -> local_theory -> Proof.state
krauss@19564
    27
krauss@19564
    28
    val setup : theory -> theory
krauss@19770
    29
    val get_congs : theory -> thm list
krauss@19564
    30
end
krauss@19564
    31
krauss@19564
    32
krauss@23203
    33
structure FundefPackage : FUNDEF_PACKAGE =
krauss@19564
    34
struct
krauss@19564
    35
krauss@21051
    36
open FundefLib
krauss@19564
    37
open FundefCommon
krauss@19564
    38
wenzelm@21435
    39
val note_theorem = LocalTheory.note Thm.theoremK
wenzelm@21435
    40
krauss@22166
    41
fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_" 
krauss@22166
    42
krauss@23819
    43
fun add_simps fnames post sort label moreatts simps lthy =
krauss@20523
    44
    let
krauss@23819
    45
      val atts = Attrib.internal (K Simplifier.simp_add) :: moreatts
krauss@23819
    46
      val spec = post simps
krauss@23819
    47
                   |> map (apfst (apsnd (append atts)))
krauss@21255
    48
krauss@22166
    49
      val (saved_spec_simps, lthy) =
krauss@23819
    50
        fold_map note_theorem spec lthy
krauss@23819
    51
krauss@22498
    52
      val saved_simps = flat (map snd saved_spec_simps)
krauss@22166
    53
      val simps_by_f = sort saved_simps
krauss@19770
    54
krauss@22166
    55
      fun add_for_f fname simps =
krauss@23819
    56
        note_theorem ((NameSpace.qualified fname label, []), simps) #> snd
krauss@19564
    57
    in
krauss@22166
    58
      (saved_simps,
krauss@22166
    59
       fold2 add_for_f fnames simps_by_f lthy)
krauss@19564
    60
    end
krauss@19564
    61
krauss@23203
    62
fun fundef_afterqed config fixes post defname cont sort_cont [[proof]] lthy =
krauss@19564
    63
    let
krauss@22733
    64
      val FundefResult {fs, R, psimps, trsimps, subset_pinducts, simple_pinducts, termination, domintros, cases, ...} = 
krauss@22166
    65
          cont (Goal.close_result proof)
krauss@22166
    66
krauss@23819
    67
      val fnames = map (fst o fst) fixes
wenzelm@21391
    68
      val qualify = NameSpace.qualified defname
krauss@23819
    69
      val addsmps = add_simps fnames post sort_cont
wenzelm@21602
    70
wenzelm@21602
    71
      val (((psimps', pinducts'), (_, [termination'])), lthy) =
krauss@21255
    72
          lthy
krauss@22170
    73
            |> addsmps "psimps" [] psimps
krauss@22170
    74
            ||> fold_option (snd oo addsmps "simps" []) trsimps
krauss@22166
    75
            ||>> note_theorem ((qualify "pinduct",
wenzelm@24861
    76
                                [Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
krauss@22166
    77
            ||>> note_theorem ((qualify "termination", []), [termination])
krauss@22166
    78
            ||> (snd o note_theorem ((qualify "cases", []), [cases]))
krauss@22166
    79
            ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
krauss@21255
    80
krauss@22170
    81
      val cdata = FundefCtxData { add_simps=addsmps, psimps=psimps',
krauss@22733
    82
                                  pinducts=snd pinducts', termination=termination', fs=fs, R=R, defname=defname }
wenzelm@22668
    83
      val cdata' = cdata |> morph_fundef_data (LocalTheory.target_morphism lthy);  (* FIXME !? *)
krauss@20523
    84
    in
krauss@22623
    85
      lthy 
krauss@22733
    86
        |> LocalTheory.declaration (fn phi => add_fundef_data (morph_fundef_data phi cdata)) (* save in target *)
krauss@22733
    87
        |> Context.proof_map (add_fundef_data cdata') (* also save in local context *)
krauss@20523
    88
    end (* FIXME: Add cases for induct and cases thm *)
wenzelm@20363
    89
krauss@19611
    90
wenzelm@24722
    91
fun prepare_spec prep fixspec eqnss lthy = (* FIXME: obsolete; use Specification.read/check_specification *)
krauss@23203
    92
    let
krauss@23203
    93
      val eqns = map (apsnd single) eqnss
krauss@19770
    94
krauss@23203
    95
      val ((fixes, _), ctxt') = prep fixspec [] lthy             
wenzelm@24722
    96
      fun prep_eqn e = the_single (snd (fst (prep [] [[e]] ctxt')))
krauss@20523
    97
krauss@23203
    98
      val spec = map prep_eqn eqns
krauss@23203
    99
                 |> map (apsnd (map (fn t => fold_rev (mk_forall o Free) (frees_in_term ctxt' t) t))) (* Add quantifiers *)
krauss@19564
   100
    in
krauss@20523
   101
      ((fixes, spec), ctxt')
krauss@19564
   102
    end
krauss@19564
   103
krauss@23203
   104
fun gen_add_fundef prep fixspec eqnss config flags lthy =
krauss@19564
   105
    let
krauss@23203
   106
      val ((fixes, spec), ctxt') = prepare_spec prep fixspec eqnss lthy
krauss@23819
   107
      val (eqs, post, sort_cont) = FundefCommon.get_preproc lthy config flags ctxt' fixes spec
krauss@22166
   108
krauss@22166
   109
      val defname = mk_defname fixes
krauss@22166
   110
krauss@23819
   111
      val ((goalstate, cont), lthy) =
krauss@23203
   112
          FundefMutual.prepare_fundef_mutual config defname fixes eqs lthy
krauss@19564
   113
krauss@23203
   114
      val afterqed = fundef_afterqed config fixes post defname cont sort_cont
krauss@20523
   115
    in
krauss@22733
   116
      lthy
krauss@22733
   117
        |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
krauss@22733
   118
        |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
krauss@20523
   119
    end
krauss@20523
   120
krauss@22733
   121
fun total_termination_afterqed data [[totality]] lthy =
krauss@20523
   122
    let
krauss@22733
   123
      val FundefCtxData { add_simps, psimps, pinducts, defname, ... } = data
krauss@19564
   124
krauss@22166
   125
      val totality = Goal.close_result totality
krauss@19564
   126
krauss@21255
   127
      val remove_domain_condition = full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
krauss@21255
   128
krauss@22166
   129
      val tsimps = map remove_domain_condition psimps
krauss@22166
   130
      val tinduct = map remove_domain_condition pinducts
krauss@19770
   131
krauss@21511
   132
      val has_guards = exists ((fn (Const ("Trueprop", _) $ _) => false | _ => true) o prop_of) tsimps
haftmann@24624
   133
      val allatts = if has_guards then [] else [Attrib.internal (K RecfunCodegen.add_default)]
wenzelm@21602
   134
krauss@21511
   135
      val qualify = NameSpace.qualified defname;
krauss@19564
   136
    in
krauss@21511
   137
      lthy
krauss@22166
   138
        |> add_simps "simps" allatts tsimps |> snd
krauss@22166
   139
        |> note_theorem ((qualify "induct", []), tinduct) |> snd
krauss@19564
   140
    end
krauss@19564
   141
krauss@20523
   142
krauss@22733
   143
fun setup_termination_proof term_opt lthy =
krauss@19564
   144
    let
krauss@22733
   145
      val data = the (case term_opt of
wenzelm@24508
   146
                        SOME t => import_fundef_data (Syntax.read_term lthy t) (Context.Proof lthy)
krauss@22733
   147
                      | NONE => import_last_fundef (Context.Proof lthy))
krauss@22733
   148
          handle Option.Option => raise ERROR ("Not a function: " ^ quote (the_default "" term_opt))
krauss@19564
   149
krauss@22166
   150
        val FundefCtxData {termination, R, ...} = data
krauss@22325
   151
        val domT = domain_type (fastype_of R)
krauss@22325
   152
        val goal = HOLogic.mk_Trueprop (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
krauss@19564
   153
    in
wenzelm@20877
   154
      lthy
krauss@22325
   155
        |> ProofContext.note_thmss_i "" [(("", [ContextRules.rule_del]), [([allI], [])])] |> snd
krauss@22325
   156
        |> ProofContext.note_thmss_i "" [(("", [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
wenzelm@21602
   157
        |> ProofContext.note_thmss_i ""
krauss@22325
   158
          [(("termination", [ContextRules.intro_bang (SOME 0)]),
wenzelm@21602
   159
            [([Goal.norm_result termination], [])])] |> snd
krauss@22733
   160
        |> Proof.theorem_i NONE (total_termination_afterqed data) [[(goal, [])]]
wenzelm@20363
   161
    end
krauss@19564
   162
krauss@19564
   163
krauss@20523
   164
val add_fundef = gen_add_fundef Specification.read_specification
wenzelm@24722
   165
val add_fundef_i = gen_add_fundef Specification.check_specification
krauss@19611
   166
krauss@19564
   167
krauss@21235
   168
(* Datatype hook to declare datatype congs as "fundef_congs" *)
krauss@21235
   169
krauss@21235
   170
wenzelm@21602
   171
fun add_case_cong n thy =
krauss@24168
   172
    Context.theory_map (FundefCtxTree.map_fundef_congs (Thm.add_thm
krauss@21235
   173
                          (DatatypePackage.get_datatype thy n |> the
krauss@21235
   174
                           |> #case_cong
wenzelm@21602
   175
                           |> safe_mk_meta_eq)))
krauss@21235
   176
                       thy
krauss@21235
   177
haftmann@24626
   178
val case_cong = fold add_case_cong
krauss@21235
   179
wenzelm@24711
   180
val setup_case_cong = DatatypePackage.interpretation case_cong
haftmann@24626
   181
krauss@19564
   182
krauss@23819
   183
krauss@23473
   184
(* ad-hoc method to convert elimination-style goals to existential statements *)
krauss@23473
   185
krauss@23473
   186
fun insert_int_goal thy subg st =
krauss@23473
   187
    let
krauss@23473
   188
      val goal = hd (prems_of st)
krauss@23473
   189
      val (ps, imp) = dest_all_all goal
krauss@23473
   190
      val cps = map (cterm_of thy) ps
krauss@23473
   191
krauss@23473
   192
      val imp_subg = fold (fn p => fn t => betapply (t,p)) ps subg
krauss@23473
   193
      val new_subg = implies $ imp_subg $ imp
krauss@23473
   194
                      |> fold_rev mk_forall ps
krauss@23473
   195
                      |> cterm_of thy 
krauss@23473
   196
                      |> assume 
krauss@23473
   197
krauss@23473
   198
      val sg2 = imp_subg
krauss@23473
   199
                 |> fold_rev mk_forall ps
krauss@23473
   200
                 |> cterm_of thy 
krauss@23473
   201
                 |> assume
krauss@23473
   202
krauss@23473
   203
      val t' = new_subg
krauss@23473
   204
                |> fold forall_elim cps
krauss@23473
   205
                |> flip implies_elim (fold forall_elim cps sg2)
krauss@23473
   206
                |> fold_rev forall_intr cps
krauss@23473
   207
krauss@23473
   208
      val st' = implies_elim st t'
krauss@23473
   209
                 |> implies_intr (cprop_of sg2)
krauss@23473
   210
                 |> implies_intr (cprop_of new_subg)
krauss@23473
   211
    in
krauss@23473
   212
      Seq.single st'
krauss@23473
   213
    end
krauss@23473
   214
krauss@23473
   215
fun mk_cases_statement thy t =
krauss@23473
   216
    let
krauss@23473
   217
      fun mk_clause t = 
krauss@23473
   218
          let 
krauss@23473
   219
            val (qs, imp) = dest_all_all t
krauss@23473
   220
          in 
krauss@23473
   221
            Logic.strip_imp_prems imp
krauss@23473
   222
             |> map (ObjectLogic.atomize_term thy)
krauss@23473
   223
             |> foldr1 HOLogic.mk_conj
krauss@23473
   224
             |> fold_rev (fn Free (v,T) => fn t => HOLogic.mk_exists (v,T,t)) qs
krauss@23473
   225
          end
krauss@23473
   226
krauss@23473
   227
      val (ps, imp) = dest_all_all t
krauss@23473
   228
    in 
krauss@23473
   229
      Logic.strip_imp_prems imp
krauss@23473
   230
       |> map mk_clause
krauss@23473
   231
       |> foldr1 HOLogic.mk_disj
krauss@23473
   232
       |> HOLogic.mk_Trueprop
krauss@23473
   233
       |> fold_rev lambda ps
krauss@23473
   234
    end
krauss@23473
   235
krauss@23473
   236
fun elim_to_cases1 ctxt st =
krauss@23473
   237
    let
krauss@23473
   238
      val thy = theory_of_thm st
krauss@23473
   239
      val [subg] = prems_of st
krauss@23473
   240
      val cex = mk_cases_statement thy subg
krauss@23473
   241
    in
krauss@23473
   242
      (insert_int_goal thy cex
krauss@23473
   243
       THEN REPEAT_ALL_NEW (Tactic.ematch_tac [disjE, exE, conjE]) 1
krauss@23473
   244
       THEN REPEAT (Goal.assume_rule_tac ctxt 1)
krauss@23473
   245
    (*   THEN REPEAT (etac thin_rl 1)*)) st
krauss@23473
   246
    end
krauss@23473
   247
krauss@23473
   248
fun elim_to_cases_tac ctxt = SELECT_GOAL (elim_to_cases1 ctxt)
krauss@23473
   249
krauss@23473
   250
val elim_to_cases_setup = Method.add_methods
krauss@23473
   251
  [("elim_to_cases", Method.ctxt_args (Method.SIMPLE_METHOD' o elim_to_cases_tac),
krauss@23473
   252
    "convert elimination-style goal to a disjunction of existentials")]
krauss@23473
   253
krauss@19564
   254
(* setup *)
krauss@19564
   255
wenzelm@21602
   256
val setup =
wenzelm@22846
   257
  Attrib.add_attributes
krauss@24168
   258
    [("fundef_cong", Attrib.add_del_args FundefCtxTree.cong_add FundefCtxTree.cong_del,
wenzelm@22846
   259
      "declaration of congruence rule for function definitions")]
haftmann@24626
   260
  #> setup_case_cong
wenzelm@22846
   261
  #> FundefRelation.setup
krauss@23473
   262
  #> elim_to_cases_setup
krauss@19564
   263
krauss@24168
   264
val get_congs = FundefCtxTree.get_fundef_congs o Context.Theory
krauss@19770
   265
krauss@21235
   266
krauss@19564
   267
(* outer syntax *)
krauss@19564
   268
krauss@19564
   269
local structure P = OuterParse and K = OuterKeyword in
krauss@19564
   270
wenzelm@24867
   271
val _ = OuterSyntax.keywords ["sequential", "otherwise"];
wenzelm@24867
   272
wenzelm@24867
   273
val _ =
krauss@19564
   274
  OuterSyntax.command "function" "define general recursive functions" K.thy_goal
krauss@22498
   275
  (fundef_parser default_config
krauss@23203
   276
     >> (fn ((config, fixes), (flags, statements)) =>
krauss@23203
   277
            Toplevel.local_theory_to_proof (target_of config) (add_fundef fixes statements config flags)
krauss@21211
   278
            #> Toplevel.print));
krauss@20523
   279
wenzelm@24867
   280
val _ =
krauss@19564
   281
  OuterSyntax.command "termination" "prove termination of a recursive function" K.thy_goal
krauss@22733
   282
  (P.opt_target -- Scan.option P.term
wenzelm@21000
   283
    >> (fn (target, name) =>
wenzelm@21000
   284
           Toplevel.print o
krauss@21211
   285
           Toplevel.local_theory_to_proof target (setup_termination_proof name)));
krauss@19564
   286
krauss@19564
   287
end;
krauss@19564
   288
krauss@19564
   289
wenzelm@19585
   290
end