src/HOL/Tools/Function/function.ML
author Lars Hupel <lars.hupel@mytum.de>
Wed, 05 Apr 2017 10:26:28 +0200
changeset 65387 5dbe02addca5
parent 63239 d562c9948dee
child 66251 cd935b7cb3fb
permissions -rw-r--r--
store totality fact in function info
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37744
3daaf23b9ab4 tuned titles
haftmann
parents: 37145
diff changeset
     1
(*  Title:      HOL/Tools/Function/function.ML
33099
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
41114
f9ae7c2abf7e tuned headers
krauss
parents: 40076
diff changeset
     4
Main entry points to the function package.
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     5
*)
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
signature FUNCTION =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     8
sig
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
     9
  include FUNCTION_DATA
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    10
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    11
  val add_function: (binding * typ option * mixfix) list ->
63064
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
    12
    Specification.multi_specs -> Function_Common.function_config ->
36522
e80a95279ef6 return info record (relative to auxiliary context!)
krauss
parents: 36520
diff changeset
    13
    (Proof.context -> tactic) -> local_theory -> info * local_theory
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    14
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    15
  val add_function_cmd: (binding * string option * mixfix) list ->
63064
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
    16
    Specification.multi_specs_cmd -> Function_Common.function_config ->
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
    17
    (Proof.context -> tactic) -> bool -> local_theory -> info * local_theory
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    18
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
    19
  val function: (binding * typ option * mixfix) list ->
63064
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
    20
    Specification.multi_specs -> Function_Common.function_config ->
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    21
    local_theory -> Proof.state
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
    22
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
    23
  val function_cmd: (binding * string option * mixfix) list ->
63064
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
    24
    Specification.multi_specs_cmd -> Function_Common.function_config ->
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
    25
    bool -> local_theory -> Proof.state
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    26
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
    27
  val prove_termination: term option -> tactic -> local_theory ->
36547
2a9d0ec8c10d return updated info record after termination proof
krauss
parents: 36522
diff changeset
    28
    info * local_theory
2a9d0ec8c10d return updated info record after termination proof
krauss
parents: 36522
diff changeset
    29
  val prove_termination_cmd: string option -> tactic -> local_theory ->
2a9d0ec8c10d return updated info record after termination proof
krauss
parents: 36522
diff changeset
    30
    info * local_theory
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    31
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
    32
  val termination : term option -> local_theory -> Proof.state
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
    33
  val termination_cmd : string option -> local_theory -> Proof.state
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    34
34232
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
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45291
diff changeset
    47
val simp_attribs =
63239
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 63183
diff changeset
    48
  @{attributes [simp, nitpick_simp]} @ [Code.add_default_eqn_attrib Code.Equation]
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    49
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45291
diff changeset
    50
val psimp_attribs =
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45291
diff changeset
    51
  @{attributes [nitpick_psimp]}
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    52
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63011
diff changeset
    53
fun note_derived (a, atts) (fname, thms) =
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63011
diff changeset
    54
  Local_Theory.note ((derived_name fname a, atts), thms) #> apfst snd
53604
c1db98d7c66f clarified
krauss
parents: 53603
diff changeset
    55
63005
d743bb1b6c23 clarified bindings;
wenzelm
parents: 63004
diff changeset
    56
fun add_simps fnames post sort extra_qualify label mod_binding moreatts simps lthy =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    57
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    58
    val spec = post simps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    59
      |> map (apfst (apsnd (fn ats => moreatts @ ats)))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    60
      |> map (apfst (apfst extra_qualify))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    61
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    62
    val (saved_spec_simps, lthy) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    63
      fold_map Local_Theory.note spec lthy
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_simps = maps snd saved_spec_simps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    66
    val simps_by_f = sort saved_simps
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    67
63004
f507e6fe1d77 prefer binding over base name;
wenzelm
parents: 62996
diff changeset
    68
    fun note fname simps =
63005
d743bb1b6c23 clarified bindings;
wenzelm
parents: 63004
diff changeset
    69
      Local_Theory.note ((mod_binding (derived_name fname label), []), simps) #> snd
63004
f507e6fe1d77 prefer binding over base name;
wenzelm
parents: 62996
diff changeset
    70
  in (saved_simps, fold2 note fnames simps_by_f lthy) end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    71
62774
cfcb20bbdbd8 reconcile object-logic constraint vs. mixfix constraint;
wenzelm
parents: 61841
diff changeset
    72
fun prepare_function do_print prep fixspec eqns config lthy =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    73
  let
62996
wenzelm
parents: 62958
diff changeset
    74
    val ((fixes0, spec0), ctxt') = prep fixspec eqns lthy
wenzelm
parents: 62958
diff changeset
    75
    val fixes = map (apfst (apfst Binding.name_of)) fixes0
wenzelm
parents: 62958
diff changeset
    76
    val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    77
    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
    78
63004
f507e6fe1d77 prefer binding over base name;
wenzelm
parents: 62996
diff changeset
    79
    val fnames = map (fst o fst) fixes0
63006
89d19aa73081 clarified bindings;
wenzelm
parents: 63005
diff changeset
    80
    val defname = Binding.conglomerate fnames;
62996
wenzelm
parents: 62958
diff changeset
    81
41846
b368a7aee46a removed support for tail-recursion from function package (now implemented by partial_function)
krauss
parents: 41417
diff changeset
    82
    val FunctionConfig {partials, default, ...} = config
41417
211dbd42f95d function (default) is legacy feature
krauss
parents: 41405
diff changeset
    83
    val _ =
44052
00f0c8782a51 slightly more uniform messages;
wenzelm
parents: 43277
diff changeset
    84
      if is_some default
00f0c8782a51 slightly more uniform messages;
wenzelm
parents: 43277
diff changeset
    85
      then legacy_feature "\"function (default)\" -- use 'partial_function' instead"
41417
211dbd42f95d function (default) is legacy feature
krauss
parents: 41405
diff changeset
    86
      else ()
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    87
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
    88
    val ((goal_state, cont), lthy') =
63011
301e631666a0 clarified bindings;
wenzelm
parents: 63006
diff changeset
    89
      Function_Mutual.prepare_function_mutual config defname fixes0 eqs lthy
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    90
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    91
    fun afterqed [[proof]] lthy =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    92
      let
60643
9173467ec5b6 clarified context;
wenzelm
parents: 60499
diff changeset
    93
        val result = cont lthy (Thm.close_derivation proof)
52384
80c00a851de5 export dom predicate in the info record
krauss
parents: 52383
diff changeset
    94
        val FunctionResult {fs, R, dom, psimps, simple_pinducts,
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
    95
                termination, domintros, cases, ...} = result
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
    96
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
    97
        val pelims = Function_Elims.mk_partial_elim_rules lthy result
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    98
59859
f9d1442c70f3 tuned signature;
wenzelm
parents: 59582
diff changeset
    99
        val concealed_partial = if partials then I else Binding.concealed
33394
9c6980f2eb39 conceal "termination" rule, used only by special tools
krauss
parents: 33369
diff changeset
   100
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   101
        val addsmps = add_simps fnames post sort_cont
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   102
53604
c1db98d7c66f clarified
krauss
parents: 53603
diff changeset
   103
        val (((((psimps', [pinducts']), [termination']), cases'), pelims'), lthy) =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   104
          lthy
59859
f9d1442c70f3 tuned signature;
wenzelm
parents: 59582
diff changeset
   105
          |> addsmps (concealed_partial o Binding.qualify false "partial")
f9d1442c70f3 tuned signature;
wenzelm
parents: 59582
diff changeset
   106
               "psimps" concealed_partial psimp_attribs psimps
63005
d743bb1b6c23 clarified bindings;
wenzelm
parents: 63004
diff changeset
   107
          ||>> Local_Theory.notes [((concealed_partial (derived_name defname "pinduct"), []),
50771
2852f997bfb5 prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents: 49967
diff changeset
   108
                simple_pinducts |> map (fn th => ([th],
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63011
diff changeset
   109
                 [Attrib.case_names cnames, Attrib.consumes (1 - Thm.nprems_of th)] @
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63011
diff changeset
   110
                 @{attributes [induct pred]})))]
63004
f507e6fe1d77 prefer binding over base name;
wenzelm
parents: 62996
diff changeset
   111
          ||>> (apfst snd o
63005
d743bb1b6c23 clarified bindings;
wenzelm
parents: 63004
diff changeset
   112
            Local_Theory.note
d743bb1b6c23 clarified bindings;
wenzelm
parents: 63004
diff changeset
   113
              ((Binding.concealed (derived_name defname "termination"), []), [termination]))
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63011
diff changeset
   114
          ||>> fold_map (note_derived ("cases", [Attrib.case_names cnames]))
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63011
diff changeset
   115
            (fnames ~~ map single cases)
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63011
diff changeset
   116
          ||>> fold_map (note_derived ("pelims", [Attrib.consumes 1, Attrib.constraints 1]))
63004
f507e6fe1d77 prefer binding over base name;
wenzelm
parents: 62996
diff changeset
   117
            (fnames ~~ pelims)
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   118
          ||> (case domintros of NONE => I | SOME thms =>
63005
d743bb1b6c23 clarified bindings;
wenzelm
parents: 63004
diff changeset
   119
                Local_Theory.note ((derived_name defname "domintros", []), thms) #> snd)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   120
63004
f507e6fe1d77 prefer binding over base name;
wenzelm
parents: 62996
diff changeset
   121
        val info =
f507e6fe1d77 prefer binding over base name;
wenzelm
parents: 62996
diff changeset
   122
          { add_simps=addsmps, fnames=fnames, case_names=cnames, psimps=psimps',
65387
5dbe02addca5 store totality fact in function info
Lars Hupel <lars.hupel@mytum.de>
parents: 63239
diff changeset
   123
          pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination', totality=NONE,
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   124
          fs=fs, R=R, dom=dom, defname=defname, is_partial=true, cases=flat cases',
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   125
          pelims=pelims',elims=NONE}
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   126
56932
11a4001b06c6 more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents: 56254
diff changeset
   127
        val _ =
11a4001b06c6 more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents: 56254
diff changeset
   128
          Proof_Display.print_consts do_print (Position.thread_data ()) lthy
11a4001b06c6 more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents: 56254
diff changeset
   129
            (K false) (map fst fixes)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   130
      in
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
   131
        (info,
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 45290
diff changeset
   132
         lthy |> Local_Theory.declaration {syntax = false, pervasive = false}
63004
f507e6fe1d77 prefer binding over base name;
wenzelm
parents: 62996
diff changeset
   133
          (fn phi => add_function_data (transform_function_data phi info)))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   134
      end
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   135
  in
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   136
    ((goal_state, afterqed), lthy')
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   137
  end
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   138
62774
cfcb20bbdbd8 reconcile object-logic constraint vs. mixfix constraint;
wenzelm
parents: 61841
diff changeset
   139
fun gen_add_function do_print prep fixspec eqns config tac lthy =
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   140
  let
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   141
    val ((goal_state, afterqed), lthy') =
62774
cfcb20bbdbd8 reconcile object-logic constraint vs. mixfix constraint;
wenzelm
parents: 61841
diff changeset
   142
      prepare_function do_print prep fixspec eqns config lthy
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   143
    val pattern_thm =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   144
      case SINGLE (tac lthy') goal_state of
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   145
        NONE => error "pattern completeness and compatibility proof failed"
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   146
      | SOME st => Goal.finish lthy' st
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   147
  in
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   148
    lthy'
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   149
    |> afterqed [[pattern_thm]]
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   150
  end
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   151
63064
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
   152
val add_function = gen_add_function false Specification.check_multi_specs
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
   153
fun add_function_cmd a b c d int = gen_add_function int Specification.read_multi_specs a b c d
60499
54a3db2ed201 avoid dynamic parsing of hardwired strings;
wenzelm
parents: 59936
diff changeset
   154
62774
cfcb20bbdbd8 reconcile object-logic constraint vs. mixfix constraint;
wenzelm
parents: 61841
diff changeset
   155
fun gen_function do_print prep fixspec eqns config lthy =
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   156
  let
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   157
    val ((goal_state, afterqed), lthy') =
62774
cfcb20bbdbd8 reconcile object-logic constraint vs. mixfix constraint;
wenzelm
parents: 61841
diff changeset
   158
      prepare_function do_print prep fixspec eqns config lthy
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   159
  in
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   160
    lthy'
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59159
diff changeset
   161
    |> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (Thm.concl_of goal_state), [])]]
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61112
diff changeset
   162
    |> Proof.refine_singleton (Method.primitive_text (K (K goal_state)))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   163
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   164
63064
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
   165
val function = gen_function false Specification.check_multi_specs
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
   166
fun function_cmd a b c int = gen_function int Specification.read_multi_specs a b c
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   167
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   168
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
   169
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   170
    val term_opt = Option.map (prep_term lthy) raw_term_opt
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   171
    val info =
49967
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   172
      (case term_opt of
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   173
        SOME t =>
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   174
          (case import_function_data t lthy of
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   175
            SOME info => info
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   176
          | NONE => error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   177
      | NONE =>
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   178
          (case import_last_function lthy of
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   179
            SOME info => info
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   180
          | NONE => error "Not a function"))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   181
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   182
    val { termination, fs, R, add_simps, case_names, psimps,
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   183
      pinducts, defname, fnames, cases, dom, pelims, ...} = info
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   184
    val domT = domain_type (fastype_of R)
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   185
    val goal = HOLogic.mk_Trueprop (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   186
    fun afterqed [[totality]] lthy =
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   187
      let
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   188
        val totality = Thm.close_derivation totality
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   189
        val remove_domain_condition =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   190
          full_simplify (put_simpset HOL_basic_ss lthy
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   191
            addsimps [totality, @{thm True_implies_equals}])
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   192
        val tsimps = map remove_domain_condition psimps
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   193
        val tinduct = map remove_domain_condition pinducts
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   194
        val telims = map (map remove_domain_condition) pelims
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   195
      in
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   196
        lthy
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   197
        |> add_simps I "simps" I simp_attribs tsimps
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   198
        ||>> Local_Theory.note
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63011
diff changeset
   199
          ((derived_name defname "induct", [Attrib.case_names case_names]), tinduct)
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63011
diff changeset
   200
        ||>> fold_map (note_derived ("elims", [Attrib.consumes 1, Attrib.constraints 1]))
63004
f507e6fe1d77 prefer binding over base name;
wenzelm
parents: 62996
diff changeset
   201
          (fnames ~~ telims)
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   202
        |-> (fn ((simps,(_,inducts)), elims) => fn lthy =>
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   203
          let val info' = { is_partial=false, defname=defname, fnames=fnames, add_simps=add_simps,
52384
80c00a851de5 export dom predicate in the info record
krauss
parents: 52383
diff changeset
   204
            case_names=case_names, fs=fs, R=R, dom=dom, psimps=psimps, pinducts=pinducts,
65387
5dbe02addca5 store totality fact in function info
Lars Hupel <lars.hupel@mytum.de>
parents: 63239
diff changeset
   205
            simps=SOME simps, inducts=SOME inducts, termination=termination, totality=SOME totality,
5dbe02addca5 store totality fact in function info
Lars Hupel <lars.hupel@mytum.de>
parents: 63239
diff changeset
   206
            cases=cases, pelims=pelims, elims=SOME elims}
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   207
          in
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   208
            (info',
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   209
             lthy
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   210
             |> Local_Theory.declaration {syntax = false, pervasive = false}
63004
f507e6fe1d77 prefer binding over base name;
wenzelm
parents: 62996
diff changeset
   211
               (fn phi => add_function_data (transform_function_data phi info'))
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   212
             |> Spec_Rules.add Spec_Rules.Equational (fs, tsimps))
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   213
          end)
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   214
      end
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   215
  in
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   216
    (goal, afterqed, termination)
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   217
  end
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   218
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   219
fun gen_prove_termination prep_term raw_term_opt tac lthy =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   220
  let
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   221
    val (goal, afterqed, termination) =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   222
      prepare_termination_proof prep_term raw_term_opt lthy
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   223
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   224
    val totality = Goal.prove lthy [] [] goal (K tac)
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   225
  in
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   226
    afterqed [[totality]] lthy
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   227
end
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   228
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   229
val prove_termination = gen_prove_termination Syntax.check_term
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   230
val prove_termination_cmd = gen_prove_termination Syntax.read_term
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   231
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   232
fun gen_termination prep_term raw_term_opt lthy =
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   233
  let
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   234
    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
   235
  in
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   236
    lthy
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41846
diff changeset
   237
    |> Proof_Context.note_thmss ""
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   238
       [((Binding.empty, [Context_Rules.rule_del]), [([allI], [])])] |> snd
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41846
diff changeset
   239
    |> Proof_Context.note_thmss ""
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   240
       [((Binding.empty, [Context_Rules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41846
diff changeset
   241
    |> Proof_Context.note_thmss ""
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   242
       [((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]),
54883
dd04a8b654fc proper context for norm_hhf and derived operations;
wenzelm
parents: 53614
diff changeset
   243
         [([Goal.norm_result lthy termination], [])])] |> snd
36547
2a9d0ec8c10d return updated info record after termination proof
krauss
parents: 36522
diff changeset
   244
    |> Proof.theorem NONE (snd oo afterqed) [[(goal, [])]]
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   245
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   246
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   247
val termination = gen_termination Syntax.check_term
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   248
val termination_cmd = gen_termination Syntax.read_term
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   249
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   250
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   251
(* Datatype hook to declare datatype congs as "function_congs" *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   252
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   253
fun add_case_cong n thy =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   254
  let
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 57959
diff changeset
   255
    val cong = #case_cong (Old_Datatype_Data.the_info thy n)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   256
      |> safe_mk_meta_eq
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   257
  in
61112
e966c311e9a7 trim context for persistent storage;
wenzelm
parents: 60643
diff changeset
   258
    Context.theory_map (Function_Context_Tree.add_function_cong cong) thy
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   259
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   260
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58816
diff changeset
   261
val _ = Theory.setup (Old_Datatype_Data.interpretation (K (fold add_case_cong)))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   262
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   263
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58816
diff changeset
   264
(* get info *)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   265
58816
aab139c0003f modernized setup;
wenzelm
parents: 58112
diff changeset
   266
val get_congs = Function_Context_Tree.get_function_congs
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   267
59159
9312710451f5 just one data slot per program unit;
wenzelm
parents: 58826
diff changeset
   268
fun get_info ctxt t = Item_Net.retrieve (get_functions ctxt) t
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   269
  |> the_single |> snd
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   270
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36547
diff changeset
   271
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   272
(* outer syntax *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   273
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   274
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59859
diff changeset
   275
  Outer_Syntax.local_theory_to_proof' @{command_keyword function}
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45592
diff changeset
   276
    "define general recursive functions"
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45592
diff changeset
   277
    (function_parser default_config
63183
wenzelm
parents: 63064
diff changeset
   278
      >> (fn (config, (fixes, specs)) => function_cmd fixes specs config))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   279
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   280
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59859
diff changeset
   281
  Outer_Syntax.local_theory_to_proof @{command_keyword termination}
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45592
diff changeset
   282
    "prove termination of a recursive function"
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45592
diff changeset
   283
    (Scan.option Parse.term >> termination_cmd)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   284
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   285
end