src/HOL/Tools/function_package/fundef_package.ML
author bulwahn
Sat, 16 May 2009 20:17:59 +0200
changeset 31174 f1f1e9b53c81
parent 31172 74d72ba262fb
child 31177 c39994cb152a
permissions -rw-r--r--
added new kind generated_theorem for theorems which are generated by packages to distinguish between theorems from users and packages
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Tools/function_package/fundef_package.ML
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     2
    Author:     Alexander Krauss, TU Muenchen
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     3
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
     4
A package for general recursive function definitions.
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     5
Isar commands.
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     6
*)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     7
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
     8
signature FUNDEF_PACKAGE =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     9
sig
30493
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
    10
    val add_fundef :  (binding * typ option * mixfix) list
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
    11
                       -> (Attrib.binding * term) list
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
    12
                       -> FundefCommon.fundef_config
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
    13
                       -> local_theory
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
    14
                       -> Proof.state
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
    15
    val add_fundef_cmd :  (binding * string option * mixfix) list
30789
b6f6948bdcf1 code attribute for tailrec-equations, too; tuned
krauss
parents: 30788
diff changeset
    16
                      -> (Attrib.binding * string) list
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    17
                      -> FundefCommon.fundef_config
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    18
                      -> local_theory
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
    19
                      -> Proof.state
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    20
30493
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
    21
    val termination_proof : term option -> local_theory -> Proof.state
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
    22
    val termination_proof_cmd : string option -> local_theory -> Proof.state
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
    23
    val termination : term option -> local_theory -> Proof.state
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
    24
    val termination_cmd : string option -> local_theory -> Proof.state
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    25
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    26
    val setup : theory -> theory
30493
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
    27
    val get_congs : Proof.context -> thm list
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    28
end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    29
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    30
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    31
structure FundefPackage : FUNDEF_PACKAGE =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    32
struct
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    33
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 21000
diff changeset
    34
open FundefLib
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    35
open FundefCommon
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    36
30789
b6f6948bdcf1 code attribute for tailrec-equations, too; tuned
krauss
parents: 30788
diff changeset
    37
val simp_attribs = map (Attrib.internal o K)
b6f6948bdcf1 code attribute for tailrec-equations, too; tuned
krauss
parents: 30788
diff changeset
    38
    [Simplifier.simp_add,
b6f6948bdcf1 code attribute for tailrec-equations, too; tuned
krauss
parents: 30788
diff changeset
    39
     Code.add_default_eqn_attribute,
31172
74d72ba262fb added collection of simplification rules of recursive functions for quickcheck
bulwahn
parents: 30907
diff changeset
    40
     Nitpick_Const_Simp_Thms.add,
74d72ba262fb added collection of simplification rules of recursive functions for quickcheck
bulwahn
parents: 30907
diff changeset
    41
     Quickcheck_RecFun_Simp_Thms.add]
30789
b6f6948bdcf1 code attribute for tailrec-equations, too; tuned
krauss
parents: 30788
diff changeset
    42
b6f6948bdcf1 code attribute for tailrec-equations, too; tuned
krauss
parents: 30788
diff changeset
    43
val psimp_attribs = map (Attrib.internal o K)
b6f6948bdcf1 code attribute for tailrec-equations, too; tuned
krauss
parents: 30788
diff changeset
    44
    [Simplifier.simp_add,
b6f6948bdcf1 code attribute for tailrec-equations, too; tuned
krauss
parents: 30788
diff changeset
    45
     Nitpick_Const_Psimp_Thms.add]
b6f6948bdcf1 code attribute for tailrec-equations, too; tuned
krauss
parents: 30788
diff changeset
    46
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
    47
fun note_theorem ((name, atts), ths) =
31174
f1f1e9b53c81 added new kind generated_theorem for theorems which are generated by packages to distinguish between theorems from users and packages
bulwahn
parents: 31172
diff changeset
    48
  LocalTheory.note Thm.generated_theoremK ((Binding.qualified_name name, atts), ths)
21435
883337ea2f3b LocalTheory.axioms/notes/defs: proper kind;
wenzelm
parents: 21391
diff changeset
    49
30789
b6f6948bdcf1 code attribute for tailrec-equations, too; tuned
krauss
parents: 30788
diff changeset
    50
fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    51
26529
03ad378ed5f0 Function package no longer overwrites theorems.
krauss
parents: 26171
diff changeset
    52
fun add_simps fnames post sort extra_qualify label moreatts simps lthy =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    53
    let
23819
2040846d1bbe some interface cleanup
krauss
parents: 23473
diff changeset
    54
      val spec = post simps
30789
b6f6948bdcf1 code attribute for tailrec-equations, too; tuned
krauss
parents: 30788
diff changeset
    55
                   |> map (apfst (apsnd (fn ats => moreatts @ ats)))
26529
03ad378ed5f0 Function package no longer overwrites theorems.
krauss
parents: 26171
diff changeset
    56
                   |> map (apfst (apfst extra_qualify))
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
    57
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    58
      val (saved_spec_simps, lthy) =
31174
f1f1e9b53c81 added new kind generated_theorem for theorems which are generated by packages to distinguish between theorems from users and packages
bulwahn
parents: 31172
diff changeset
    59
        fold_map (LocalTheory.note Thm.generated_theoremK) spec lthy
23819
2040846d1bbe some interface cleanup
krauss
parents: 23473
diff changeset
    60
22498
62cdd4b3e96b made function syntax strict, requiring | to separate equations; cleanup
krauss
parents: 22325
diff changeset
    61
      val saved_simps = flat (map snd saved_spec_simps)
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    62
      val simps_by_f = sort saved_simps
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    63
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    64
      fun add_for_f fname simps =
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30223
diff changeset
    65
        note_theorem ((Long_Name.qualify fname label, []), simps) #> snd
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    66
    in
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    67
      (saved_simps,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    68
       fold2 add_for_f fnames simps_by_f lthy)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    69
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    70
30787
5b7a5a05c7aa remove "otherwise" feature from function package which was never really documented or used
krauss
parents: 30761
diff changeset
    71
fun gen_add_fundef is_external prep default_constraint fixspec eqns config lthy =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    72
    let
29051
b9c5726e79ab constrain type inference to sort "type"
krauss
parents: 29006
diff changeset
    73
      val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
30486
9cdc7ce0e389 simplified preparation and outer parsing of specification;
wenzelm
parents: 30435
diff changeset
    74
      val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 29866
diff changeset
    75
      val fixes = map (apfst (apfst Binding.name_of)) fixes0;
30790
350bb108406d bstring -> binding
krauss
parents: 30789
diff changeset
    76
      val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
30787
5b7a5a05c7aa remove "otherwise" feature from function package which was never really documented or used
krauss
parents: 30761
diff changeset
    77
      val (eqs, post, sort_cont, cnames) = FundefCommon.get_preproc lthy config ctxt' fixes spec
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    78
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    79
      val defname = mk_defname fixes
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    80
23819
2040846d1bbe some interface cleanup
krauss
parents: 23473
diff changeset
    81
      val ((goalstate, cont), lthy) =
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    82
          FundefMutual.prepare_fundef_mutual config defname fixes eqs lthy
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    83
30907
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    84
      fun afterqed [[proof]] lthy =
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    85
        let
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    86
          val FundefResult {fs, R, psimps, trsimps,  simple_pinducts, termination,
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    87
                            domintros, cases, ...} =
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    88
          cont (Thm.close_derivation proof)
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    89
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    90
          val fnames = map (fst o fst) fixes
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    91
          val qualify = Long_Name.qualify defname
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    92
          val addsmps = add_simps fnames post sort_cont
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    93
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    94
          val (((psimps', pinducts'), (_, [termination'])), lthy) =
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    95
            lthy
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    96
            |> addsmps (Binding.qualify false "partial") "psimps"
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    97
                 psimp_attribs psimps
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    98
            ||> fold_option (snd oo addsmps I "simps" simp_attribs) trsimps
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
    99
            ||>> note_theorem ((qualify "pinduct",
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   100
                   [Attrib.internal (K (RuleCases.case_names cnames)),
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   101
                    Attrib.internal (K (RuleCases.consumes 1)),
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   102
                    Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   103
            ||>> note_theorem ((qualify "termination", []), [termination])
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   104
            ||> (snd o note_theorem ((qualify "cases",
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   105
                   [Attrib.internal (K (RuleCases.case_names cnames))]), [cases]))
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   106
            ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   107
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   108
          val cdata = FundefCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps',
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   109
                                      pinducts=snd pinducts', termination=termination',
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   110
                                      fs=fs, R=R, defname=defname }
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   111
          val _ =
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   112
            if not is_external then ()
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   113
            else Specification.print_consts lthy (K false) (map fst fixes)
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   114
        in
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   115
          lthy
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   116
          |> LocalTheory.declaration (add_fundef_data o morph_fundef_data cdata)
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   117
        end
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   118
    in
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   119
      lthy
26594
1c676ae50311 fundef_afterqed: removed unused config, added do_print flag;
wenzelm
parents: 26580
diff changeset
   120
        |> is_external ? LocalTheory.set_group (serial_string ())
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   121
        |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   122
        |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   123
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   124
30493
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   125
val add_fundef = gen_add_fundef false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   126
val add_fundef_cmd = gen_add_fundef true Specification.read_spec "_::type"
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   127
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   128
fun gen_termination_proof prep_term raw_term_opt lthy =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   129
    let
30493
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   130
      val term_opt = Option.map (prep_term lthy) raw_term_opt
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   131
      val data = the (case term_opt of
30493
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   132
                        SOME t => (import_fundef_data t lthy
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   133
                          handle Option.Option =>
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   134
                            error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   135
                      | NONE => (import_last_fundef lthy handle Option.Option => error "Not a function"))
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   136
30907
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   137
        val FundefCtxData { termination, R, add_simps, case_names, psimps,
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   138
                            pinducts, defname, ...} = data
22325
be61bd159a99 changed termination goal to use object quantifier
krauss
parents: 22170
diff changeset
   139
        val domT = domain_type (fastype_of R)
30907
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   140
        val goal = HOLogic.mk_Trueprop
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   141
                     (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   142
        fun afterqed [[totality]] lthy =
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   143
          let
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   144
            val totality = Thm.close_derivation totality
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   145
            val remove_domain_condition =
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   146
              full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   147
            val tsimps = map remove_domain_condition psimps
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   148
            val tinduct = map remove_domain_condition pinducts
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   149
            val qualify = Long_Name.qualify defname;
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   150
          in
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   151
            lthy
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   152
            |> add_simps I "simps" simp_attribs tsimps |> snd
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   153
            |> note_theorem
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   154
               ((qualify "induct",
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   155
                 [Attrib.internal (K (RuleCases.case_names case_names))]),
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   156
                tinduct) |> snd
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   157
          end
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   158
    in
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
   159
      lthy
30907
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   160
      |> ProofContext.note_thmss ""
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   161
         [((Binding.empty, [ContextRules.rule_del]), [([allI], [])])] |> snd
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   162
      |> ProofContext.note_thmss ""
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   163
         [((Binding.empty, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   164
      |> ProofContext.note_thmss ""
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   165
         [((Binding.name "termination", [ContextRules.intro_bang (SOME 0)]),
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   166
           [([Goal.norm_result termination], [])])] |> snd
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   167
      |> Proof.theorem_i NONE afterqed [[(goal, [])]]
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   168
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   169
30493
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   170
val termination_proof = gen_termination_proof Syntax.check_term;
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   171
val termination_proof_cmd = gen_termination_proof Syntax.read_term;
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   172
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   173
fun termination term_opt lthy =
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   174
  lthy
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   175
  |> LocalTheory.set_group (serial_string ())
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   176
  |> termination_proof term_opt;
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   177
26171
5426a823455c more precise handling of "group" for termination;
wenzelm
parents: 26129
diff changeset
   178
fun termination_cmd term_opt lthy =
5426a823455c more precise handling of "group" for termination;
wenzelm
parents: 26129
diff changeset
   179
  lthy
5426a823455c more precise handling of "group" for termination;
wenzelm
parents: 26129
diff changeset
   180
  |> LocalTheory.set_group (serial_string ())
30493
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   181
  |> termination_proof_cmd term_opt;
19611
da2a0014c461 Function Package: Quick-and-dirty-fixed strange "Proved a different theorem bug"
krauss
parents: 19585
diff changeset
   182
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   183
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   184
(* Datatype hook to declare datatype congs as "fundef_congs" *)
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   185
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   186
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   187
fun add_case_cong n thy =
24168
86a03a092062 simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
krauss
parents: 24039
diff changeset
   188
    Context.theory_map (FundefCtxTree.map_fundef_congs (Thm.add_thm
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   189
                          (DatatypePackage.get_datatype thy n |> the
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   190
                           |> #case_cong
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   191
                           |> safe_mk_meta_eq)))
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   192
                       thy
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   193
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
   194
val case_cong = fold add_case_cong
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   195
24711
e8bba7723858 simplified interpretation setup;
wenzelm
parents: 24626
diff changeset
   196
val setup_case_cong = DatatypePackage.interpretation case_cong
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
   197
30493
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   198
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   199
(* setup *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   200
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   201
val setup =
30907
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   202
  Attrib.setup @{binding fundef_cong}
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   203
    (Attrib.add_del FundefCtxTree.cong_add FundefCtxTree.cong_del)
63b8b2b52f56 inlined afterqeds to improve clarity; tuned
krauss
parents: 30790
diff changeset
   204
    "declaration of congruence rule for function definitions"
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
   205
  #> setup_case_cong
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22733
diff changeset
   206
  #> FundefRelation.setup
26749
397a1aeede7d * New attribute "termination_simp": Simp rules for termination proofs
krauss
parents: 26628
diff changeset
   207
  #> FundefCommon.TerminationSimps.setup
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   208
30493
b8570ea1ce25 provide regular ML interfaces for Isar source language elements;
wenzelm
parents: 30486
diff changeset
   209
val get_congs = FundefCtxTree.get_fundef_congs
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
   210
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   211
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   212
(* outer syntax *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   213
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   214
local structure P = OuterParse and K = OuterKeyword in
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   215
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24861
diff changeset
   216
val _ =
26989
9b2acb536228 uniform treatment of target, not as config;
wenzelm
parents: 26749
diff changeset
   217
  OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
22498
62cdd4b3e96b made function syntax strict, requiring | to separate equations; cleanup
krauss
parents: 22325
diff changeset
   218
  (fundef_parser default_config
30787
5b7a5a05c7aa remove "otherwise" feature from function package which was never really documented or used
krauss
parents: 30761
diff changeset
   219
     >> (fn ((config, fixes), statements) => add_fundef_cmd fixes statements config));
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   220
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24861
diff changeset
   221
val _ =
26989
9b2acb536228 uniform treatment of target, not as config;
wenzelm
parents: 26749
diff changeset
   222
  OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
9b2acb536228 uniform treatment of target, not as config;
wenzelm
parents: 26749
diff changeset
   223
  (Scan.option P.term >> termination_cmd);
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   224
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   225
end;
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   226
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   227
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19583
diff changeset
   228
end