src/HOL/Tools/Function/function.ML
author Manuel Eberl
Sun, 08 Sep 2013 22:32:47 +0200
changeset 53603 59ef06cda7b9
parent 52384 80c00a851de5
child 53604 c1db98d7c66f
permissions -rw-r--r--
generate elim rules for elimination of function equalities; added fun_cases command; recover proper cases rules for mutual recursive case (no sum types)
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 ->
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    12
    (Attrib.binding * term) list -> 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 ->
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
    16
    (Attrib.binding * string) list -> 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 ->
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    20
    (Attrib.binding * term) list -> Function_Common.function_config ->
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 ->
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    24
    (Attrib.binding * string) list -> 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 setup : theory -> theory
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    36
  val get_congs : Proof.context -> thm list
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
    37
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    38
  val get_info : Proof.context -> term -> info
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    39
end
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
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    42
structure Function : FUNCTION =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    43
struct
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
open Function_Lib
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    46
open Function_Common
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    47
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45291
diff changeset
    48
val simp_attribs =
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45291
diff changeset
    49
  @{attributes [simp, nitpick_simp]} @ [Attrib.internal (K Code.add_default_eqn_attribute)]
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    50
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45291
diff changeset
    51
val psimp_attribs =
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45291
diff changeset
    52
  @{attributes [nitpick_psimp]}
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    53
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    54
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
    55
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    56
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
    57
  simps lthy =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    58
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    59
    val spec = post simps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    60
      |> map (apfst (apsnd (fn ats => moreatts @ ats)))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    61
      |> map (apfst (apfst extra_qualify))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    62
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    63
    val (saved_spec_simps, lthy) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    64
      fold_map Local_Theory.note spec lthy
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    65
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    66
    val saved_simps = maps snd saved_spec_simps
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    67
    val simps_by_f = sort saved_simps
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    68
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    69
    fun add_for_f fname simps =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    70
      Local_Theory.note
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    71
        ((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
    72
      #> snd
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    73
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    74
    (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
    75
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    76
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
    77
fun prepare_function do_print prep default_constraint fixspec eqns config lthy =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    78
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    79
    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
    80
    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
    81
    val fixes = map (apfst (apfst Binding.name_of)) fixes0;
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    82
    val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    83
    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
    84
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    85
    val defname = mk_defname fixes
41846
b368a7aee46a removed support for tail-recursion from function package (now implemented by partial_function)
krauss
parents: 41417
diff changeset
    86
    val FunctionConfig {partials, default, ...} = config
41417
211dbd42f95d function (default) is legacy feature
krauss
parents: 41405
diff changeset
    87
    val _ =
44052
00f0c8782a51 slightly more uniform messages;
wenzelm
parents: 43277
diff changeset
    88
      if is_some default
00f0c8782a51 slightly more uniform messages;
wenzelm
parents: 43277
diff changeset
    89
      then legacy_feature "\"function (default)\" -- use 'partial_function' instead"
41417
211dbd42f95d function (default) is legacy feature
krauss
parents: 41405
diff changeset
    90
      else ()
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    91
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
    92
    val ((goal_state, cont), lthy') =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    93
      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
    94
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    95
    fun afterqed [[proof]] lthy =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
    96
      let
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
    97
        val result = cont (Thm.close_derivation proof)
52384
80c00a851de5 export dom predicate in the info record
krauss
parents: 52383
diff changeset
    98
        val FunctionResult {fs, R, dom, psimps, simple_pinducts,
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
    99
                termination, domintros, cases, ...} = result
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   100
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   101
        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
   102
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   103
        val fnames = map (fst o fst) fixes
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   104
        fun qualify n = Binding.name n
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   105
          |> Binding.qualify true defname
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   106
        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
   107
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   108
        val addsmps = add_simps fnames post sort_cont
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   109
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   110
        (* TODO: case names *)
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   111
        fun addcases lthy =
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   112
          let fun go name thm (thms_acc, lthy) =
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   113
                 case Local_Theory.note ((Binding.name "cases" |> Binding.qualify true name,
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   114
                          [Attrib.internal (K (Rule_Cases.case_names cnames))]), [thm]) lthy
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   115
                 of ((_,thms), lthy') => (thms :: thms_acc, lthy')
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   116
              val (thms, lthy') = fold2 go fnames cases ([], lthy);
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   117
          in
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   118
            (rev thms, lthy')
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   119
          end;
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   120
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   121
        fun addpelims lthy =
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   122
          let fun go name thm (thms_acc, lthy) =
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   123
                 case Local_Theory.note ((Binding.name "pelims" |> Binding.qualify true name,
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   124
                          [Attrib.internal (K (Rule_Cases.consumes 1)),
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   125
                           Attrib.internal (K (Rule_Cases.constraints 1))]), thm) lthy
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   126
                 of ((_,thms), lthy') => (thms :: thms_acc, lthy')
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   127
              val (thms, lthy') = fold2 go fnames pelims ([], lthy);
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   128
          in
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   129
            (rev thms, lthy')
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   130
          end;
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   131
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   132
          val (((((psimps', [pinducts']), [termination']), cases'), pelims'), lthy) =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   133
          lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   134
          |> addsmps (conceal_partial o Binding.qualify false "partial")
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   135
               "psimps" conceal_partial psimp_attribs psimps
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
   136
          ||>> Local_Theory.notes [((conceal_partial (qualify "pinduct"), []),
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
   137
                simple_pinducts |> map (fn th => ([th],
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   138
                 [Attrib.internal (K (Rule_Cases.case_names cnames)),
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
   139
                  Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of th))),
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
   140
                  Attrib.internal (K (Induct.induct_pred ""))])))]
52383
71df93ff010d export cases rule in the info record
krauss
parents: 51717
diff changeset
   141
          ||>> (apfst snd o Local_Theory.note ((Binding.conceal (qualify "termination"), []), [termination]))
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   142
          ||>> addcases
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   143
          ||>> addpelims
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   144
          ||> (case domintros of NONE => I | SOME thms =>
40076
6f012a209dac some cleanup in Function_Lib
krauss
parents: 39754
diff changeset
   145
                   Local_Theory.note ((qualify "domintros", []), thms) #> snd)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   146
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   147
        val info = { add_simps=addsmps, fnames=fnames, case_names=cnames, psimps=psimps',
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   148
          pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination',
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   149
          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
   150
          pelims=pelims',elims=NONE}
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   151
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
   152
        val _ = Proof_Display.print_consts do_print lthy (K false) (map fst fixes)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   153
      in
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
   154
        (info,
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 45290
diff changeset
   155
         lthy |> Local_Theory.declaration {syntax = false, pervasive = false}
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 45290
diff changeset
   156
          (add_function_data o transform_function_data info))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   157
      end
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   158
  in
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   159
    ((goal_state, afterqed), lthy')
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   160
  end
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   161
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
   162
fun gen_add_function do_print prep default_constraint fixspec eqns config tac lthy =
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   163
  let
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   164
    val ((goal_state, afterqed), lthy') =
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
   165
      prepare_function do_print prep default_constraint fixspec eqns config lthy
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   166
    val pattern_thm =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   167
      case SINGLE (tac lthy') goal_state of
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   168
        NONE => error "pattern completeness and compatibility proof failed"
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   169
      | SOME st => Goal.finish lthy' st
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   170
  in
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   171
    lthy'
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   172
    |> afterqed [[pattern_thm]]
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   173
  end
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   174
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   175
val add_function =
37145
01aa36932739 renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
wenzelm
parents: 36960
diff changeset
   176
  gen_add_function false Specification.check_spec (Type_Infer.anyT HOLogic.typeS)
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
   177
fun add_function_cmd a b c d int = gen_add_function int Specification.read_spec "_::type" a b c d
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   178
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
   179
fun gen_function do_print prep default_constraint fixspec eqns config lthy =
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   180
  let
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   181
    val ((goal_state, afterqed), lthy') =
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
   182
      prepare_function do_print prep default_constraint fixspec eqns config lthy
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   183
  in
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   184
    lthy'
36522
e80a95279ef6 return info record (relative to auxiliary context!)
krauss
parents: 36520
diff changeset
   185
    |> 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
   186
    |> 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
   187
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   188
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   189
val function =
37145
01aa36932739 renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
wenzelm
parents: 36960
diff changeset
   190
  gen_function false Specification.check_spec (Type_Infer.anyT HOLogic.typeS)
44239
47ecd30e018d less verbosity for 'function' and 'fun': observe "int" flag more carefully (cf. a32ca9165928);
wenzelm
parents: 44192
diff changeset
   191
fun function_cmd a b c int = gen_function int Specification.read_spec "_::type" a b c
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   192
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   193
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
   194
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   195
    val term_opt = Option.map (prep_term lthy) raw_term_opt
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   196
    val info =
49967
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   197
      (case term_opt of
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   198
        SOME t =>
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   199
          (case import_function_data t lthy of
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   200
            SOME info => info
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   201
          | 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
   202
      | NONE =>
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   203
          (case import_last_function lthy of
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   204
            SOME info => info
69774b4f5b8a recovered explicit error message, which was lost in b8570ea1ce25;
wenzelm
parents: 48995
diff changeset
   205
          | NONE => error "Not a function"))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   206
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   207
    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
   208
      pinducts, defname, fnames, cases, dom, pelims, ...} = info
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   209
    val domT = domain_type (fastype_of R)
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   210
    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
   211
    fun afterqed [[totality]] lthy =
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   212
      let
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   213
        val totality = Thm.close_derivation totality
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   214
        val remove_domain_condition =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   215
          full_simplify (put_simpset HOL_basic_ss lthy
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   216
            addsimps [totality, @{thm True_implies_equals}])
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   217
        val tsimps = map remove_domain_condition psimps
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   218
        val tinduct = map remove_domain_condition pinducts
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   219
        val telims = map (map remove_domain_condition) pelims
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   220
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   221
        fun qualify n = Binding.name n
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   222
          |> Binding.qualify true defname
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   223
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   224
        fun addtelims lthy =
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   225
          let fun go name thm (thms_acc, lthy) =
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   226
                 case Local_Theory.note ((Binding.name "elims" |> Binding.qualify true name,
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   227
                          [Attrib.internal (K (Rule_Cases.consumes 1)),
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   228
                           Attrib.internal (K (Rule_Cases.constraints 1)),
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   229
                           Attrib.internal (K (Induct.cases_pred defname))]), thm) lthy
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   230
                 of ((_,thms), lthy') => (thms :: thms_acc, lthy')
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   231
              val (thms, lthy') = fold2 go fnames telims ([], lthy);
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   232
          in
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   233
            (rev thms, lthy')
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   234
          end;
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   235
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   236
      in
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   237
        lthy
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   238
        |> add_simps I "simps" I simp_attribs tsimps
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   239
        ||>> Local_Theory.note
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   240
           ((qualify "induct",
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   241
             [Attrib.internal (K (Rule_Cases.case_names case_names))]),
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   242
            tinduct)
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   243
        ||>> addtelims
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   244
        |-> (fn ((simps,(_,inducts)), elims) => fn lthy =>
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   245
          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
   246
            case_names=case_names, fs=fs, R=R, dom=dom, psimps=psimps, pinducts=pinducts,
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   247
            simps=SOME simps, inducts=SOME inducts, termination=termination, cases=cases, pelims=pelims, elims=SOME elims}
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   248
          in
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   249
            (info',
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52384
diff changeset
   250
             lthy
48995
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   251
             |> Local_Theory.declaration {syntax = false, pervasive = false}
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   252
               (add_function_data o transform_function_data info')
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   253
             |> Spec_Rules.add Spec_Rules.Equational (fs, tsimps))
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   254
          end)
0e1cab4a334e more precise indentation;
wenzelm
parents: 47701
diff changeset
   255
      end
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   256
  in
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   257
    (goal, afterqed, termination)
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   258
  end
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   259
36520
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   260
fun gen_prove_termination prep_term raw_term_opt tac lthy =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   261
  let
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   262
    val (goal, afterqed, termination) =
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   263
      prepare_termination_proof prep_term raw_term_opt lthy
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   264
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   265
    val totality = Goal.prove lthy [] [] goal (K tac)
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   266
  in
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   267
    afterqed [[totality]] lthy
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   268
end
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   269
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   270
val prove_termination = gen_prove_termination Syntax.check_term
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   271
val prove_termination_cmd = gen_prove_termination Syntax.read_term
772ed73e1d61 function: sane interface for programmatic use
krauss
parents: 36519
diff changeset
   272
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   273
fun gen_termination prep_term raw_term_opt lthy =
36518
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   274
  let
a33b986f2e22 function: better separate Isar integration from actual functionality
krauss
parents: 36323
diff changeset
   275
    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
   276
  in
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   277
    lthy
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41846
diff changeset
   278
    |> Proof_Context.note_thmss ""
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   279
       [((Binding.empty, [Context_Rules.rule_del]), [([allI], [])])] |> snd
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41846
diff changeset
   280
    |> Proof_Context.note_thmss ""
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   281
       [((Binding.empty, [Context_Rules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41846
diff changeset
   282
    |> Proof_Context.note_thmss ""
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   283
       [((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   284
         [([Goal.norm_result termination], [])])] |> snd
36547
2a9d0ec8c10d return updated info record after termination proof
krauss
parents: 36522
diff changeset
   285
    |> Proof.theorem NONE (snd oo afterqed) [[(goal, [])]]
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   286
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   287
36519
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   288
val termination = gen_termination Syntax.check_term
46bf776a81e0 ML interface uses plain command names, following conventions from typedef
krauss
parents: 36518
diff changeset
   289
val termination_cmd = gen_termination Syntax.read_term
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   290
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   291
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   292
(* Datatype hook to declare datatype congs as "function_congs" *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   293
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   294
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   295
fun add_case_cong n thy =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   296
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   297
    val cong = #case_cong (Datatype.the_info thy n)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   298
      |> safe_mk_meta_eq
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   299
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   300
    Context.theory_map
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34231
diff changeset
   301
      (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
   302
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   303
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   304
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
   305
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   306
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   307
(* setup *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   308
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   309
val setup =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   310
  Attrib.setup @{binding fundef_cong}
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   311
    (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
   312
    "declaration of congruence rule for function definitions"
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   313
  #> setup_case_cong
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   314
  #> Function_Common.Termination_Simps.setup
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   315
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   316
val get_congs = Function_Ctx_Tree.get_function_congs
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   317
34230
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   318
fun get_info ctxt t = Item_Net.retrieve (get_function ctxt) t
b0d21ae2528e more official data record Function.info
krauss
parents: 33726
diff changeset
   319
  |> the_single |> snd
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   320
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36547
diff changeset
   321
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   322
(* outer syntax *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   323
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   324
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45592
diff changeset
   325
  Outer_Syntax.local_theory_to_proof' @{command_spec "function"}
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45592
diff changeset
   326
    "define general recursive functions"
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45592
diff changeset
   327
    (function_parser default_config
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45592
diff changeset
   328
      >> (fn ((config, fixes), statements) => function_cmd fixes statements config))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   329
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   330
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45592
diff changeset
   331
  Outer_Syntax.local_theory_to_proof @{command_spec "termination"}
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45592
diff changeset
   332
    "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
   333
    (Scan.option Parse.term >> termination_cmd)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   334
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   335
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   336
end