src/HOL/Tools/Function/function.ML
author wenzelm
Sun, 25 Apr 2010 15:52:03 +0200
changeset 36323 655e2d74de3a
parent 35756 cfde251d03a5
child 36518 a33b986f2e22
permissions -rw-r--r--
modernized naming conventions of main Isar proof elements;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Function/fundef.ML
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     2
    Author:     Alexander Krauss, TU Muenchen
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     3
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     4
A package for general recursive function definitions.
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     5
Isar commands.
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     6
*)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     7
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     8
signature FUNCTION =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     9
sig
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    10
  include FUNCTION_DATA
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    11
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    12
  val add_function: (binding * typ option * mixfix) list ->
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    13
    (Attrib.binding * term) list -> Function_Common.function_config ->
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    14
    local_theory -> Proof.state
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
    15
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    16
  val add_function_cmd: (binding * string option * mixfix) list ->
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    17
    (Attrib.binding * string) list -> Function_Common.function_config ->
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    18
    local_theory -> Proof.state
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    19
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    20
  val termination_proof : term option -> local_theory -> Proof.state
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    21
  val termination_proof_cmd : string option -> local_theory -> Proof.state
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    22
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    23
  val setup : theory -> theory
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    24
  val get_congs : Proof.context -> thm list
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
    25
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    26
  val get_info : Proof.context -> term -> info
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    27
end
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    28
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    29
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    30
structure Function : FUNCTION =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    31
struct
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    32
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    33
open Function_Lib
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    34
open Function_Common
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    35
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    36
val simp_attribs = map (Attrib.internal o K)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    37
  [Simplifier.simp_add,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    38
   Code.add_default_eqn_attribute,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    39
   Nitpick_Simps.add]
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    40
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    41
val psimp_attribs = map (Attrib.internal o K)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    42
  [Simplifier.simp_add,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    43
   Nitpick_Psimps.add]
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    44
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    45
fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    46
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    47
fun add_simps fnames post sort extra_qualify label mod_binding moreatts
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    48
  simps lthy =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    49
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    50
    val spec = post simps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    51
      |> map (apfst (apsnd (fn ats => moreatts @ ats)))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    52
      |> map (apfst (apfst extra_qualify))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    53
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    54
    val (saved_spec_simps, lthy) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    55
      fold_map Local_Theory.note spec lthy
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    56
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    57
    val saved_simps = maps snd saved_spec_simps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    58
    val simps_by_f = sort saved_simps
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    59
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    60
    fun add_for_f fname simps =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    61
      Local_Theory.note
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    62
        ((mod_binding (Binding.qualify true fname (Binding.name label)), []), simps)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    63
      #> snd
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    64
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    65
    (saved_simps, fold2 add_for_f fnames simps_by_f lthy)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    66
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    67
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    68
fun gen_add_function is_external prep default_constraint fixspec eqns config lthy =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    69
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    70
    val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    71
    val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    72
    val fixes = map (apfst (apfst Binding.name_of)) fixes0;
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    73
    val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    74
    val (eqs, post, sort_cont, cnames) = get_preproc lthy config ctxt' fixes spec
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    75
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    76
    val defname = mk_defname fixes
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    77
    val FunctionConfig {partials, ...} = config
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    78
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    79
    val ((goalstate, cont), lthy) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    80
      Function_Mutual.prepare_function_mutual config defname fixes eqs lthy
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    81
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    82
    fun afterqed [[proof]] lthy =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    83
      let
35756
cfde251d03a5 refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
bulwahn
parents: 35410
diff changeset
    84
        val FunctionResult {fs, R, psimps, trsimps, simple_pinducts,
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    85
          termination, domintros, cases, ...} =
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    86
          cont (Thm.close_derivation proof)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    87
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    88
        val fnames = map (fst o fst) fixes
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    89
        fun qualify n = Binding.name n
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    90
          |> Binding.qualify true defname
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    91
        val conceal_partial = if partials then I else Binding.conceal
33394
9c6980f2eb39 conceal "termination" rule, used only by special tools
krauss
parents: 33369
diff changeset
    92
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    93
        val addsmps = add_simps fnames post sort_cont
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    94
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    95
        val (((psimps', pinducts'), (_, [termination'])), lthy) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    96
          lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    97
          |> addsmps (conceal_partial o Binding.qualify false "partial")
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    98
               "psimps" conceal_partial psimp_attribs psimps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    99
          ||> fold_option (snd oo addsmps I "simps" I simp_attribs) trsimps
35756
cfde251d03a5 refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
bulwahn
parents: 35410
diff changeset
   100
          ||> fold_option (Spec_Rules.add Spec_Rules.Equational o pair fs) trsimps
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   101
          ||>> Local_Theory.note ((conceal_partial (qualify "pinduct"),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   102
                 [Attrib.internal (K (Rule_Cases.case_names cnames)),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   103
                  Attrib.internal (K (Rule_Cases.consumes 1)),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   104
                  Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   105
          ||>> Local_Theory.note ((Binding.conceal (qualify "termination"), []), [termination])
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   106
          ||> (snd o Local_Theory.note ((qualify "cases",
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   107
                 [Attrib.internal (K (Rule_Cases.case_names cnames))]), [cases]))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   108
          ||> fold_option (snd oo curry Local_Theory.note (qualify "domintros", [])) domintros
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   109
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   110
        val info = { add_simps=addsmps, case_names=cnames, psimps=psimps',
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   111
          pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination',
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   112
          fs=fs, R=R, defname=defname, is_partial=true }
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   113
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   114
        val _ =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   115
          if not is_external then ()
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   116
          else Specification.print_consts lthy (K false) (map fst fixes)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   117
      in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   118
        lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   119
        |> Local_Theory.declaration false (add_function_data o morph_function_data info)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   120
      end
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   121
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   122
    lthy
36323
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 35756
diff changeset
   123
    |> Proof.theorem NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   124
    |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   125
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   126
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   127
val add_function =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   128
  gen_add_function false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   129
val add_function_cmd = gen_add_function true Specification.read_spec "_::type"
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34950
diff changeset
   130
                                                
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   131
fun gen_termination_proof prep_term raw_term_opt lthy =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   132
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   133
    val term_opt = Option.map (prep_term lthy) raw_term_opt
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   134
    val info = the (case term_opt of
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   135
                      SOME t => (import_function_data t lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   136
                        handle Option.Option =>
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   137
                          error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   138
                    | NONE => (import_last_function lthy handle Option.Option => error "Not a function"))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   139
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   140
      val { termination, fs, R, add_simps, case_names, psimps,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   141
        pinducts, defname, ...} = info
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   142
      val domT = domain_type (fastype_of R)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   143
      val goal = HOLogic.mk_Trueprop
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   144
                   (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   145
      fun afterqed [[totality]] lthy =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   146
        let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   147
          val totality = Thm.close_derivation totality
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   148
          val remove_domain_condition =
35410
1ea89d2a1bd4 more antiquotations;
wenzelm
parents: 35324
diff changeset
   149
            full_simplify (HOL_basic_ss addsimps [totality, @{thm True_implies_equals}])
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   150
          val tsimps = map remove_domain_condition psimps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   151
          val tinduct = map remove_domain_condition pinducts
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   152
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   153
          fun qualify n = Binding.name n
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   154
            |> Binding.qualify true defname
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   155
        in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   156
          lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   157
          |> add_simps I "simps" I simp_attribs tsimps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   158
          ||>> Local_Theory.note
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   159
             ((qualify "induct",
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   160
               [Attrib.internal (K (Rule_Cases.case_names case_names))]),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   161
              tinduct)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   162
          |-> (fn (simps, (_, inducts)) =>
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   163
            let val info' = { is_partial=false, defname=defname, add_simps=add_simps,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   164
              case_names=case_names, fs=fs, R=R, psimps=psimps, pinducts=pinducts,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   165
              simps=SOME simps, inducts=SOME inducts, termination=termination }
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   166
            in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   167
              Local_Theory.declaration false (add_function_data o morph_function_data info')
35756
cfde251d03a5 refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
bulwahn
parents: 35410
diff changeset
   168
              #> Spec_Rules.add Spec_Rules.Equational (fs, tsimps)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   169
            end)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   170
        end
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   171
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   172
    lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   173
    |> ProofContext.note_thmss ""
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   174
       [((Binding.empty, [Context_Rules.rule_del]), [([allI], [])])] |> snd
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   175
    |> ProofContext.note_thmss ""
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   176
       [((Binding.empty, [Context_Rules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   177
    |> ProofContext.note_thmss ""
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   178
       [((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   179
         [([Goal.norm_result termination], [])])] |> snd
36323
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 35756
diff changeset
   180
    |> Proof.theorem NONE afterqed [[(goal, [])]]
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   181
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   182
33726
0878aecbf119 eliminated slightly odd name space grouping -- now managed by Isar toplevel;
wenzelm
parents: 33671
diff changeset
   183
val termination_proof = gen_termination_proof Syntax.check_term
0878aecbf119 eliminated slightly odd name space grouping -- now managed by Isar toplevel;
wenzelm
parents: 33671
diff changeset
   184
val termination_proof_cmd = gen_termination_proof Syntax.read_term
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   185
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   186
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   187
(* Datatype hook to declare datatype congs as "function_congs" *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   188
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   189
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   190
fun add_case_cong n thy =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   191
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   192
    val cong = #case_cong (Datatype.the_info thy n)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   193
      |> safe_mk_meta_eq
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   194
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   195
    Context.theory_map
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   196
      (Function_Ctx_Tree.map_function_congs (Thm.add_thm cong)) thy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   197
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   198
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   199
val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   200
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   201
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   202
(* setup *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   203
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   204
val setup =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   205
  Attrib.setup @{binding fundef_cong}
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   206
    (Attrib.add_del Function_Ctx_Tree.cong_add Function_Ctx_Tree.cong_del)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   207
    "declaration of congruence rule for function definitions"
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   208
  #> setup_case_cong
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   209
  #> Function_Relation.setup
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   210
  #> Function_Common.Termination_Simps.setup
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   211
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   212
val get_congs = Function_Ctx_Tree.get_function_congs
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   213
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   214
fun get_info ctxt t = Item_Net.retrieve (get_function ctxt) t
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   215
  |> the_single |> snd
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   216
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   217
(* outer syntax *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   218
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   219
local structure P = OuterParse and K = OuterKeyword in
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   220
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   221
val _ =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   222
  OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   223
  (function_parser default_config
33726
0878aecbf119 eliminated slightly odd name space grouping -- now managed by Isar toplevel;
wenzelm
parents: 33671
diff changeset
   224
     >> (fn ((config, fixes), statements) => add_function_cmd fixes statements config))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   225
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   226
val _ =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   227
  OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
33726
0878aecbf119 eliminated slightly odd name space grouping -- now managed by Isar toplevel;
wenzelm
parents: 33671
diff changeset
   228
  (Scan.option P.term >> termination_proof_cmd)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   229
33726
0878aecbf119 eliminated slightly odd name space grouping -- now managed by Isar toplevel;
wenzelm
parents: 33671
diff changeset
   230
end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   231
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   232
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   233
end