src/HOL/Tools/function_package/fundef_package.ML
author krauss
Fri, 22 Jun 2007 10:23:37 +0200
changeset 23473 997bca36d4fe
parent 23203 a5026e73cfcf
child 23819 2040846d1bbe
permissions -rw-r--r--
new method "elim_to_cases" provides ad-hoc conversion of obtain-style elimination goals to a disjunction of existentials.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Tools/function_package/fundef_package.ML
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     2
    ID:         $Id$
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     3
    Author:     Alexander Krauss, TU Muenchen
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     4
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
     5
A package for general recursive function definitions.
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     6
Isar commands.
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     7
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     8
*)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     9
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
    10
signature FUNDEF_PACKAGE =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    11
sig
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    12
    val add_fundef :  (string * string option * mixfix) list
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    13
                      -> ((bstring * Attrib.src list) * string) list 
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    14
                      -> FundefCommon.fundef_config
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    15
                      -> bool list
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    16
                      -> local_theory
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
    17
                      -> Proof.state
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    18
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    19
    val add_fundef_i:  (string * typ option * mixfix) list
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    20
                       -> ((bstring * Attrib.src list) * term) list
22170
5682eeaefaf4 fixes smlnj-related problem, updated signature
krauss
parents: 22166
diff changeset
    21
                       -> FundefCommon.fundef_config
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    22
                       -> bool list
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    23
                       -> local_theory
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
    24
                       -> Proof.state
21211
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
    25
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
    26
    val setup_termination_proof : string option -> local_theory -> Proof.state
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    27
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    28
    val cong_add: attribute
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    29
    val cong_del: attribute
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
    30
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    31
    val setup : theory -> theory
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
    32
    val setup_case_cong_hook : theory -> theory
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
    33
    val get_congs : theory -> thm list
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    34
end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    35
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    36
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    37
structure FundefPackage : FUNDEF_PACKAGE =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    38
struct
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    39
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 21000
diff changeset
    40
open FundefLib
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    41
open FundefCommon
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    42
21435
883337ea2f3b LocalTheory.axioms/notes/defs: proper kind;
wenzelm
parents: 21391
diff changeset
    43
val note_theorem = LocalTheory.note Thm.theoremK
883337ea2f3b LocalTheory.axioms/notes/defs: proper kind;
wenzelm
parents: 21391
diff changeset
    44
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    45
fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_" 
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    46
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    47
fun add_simps fixes post sort label moreatts simps lthy =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    48
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    49
      val fnames = map (fst o fst) fixes
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
    50
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    51
      val (saved_spec_simps, lthy) =
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    52
        fold_map note_theorem (post simps) lthy
22498
62cdd4b3e96b made function syntax strict, requiring | to separate equations; cleanup
krauss
parents: 22325
diff changeset
    53
      val saved_simps = flat (map snd saved_spec_simps)
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
    54
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    55
      val simps_by_f = sort saved_simps
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    56
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    57
      fun add_for_f fname simps =
21435
883337ea2f3b LocalTheory.axioms/notes/defs: proper kind;
wenzelm
parents: 21391
diff changeset
    58
        note_theorem
21658
5e31241e1e3c Attrib.internal: morphism;
wenzelm
parents: 21602
diff changeset
    59
          ((NameSpace.qualified fname label, Attrib.internal (K Simplifier.simp_add) :: moreatts),
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    60
            simps) #> snd
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    61
    in
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    62
      (saved_simps,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    63
       fold2 add_for_f fnames simps_by_f lthy)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    64
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    65
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    66
fun fundef_afterqed config fixes post defname cont sort_cont [[proof]] lthy =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    67
    let
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
    68
      val FundefResult {fs, R, psimps, trsimps, subset_pinducts, simple_pinducts, termination, domintros, cases, ...} = 
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    69
          cont (Goal.close_result proof)
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    70
21391
a8809f46bd7f replaced NameSpace.append by NameSpace.qualified, which handles empty names as expected;
wenzelm
parents: 21359
diff changeset
    71
      val qualify = NameSpace.qualified defname
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    72
      val addsmps = add_simps fixes post sort_cont
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
    73
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
    74
      val (((psimps', pinducts'), (_, [termination'])), lthy) =
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
    75
          lthy
22170
5682eeaefaf4 fixes smlnj-related problem, updated signature
krauss
parents: 22166
diff changeset
    76
            |> addsmps "psimps" [] psimps
5682eeaefaf4 fixes smlnj-related problem, updated signature
krauss
parents: 22166
diff changeset
    77
            ||> fold_option (snd oo addsmps "simps" []) trsimps
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    78
            ||>> note_theorem ((qualify "pinduct",
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    79
                                [Attrib.internal (K (InductAttrib.induct_set ""))]), simple_pinducts)
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    80
            ||>> note_theorem ((qualify "termination", []), [termination])
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    81
            ||> (snd o note_theorem ((qualify "cases", []), [cases]))
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    82
            ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
    83
22170
5682eeaefaf4 fixes smlnj-related problem, updated signature
krauss
parents: 22166
diff changeset
    84
      val cdata = FundefCtxData { add_simps=addsmps, psimps=psimps',
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
    85
                                  pinducts=snd pinducts', termination=termination', fs=fs, R=R, defname=defname }
22668
8183ee579232 data declaration: removed obsolete target_morphism (still required for local data!?);
wenzelm
parents: 22635
diff changeset
    86
      val cdata' = cdata |> morph_fundef_data (LocalTheory.target_morphism lthy);  (* FIXME !? *)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    87
    in
22623
5fcee5b319a2 proper handling of morphisms
krauss
parents: 22498
diff changeset
    88
      lthy 
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
    89
        |> LocalTheory.declaration (fn phi => add_fundef_data (morph_fundef_data phi cdata)) (* save in target *)
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
    90
        |> Context.proof_map (add_fundef_data cdata') (* also save in local context *)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    91
    end (* FIXME: Add cases for induct and cases thm *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    92
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    93
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    94
fun prepare_spec prep fixspec eqnss lthy = (* FIXME: obsolete when "read_specification" etc. is there *)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    95
    let
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    96
      val eqns = map (apsnd single) eqnss
22498
62cdd4b3e96b made function syntax strict, requiring | to separate equations; cleanup
krauss
parents: 22325
diff changeset
    97
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    98
      val ((fixes, _), ctxt') = prep fixspec [] lthy             
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
    99
      fun prep_eqn e = the_single (snd (fst (prep [] [e] ctxt')))
23189
4574ab8f3b21 simplified interfaces, some restructuring
krauss
parents: 22846
diff changeset
   100
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
   101
      val spec = map prep_eqn eqns
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
   102
                 |> map (apsnd (map (fn t => fold_rev (mk_forall o Free) (frees_in_term ctxt' t) t))) (* Add quantifiers *)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   103
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   104
      ((fixes, spec), ctxt')
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   105
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   106
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
   107
fun gen_add_fundef prep fixspec eqnss config flags lthy =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   108
    let
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
   109
      val ((fixes, spec), ctxt') = prepare_spec prep fixspec eqnss lthy
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
   110
      val (eqs, post) = FundefCommon.get_preproc lthy config flags ctxt' fixes spec
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   111
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   112
      val defname = mk_defname fixes
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   113
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
   114
      val ((goalstate, cont, sort_cont), lthy) =
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
   115
          FundefMutual.prepare_fundef_mutual config defname fixes eqs lthy
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   116
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
   117
      val afterqed = fundef_afterqed config fixes post defname cont sort_cont
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   118
    in
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   119
      lthy
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   120
        |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   121
        |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   122
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   123
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   124
fun total_termination_afterqed data [[totality]] lthy =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   125
    let
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   126
      val FundefCtxData { add_simps, psimps, pinducts, defname, ... } = data
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
   127
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   128
      val totality = Goal.close_result totality
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   129
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
   130
      val remove_domain_condition = full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   131
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   132
      val tsimps = map remove_domain_condition psimps
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   133
      val tinduct = map remove_domain_condition pinducts
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
   134
21511
16c62deb1adf The function package declares the [code] attribute automatically again.
krauss
parents: 21491
diff changeset
   135
      val has_guards = exists ((fn (Const ("Trueprop", _) $ _) => false | _ => true) o prop_of) tsimps
21658
5e31241e1e3c Attrib.internal: morphism;
wenzelm
parents: 21602
diff changeset
   136
      val allatts = if has_guards then [] else [Attrib.internal (K (RecfunCodegen.add NONE))]
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   137
21511
16c62deb1adf The function package declares the [code] attribute automatically again.
krauss
parents: 21491
diff changeset
   138
      val qualify = NameSpace.qualified defname;
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   139
    in
21511
16c62deb1adf The function package declares the [code] attribute automatically again.
krauss
parents: 21491
diff changeset
   140
      lthy
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   141
        |> add_simps "simps" allatts tsimps |> snd
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   142
        |> note_theorem ((qualify "induct", []), tinduct) |> snd
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
   143
    end
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   144
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   145
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   146
fun setup_termination_proof term_opt lthy =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   147
    let
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   148
      val data = the (case term_opt of
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   149
                        SOME t => import_fundef_data (ProofContext.read_term lthy t) (Context.Proof lthy)
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   150
                      | NONE => import_last_fundef (Context.Proof lthy))
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   151
          handle Option.Option => raise ERROR ("Not a function: " ^ quote (the_default "" term_opt))
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   152
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   153
        val FundefCtxData {termination, R, ...} = data
22325
be61bd159a99 changed termination goal to use object quantifier
krauss
parents: 22170
diff changeset
   154
        val domT = domain_type (fastype_of R)
be61bd159a99 changed termination goal to use object quantifier
krauss
parents: 22170
diff changeset
   155
        val goal = HOLogic.mk_Trueprop (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   156
    in
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
   157
      lthy
22325
be61bd159a99 changed termination goal to use object quantifier
krauss
parents: 22170
diff changeset
   158
        |> ProofContext.note_thmss_i "" [(("", [ContextRules.rule_del]), [([allI], [])])] |> snd
be61bd159a99 changed termination goal to use object quantifier
krauss
parents: 22170
diff changeset
   159
        |> ProofContext.note_thmss_i "" [(("", [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   160
        |> ProofContext.note_thmss_i ""
22325
be61bd159a99 changed termination goal to use object quantifier
krauss
parents: 22170
diff changeset
   161
          [(("termination", [ContextRules.intro_bang (SOME 0)]),
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   162
            [([Goal.norm_result termination], [])])] |> snd
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   163
        |> Proof.theorem_i NONE (total_termination_afterqed data) [[(goal, [])]]
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   164
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   165
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   166
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   167
val add_fundef = gen_add_fundef Specification.read_specification
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   168
val add_fundef_i = gen_add_fundef Specification.cert_specification
19611
da2a0014c461 Function Package: Quick-and-dirty-fixed strange "Proved a different theorem bug"
krauss
parents: 19585
diff changeset
   169
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   170
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   171
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   172
(* congruence rules *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   173
19617
7cb4b67d4b97 avoid raw equality on type thm;
wenzelm
parents: 19611
diff changeset
   174
val cong_add = Thm.declaration_attribute (map_fundef_congs o Drule.add_rule o safe_mk_meta_eq);
7cb4b67d4b97 avoid raw equality on type thm;
wenzelm
parents: 19611
diff changeset
   175
val cong_del = Thm.declaration_attribute (map_fundef_congs o Drule.del_rule o safe_mk_meta_eq);
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   176
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   177
(* Datatype hook to declare datatype congs as "fundef_congs" *)
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   178
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   179
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   180
fun add_case_cong n thy =
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   181
    Context.theory_map (map_fundef_congs (Drule.add_rule
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   182
                          (DatatypePackage.get_datatype thy n |> the
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   183
                           |> #case_cong
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   184
                           |> safe_mk_meta_eq)))
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   185
                       thy
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   186
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   187
val case_cong_hook = fold add_case_cong
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   188
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   189
val setup_case_cong_hook =
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   190
    DatatypeHooks.add case_cong_hook
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   191
    #> (fn thy => case_cong_hook (Symtab.keys (DatatypePackage.get_datatypes thy)) thy)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   192
23473
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   193
(* ad-hoc method to convert elimination-style goals to existential statements *)
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   194
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   195
fun insert_int_goal thy subg st =
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   196
    let
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   197
      val goal = hd (prems_of st)
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   198
      val (ps, imp) = dest_all_all goal
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   199
      val cps = map (cterm_of thy) ps
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   200
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   201
      val imp_subg = fold (fn p => fn t => betapply (t,p)) ps subg
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   202
      val new_subg = implies $ imp_subg $ imp
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   203
                      |> fold_rev mk_forall ps
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   204
                      |> cterm_of thy 
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   205
                      |> assume 
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   206
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   207
      val sg2 = imp_subg
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   208
                 |> fold_rev mk_forall ps
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   209
                 |> cterm_of thy 
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   210
                 |> assume
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   211
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   212
      val t' = new_subg
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   213
                |> fold forall_elim cps
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   214
                |> flip implies_elim (fold forall_elim cps sg2)
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   215
                |> fold_rev forall_intr cps
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   216
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   217
      val st' = implies_elim st t'
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   218
                 |> implies_intr (cprop_of sg2)
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   219
                 |> implies_intr (cprop_of new_subg)
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   220
    in
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   221
      Seq.single st'
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   222
    end
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   223
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   224
fun mk_cases_statement thy t =
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   225
    let
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   226
      fun mk_clause t = 
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   227
          let 
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   228
            val (qs, imp) = dest_all_all t
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   229
          in 
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   230
            Logic.strip_imp_prems imp
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   231
             |> map (ObjectLogic.atomize_term thy)
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   232
             |> foldr1 HOLogic.mk_conj
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   233
             |> fold_rev (fn Free (v,T) => fn t => HOLogic.mk_exists (v,T,t)) qs
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   234
          end
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   235
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   236
      val (ps, imp) = dest_all_all t
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   237
    in 
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   238
      Logic.strip_imp_prems imp
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   239
       |> map mk_clause
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   240
       |> foldr1 HOLogic.mk_disj
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   241
       |> HOLogic.mk_Trueprop
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   242
       |> fold_rev lambda ps
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   243
    end
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   244
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   245
fun elim_to_cases1 ctxt st =
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   246
    let
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   247
      val thy = theory_of_thm st
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   248
      val [subg] = prems_of st
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   249
      val cex = mk_cases_statement thy subg
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   250
    in
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   251
      (insert_int_goal thy cex
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   252
       THEN REPEAT_ALL_NEW (Tactic.ematch_tac [disjE, exE, conjE]) 1
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   253
       THEN REPEAT (Goal.assume_rule_tac ctxt 1)
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   254
    (*   THEN REPEAT (etac thin_rl 1)*)) st
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   255
    end
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   256
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   257
fun elim_to_cases_tac ctxt = SELECT_GOAL (elim_to_cases1 ctxt)
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   258
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   259
val elim_to_cases_setup = Method.add_methods
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   260
  [("elim_to_cases", Method.ctxt_args (Method.SIMPLE_METHOD' o elim_to_cases_tac),
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   261
    "convert elimination-style goal to a disjunction of existentials")]
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   262
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   263
(* setup *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   264
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   265
val setup =
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22733
diff changeset
   266
  Attrib.add_attributes
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22733
diff changeset
   267
    [("fundef_cong", Attrib.add_del_args cong_add cong_del,
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22733
diff changeset
   268
      "declaration of congruence rule for function definitions")]
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22733
diff changeset
   269
  #> setup_case_cong_hook
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22733
diff changeset
   270
  #> FundefRelation.setup
23473
997bca36d4fe new method "elim_to_cases" provides ad-hoc conversion of obtain-style
krauss
parents: 23203
diff changeset
   271
  #> elim_to_cases_setup
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   272
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
   273
val get_congs = FundefCommon.get_fundef_congs o Context.Theory
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
   274
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
   275
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   276
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   277
(* outer syntax *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   278
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   279
local structure P = OuterParse and K = OuterKeyword in
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   280
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   281
val functionP =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   282
  OuterSyntax.command "function" "define general recursive functions" K.thy_goal
22498
62cdd4b3e96b made function syntax strict, requiring | to separate equations; cleanup
krauss
parents: 22325
diff changeset
   283
  (fundef_parser default_config
23203
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
   284
     >> (fn ((config, fixes), (flags, statements)) =>
a5026e73cfcf "function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
krauss
parents: 23189
diff changeset
   285
            Toplevel.local_theory_to_proof (target_of config) (add_fundef fixes statements config flags)
21211
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
   286
            #> Toplevel.print));
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   287
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   288
val terminationP =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   289
  OuterSyntax.command "termination" "prove termination of a recursive function" K.thy_goal
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22725
diff changeset
   290
  (P.opt_target -- Scan.option P.term
21000
54c42dfbcafa Toplevel.local_theory_to_proof: proper target;
wenzelm
parents: 20877
diff changeset
   291
    >> (fn (target, name) =>
54c42dfbcafa Toplevel.local_theory_to_proof: proper target;
wenzelm
parents: 20877
diff changeset
   292
           Toplevel.print o
21211
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
   293
           Toplevel.local_theory_to_proof target (setup_termination_proof name)));
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   294
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   295
val _ = OuterSyntax.add_parsers [functionP];
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   296
val _ = OuterSyntax.add_parsers [terminationP];
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20638
diff changeset
   297
val _ = OuterSyntax.add_keywords ["sequential", "otherwise"];
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   298
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   299
end;
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   300
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   301
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19583
diff changeset
   302
end