src/HOL/Tools/Function/function.ML
author krauss
Wed, 28 Apr 2010 16:13:17 +0200
changeset 36522 e80a95279ef6
parent 36520 772ed73e1d61
child 36547 2a9d0ec8c10d
permissions -rw-r--r--
return info record (relative to auxiliary context!)
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
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    12
  val add_function: (binding * typ option * mixfix) list ->
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    13
    (Attrib.binding * term) list -> Function_Common.function_config ->
36522
e80a95279ef6 return info record (relative to auxiliary context!)
krauss
parents: 36520
diff changeset
    14
    (Proof.context -> tactic) -> local_theory -> info * local_theory
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    15
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    16
  val add_function_cmd: (binding * string option * mixfix) list ->
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    17
    (Attrib.binding * string) list -> Function_Common.function_config ->
36522
e80a95279ef6 return info record (relative to auxiliary context!)
krauss
parents: 36520
diff changeset
    18
    (Proof.context -> tactic) -> local_theory -> info * local_theory
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    19
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
    20
  val function: (binding * typ option * mixfix) list ->
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    21
    (Attrib.binding * term) list -> Function_Common.function_config ->
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    22
    local_theory -> Proof.state
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
    23
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
    24
  val function_cmd: (binding * string option * mixfix) list ->
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    25
    (Attrib.binding * string) list -> Function_Common.function_config ->
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    26
    local_theory -> Proof.state
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    27
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    28
  val prove_termination: term option -> tactic -> local_theory -> local_theory
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    29
  val prove_termination_cmd: string option -> tactic -> local_theory -> local_theory
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    30
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
    31
  val termination : term option -> local_theory -> Proof.state
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
    32
  val termination_cmd : string option -> local_theory -> Proof.state
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    33
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    34
  val setup : theory -> theory
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    35
  val get_congs : Proof.context -> thm list
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
    36
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    37
  val get_info : Proof.context -> term -> info
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    38
end
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    39
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
structure Function : FUNCTION =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    42
struct
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    43
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    44
open Function_Lib
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    45
open Function_Common
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    46
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    47
val simp_attribs = map (Attrib.internal o K)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    48
  [Simplifier.simp_add,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    49
   Code.add_default_eqn_attribute,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    50
   Nitpick_Simps.add]
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    51
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    52
val psimp_attribs = map (Attrib.internal o K)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    53
  [Simplifier.simp_add,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    54
   Nitpick_Psimps.add]
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    55
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    56
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
    57
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    58
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
    59
  simps lthy =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    60
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    61
    val spec = post simps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    62
      |> map (apfst (apsnd (fn ats => moreatts @ ats)))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    63
      |> map (apfst (apfst extra_qualify))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    64
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    65
    val (saved_spec_simps, lthy) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    66
      fold_map Local_Theory.note spec lthy
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    67
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    68
    val saved_simps = maps snd saved_spec_simps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    69
    val simps_by_f = sort saved_simps
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    70
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    71
    fun add_for_f fname simps =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    72
      Local_Theory.note
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    73
        ((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
    74
      #> snd
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    75
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    76
    (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
    77
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    78
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
    79
fun prepare_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
    80
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    81
    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
    82
    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
    83
    val fixes = map (apfst (apfst Binding.name_of)) fixes0;
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    84
    val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    85
    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
    86
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    87
    val defname = mk_defname fixes
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    88
    val FunctionConfig {partials, ...} = config
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    89
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
    90
    val ((goal_state, cont), lthy') =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    91
      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
    92
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    93
    fun afterqed [[proof]] lthy =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    94
      let
35756
cfde251d03a5 refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
bulwahn
parents: 35410
diff changeset
    95
        val FunctionResult {fs, R, psimps, trsimps, simple_pinducts,
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    96
          termination, domintros, cases, ...} =
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    97
          cont (Thm.close_derivation proof)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    98
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    99
        val fnames = map (fst o fst) fixes
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   100
        fun qualify n = Binding.name n
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   101
          |> Binding.qualify true defname
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   102
        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
   103
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   104
        val addsmps = add_simps fnames post sort_cont
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   105
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   106
        val (((psimps', pinducts'), (_, [termination'])), lthy) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   107
          lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   108
          |> addsmps (conceal_partial o Binding.qualify false "partial")
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   109
               "psimps" conceal_partial psimp_attribs psimps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   110
          ||> 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
   111
          ||> 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
   112
          ||>> Local_Theory.note ((conceal_partial (qualify "pinduct"),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   113
                 [Attrib.internal (K (Rule_Cases.case_names cnames)),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   114
                  Attrib.internal (K (Rule_Cases.consumes 1)),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   115
                  Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   116
          ||>> Local_Theory.note ((Binding.conceal (qualify "termination"), []), [termination])
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   117
          ||> (snd o Local_Theory.note ((qualify "cases",
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   118
                 [Attrib.internal (K (Rule_Cases.case_names cnames))]), [cases]))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   119
          ||> 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
   120
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   121
        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
   122
          pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination',
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   123
          fs=fs, R=R, defname=defname, is_partial=true }
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   124
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   125
        val _ =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   126
          if not is_external then ()
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   127
          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
   128
      in
36522
e80a95279ef6 return info record (relative to auxiliary context!)
krauss
parents: 36520
diff changeset
   129
        (info, 
e80a95279ef6 return info record (relative to auxiliary context!)
krauss
parents: 36520
diff changeset
   130
         lthy |> Local_Theory.declaration false (add_function_data o morph_function_data info))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   131
      end
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   132
  in
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   133
    ((goal_state, afterqed), lthy')
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   134
  end
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   135
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   136
fun gen_add_function is_external prep default_constraint fixspec eqns config tac lthy =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   137
  let
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   138
    val ((goal_state, afterqed), lthy') =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   139
      prepare_function is_external prep default_constraint fixspec eqns config lthy
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   140
    val pattern_thm =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   141
      case SINGLE (tac lthy') goal_state of
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   142
        NONE => error "pattern completeness and compatibility proof failed"
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   143
      | SOME st => Goal.finish lthy' st
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   144
  in
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   145
    lthy'
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   146
    |> afterqed [[pattern_thm]]
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   147
  end
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   148
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   149
val add_function =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   150
  gen_add_function false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   151
val add_function_cmd = gen_add_function true Specification.read_spec "_::type"
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   152
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   153
fun gen_function is_external prep default_constraint fixspec eqns config lthy =
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   154
  let
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   155
    val ((goal_state, afterqed), lthy') =
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   156
      prepare_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
   157
  in
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   158
    lthy'
36522
e80a95279ef6 return info record (relative to auxiliary context!)
krauss
parents: 36520
diff changeset
   159
    |> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (concl_of goal_state), [])]]
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   160
    |> Proof.refine (Method.primitive_text (K goal_state)) |> Seq.hd
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   161
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   162
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   163
val function =
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   164
  gen_function false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   165
val function_cmd = gen_function true Specification.read_spec "_::type"
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   166
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   167
fun prepare_termination_proof prep_term raw_term_opt lthy =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   168
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   169
    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
   170
    val info = the (case term_opt of
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   171
                      SOME t => (import_function_data t lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   172
                        handle Option.Option =>
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   173
                          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
   174
                    | 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
   175
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   176
      val { termination, fs, R, add_simps, case_names, psimps,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   177
        pinducts, defname, ...} = info
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   178
      val domT = domain_type (fastype_of R)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   179
      val goal = HOLogic.mk_Trueprop
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   180
                   (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
   181
      fun afterqed [[totality]] lthy =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   182
        let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   183
          val totality = Thm.close_derivation totality
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   184
          val remove_domain_condition =
35410
1ea89d2a1bd4 more antiquotations;
wenzelm
parents: 35324
diff changeset
   185
            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
   186
          val tsimps = map remove_domain_condition psimps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   187
          val tinduct = map remove_domain_condition pinducts
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   188
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   189
          fun qualify n = Binding.name n
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   190
            |> Binding.qualify true defname
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   191
        in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   192
          lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   193
          |> add_simps I "simps" I simp_attribs tsimps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   194
          ||>> Local_Theory.note
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   195
             ((qualify "induct",
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   196
               [Attrib.internal (K (Rule_Cases.case_names case_names))]),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   197
              tinduct)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   198
          |-> (fn (simps, (_, inducts)) =>
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   199
            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
   200
              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
   201
              simps=SOME simps, inducts=SOME inducts, termination=termination }
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   202
            in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   203
              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
   204
              #> Spec_Rules.add Spec_Rules.Equational (fs, tsimps)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   205
            end)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   206
        end
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   207
  in
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   208
    (goal, afterqed, termination)
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   209
  end
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   210
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   211
fun gen_prove_termination prep_term raw_term_opt tac lthy =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   212
  let
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   213
    val (goal, afterqed, termination) =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   214
      prepare_termination_proof prep_term raw_term_opt lthy
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   215
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   216
    val totality = Goal.prove lthy [] [] goal (K tac)
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   217
  in
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   218
    afterqed [[totality]] lthy
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   219
end
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   220
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   221
val prove_termination = gen_prove_termination Syntax.check_term
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   222
val prove_termination_cmd = gen_prove_termination Syntax.read_term
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   223
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   224
fun gen_termination prep_term raw_term_opt lthy =
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   225
  let
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   226
    val (goal, afterqed, termination) = prepare_termination_proof prep_term raw_term_opt lthy
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   227
  in
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   228
    lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   229
    |> ProofContext.note_thmss ""
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   230
       [((Binding.empty, [Context_Rules.rule_del]), [([allI], [])])] |> snd
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   231
    |> ProofContext.note_thmss ""
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   232
       [((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
   233
    |> ProofContext.note_thmss ""
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   234
       [((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   235
         [([Goal.norm_result termination], [])])] |> snd
36323
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 35756
diff changeset
   236
    |> Proof.theorem NONE afterqed [[(goal, [])]]
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   237
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   238
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   239
val termination = gen_termination Syntax.check_term
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   240
val termination_cmd = gen_termination Syntax.read_term
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   241
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   242
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   243
(* Datatype hook to declare datatype congs as "function_congs" *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   244
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   245
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   246
fun add_case_cong n thy =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   247
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   248
    val cong = #case_cong (Datatype.the_info thy n)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   249
      |> safe_mk_meta_eq
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   250
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   251
    Context.theory_map
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   252
      (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
   253
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   254
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   255
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
   256
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   257
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   258
(* setup *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   259
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   260
val setup =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   261
  Attrib.setup @{binding fundef_cong}
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   262
    (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
   263
    "declaration of congruence rule for function definitions"
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   264
  #> setup_case_cong
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   265
  #> Function_Relation.setup
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   266
  #> Function_Common.Termination_Simps.setup
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   267
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   268
val get_congs = Function_Ctx_Tree.get_function_congs
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   269
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   270
fun get_info ctxt t = Item_Net.retrieve (get_function ctxt) t
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   271
  |> the_single |> snd
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   272
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   273
(* outer syntax *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   274
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   275
local structure P = OuterParse and K = OuterKeyword in
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   276
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   277
val _ =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   278
  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
   279
  (function_parser default_config
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   280
     >> (fn ((config, fixes), statements) => function_cmd fixes statements config))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   281
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   282
val _ =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   283
  OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   284
  (Scan.option P.term >> termination_cmd)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   285
33726
0878aecbf119 eliminated slightly odd name space grouping -- now managed by Isar toplevel;
wenzelm
parents: 33671
diff changeset
   286
end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   287
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   288
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   289
end