src/HOL/Tools/function_package/fundef_package.ML
author krauss
Mon, 22 Jan 2007 17:29:43 +0100
changeset 22166 0a50d4db234a
parent 22102 a89ef7144729
child 22170 5682eeaefaf4
permissions -rw-r--r--
* Preliminary implementation of tail recursion * Clarified internal interfaces
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
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    13
                      -> ((bstring * Attrib.src list) * string list) list list
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    14
                      -> FundefCommon.fundef_config
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    15
                      -> local_theory
21211
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
    16
                      -> string * Proof.state
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    17
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    18
    val add_fundef_i:  (string * typ option * mixfix) list
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    19
                       -> ((bstring * Attrib.src list) * term list) list list
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    20
                       -> bool
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
    21
                       -> local_theory
21211
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
    22
                       -> string * Proof.state
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
    23
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
    24
    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
    25
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    26
    val cong_add: attribute
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    27
    val cong_del: attribute
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
    28
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    29
    val setup : theory -> theory
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
    30
    val setup_case_cong_hook : theory -> theory
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
    31
    val get_congs : theory -> thm list
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    32
end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    33
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    34
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    35
structure FundefPackage  =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    36
struct
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    37
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 21000
diff changeset
    38
open FundefLib
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    39
open FundefCommon
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    40
21435
883337ea2f3b LocalTheory.axioms/notes/defs: proper kind;
wenzelm
parents: 21391
diff changeset
    41
val note_theorem = LocalTheory.note Thm.theoremK
883337ea2f3b LocalTheory.axioms/notes/defs: proper kind;
wenzelm
parents: 21391
diff changeset
    42
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    43
fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_" 
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    44
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    45
fun burrow_snd f ps = (* ('a list -> 'b list) -> ('c * 'a) list -> ('c * 'b) list *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    46
    let val (xs, ys) = split_list ps
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    47
    in xs ~~ f ys end
20270
3abe7dae681e Function package can now do automatic splits of overlapping datatype patterns
krauss
parents: 19938
diff changeset
    48
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    49
fun restore_spec_structure reps spec =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    50
    (burrow o burrow_snd o burrow o K) reps spec
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
    51
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    52
fun add_simps fixes spec sort label moreatts simps lthy =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    53
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    54
      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
    55
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    56
      val (saved_spec_simps, lthy) =
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    57
        fold_map (fold_map note_theorem) (restore_spec_structure simps spec) lthy
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    58
      val saved_simps = flat (map snd (flat saved_spec_simps))
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
    59
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    60
      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
    61
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    62
      fun add_for_f fname simps =
21435
883337ea2f3b LocalTheory.axioms/notes/defs: proper kind;
wenzelm
parents: 21391
diff changeset
    63
        note_theorem
21658
5e31241e1e3c Attrib.internal: morphism;
wenzelm
parents: 21602
diff changeset
    64
          ((NameSpace.qualified fname label, Attrib.internal (K Simplifier.simp_add) :: moreatts),
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    65
            simps) #> snd
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    66
    in
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    67
      (saved_simps,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    68
       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
    69
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    70
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    71
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    72
fun fundef_afterqed config fixes spec 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
    73
    let
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    74
      val FundefResult {f, R, psimps, trsimps, subset_pinducts, simple_pinducts, termination, domintros, cases, ...} = 
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    75
          cont (Goal.close_result proof)
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    76
21391
a8809f46bd7f replaced NameSpace.append by NameSpace.qualified, which handles empty names as expected;
wenzelm
parents: 21359
diff changeset
    77
      val qualify = NameSpace.qualified defname
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    78
      val addsimps = add_simps fixes spec sort_cont
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
    79
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
    80
      val (((psimps', pinducts'), (_, [termination'])), lthy) =
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
    81
          lthy
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    82
            |> addsimps "psimps" [] psimps
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    83
            ||> fold_option (snd oo addsimps "simps" []) trsimps
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    84
            ||>> note_theorem ((qualify "pinduct",
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    85
                                [Attrib.internal (K (InductAttrib.induct_set ""))]), simple_pinducts)
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    86
            ||>> note_theorem ((qualify "termination", []), [termination])
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    87
            ||> (snd o note_theorem ((qualify "cases", []), [cases]))
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    88
            ||> 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
    89
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
    90
      val cdata = FundefCtxData { add_simps=addsimps, psimps=psimps',
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
    91
                                  pinducts=snd pinducts', termination=termination', f=f, R=R }
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    92
    in
21491
574e63799847 declarations: pass morphism (dummy);
wenzelm
parents: 21435
diff changeset
    93
      lthy  (* FIXME proper handling of morphism *)
574e63799847 declarations: pass morphism (dummy);
wenzelm
parents: 21435
diff changeset
    94
        |> LocalTheory.declaration (fn phi => add_fundef_data defname cdata) (* save in target *)
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
    95
        |> Context.proof_map (add_fundef_data defname 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
    96
    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
    97
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    98
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    99
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   100
fun prep_with_flags prep fixspec eqnss_flags global_flag lthy =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   101
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   102
      val eqnss = map (map (apsnd (map fst))) eqnss_flags
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   103
      val flags = map (map (map (fn (_, f) => global_flag orelse f) o snd)) eqnss_flags
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   104
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   105
      val ((fixes, _), ctxt') = prep fixspec [] lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   106
      val spec = map (fn eqns => snd (fst (prep [] eqns ctxt'))) eqnss
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   107
                     |> map (map (apsnd (map (fn t => fold_rev (mk_forall o Free) (frees_in_term ctxt' t) t)))) (* Add quantifiers *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   108
                     |> map2 (map2 (fn fs => fn (r, thms) => (r, fs ~~ thms))) flags
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
   109
                     |> (burrow o burrow_snd o burrow)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   110
                          (FundefSplit.split_some_equations lthy)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   111
                     |> map (map (apsnd flat))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   112
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   113
      ((fixes, spec), ctxt')
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   114
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   115
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   116
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20638
diff changeset
   117
fun gen_add_fundef prep_spec fixspec eqnss_flags config lthy =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   118
    let
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   119
      val FundefConfig {sequential, default, tailrec, ...} = config
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20638
diff changeset
   120
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20638
diff changeset
   121
      val ((fixes, spec), ctxt') = prep_with_flags prep_spec fixspec eqnss_flags sequential lthy
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   122
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   123
      val defname = mk_defname fixes
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   124
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   125
      val t_eqns = spec
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   126
                     |> flat |> map snd |> flat (* flatten external structure *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   127
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   128
      val ((goalstate, cont, sort_cont), lthy) =
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   129
          FundefMutual.prepare_fundef_mutual config defname fixes t_eqns default lthy
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   130
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   131
      val afterqed = fundef_afterqed config fixes spec defname cont sort_cont
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   132
    in
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   133
      (defname, lthy
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   134
         |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   135
         |> 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
   136
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   137
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   138
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   139
fun total_termination_afterqed defname data [[totality]] lthy =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   140
    let
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   141
      val FundefCtxData { add_simps, psimps, pinducts, ... } = data
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
   142
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   143
      val totality = Goal.close_result totality
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   144
21255
617fdb08abe9 added profiling code, improved handling of proof terms, generation of domain
krauss
parents: 21240
diff changeset
   145
      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
   146
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   147
      val tsimps = map remove_domain_condition psimps
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   148
      val tinduct = map remove_domain_condition pinducts
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
   149
21511
16c62deb1adf The function package declares the [code] attribute automatically again.
krauss
parents: 21491
diff changeset
   150
        (* FIXME: How to generate code from (possibly) local contexts*)
16c62deb1adf The function package declares the [code] attribute automatically again.
krauss
parents: 21491
diff changeset
   151
      val has_guards = exists ((fn (Const ("Trueprop", _) $ _) => false | _ => true) o prop_of) tsimps
21658
5e31241e1e3c Attrib.internal: morphism;
wenzelm
parents: 21602
diff changeset
   152
      val allatts = if has_guards then [] else [Attrib.internal (K (RecfunCodegen.add NONE))]
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   153
21511
16c62deb1adf The function package declares the [code] attribute automatically again.
krauss
parents: 21491
diff changeset
   154
      val qualify = NameSpace.qualified defname;
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   155
    in
21511
16c62deb1adf The function package declares the [code] attribute automatically again.
krauss
parents: 21491
diff changeset
   156
      lthy
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   157
        |> add_simps "simps" allatts tsimps |> snd
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   158
        |> 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
   159
    end
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   160
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   161
21211
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
   162
fun setup_termination_proof name_opt lthy =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   163
    let
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   164
        val defname = the_default (get_last_fundef (Context.Proof lthy)) name_opt
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   165
        val data = the (get_fundef_data defname (Context.Proof lthy))
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   166
                   handle Option.Option => raise ERROR ("No such function definition: " ^ defname)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   167
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   168
        val FundefCtxData {termination, R, ...} = data
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   169
        val goal = FundefTermination.mk_total_termination_goal R
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   170
    in
20877
368b997ad67e removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents: 20654
diff changeset
   171
      lthy
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   172
        |> ProofContext.note_thmss_i ""
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   173
          [(("termination", [ContextRules.intro_query NONE]),
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   174
            [([Goal.norm_result termination], [])])] |> snd
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21255
diff changeset
   175
        |> set_termination_rule termination
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 22102
diff changeset
   176
        |> Proof.theorem_i NONE (total_termination_afterqed defname data) [[(goal, [])]]
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   177
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   178
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   179
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   180
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
   181
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
   182
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   183
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   184
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   185
(* congruence rules *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   186
19617
7cb4b67d4b97 avoid raw equality on type thm;
wenzelm
parents: 19611
diff changeset
   187
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
   188
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
   189
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   190
(* 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
   191
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   192
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   193
fun add_case_cong n thy =
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   194
    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
   195
                          (DatatypePackage.get_datatype thy n |> the
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   196
                           |> #case_cong
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   197
                           |> safe_mk_meta_eq)))
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   198
                       thy
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   199
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   200
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
   201
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   202
val setup_case_cong_hook =
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   203
    DatatypeHooks.add case_cong_hook
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   204
    #> (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
   205
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   206
(* setup *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   207
21602
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   208
val setup =
cb13024d0e36 Goal.norm/close_result;
wenzelm
parents: 21511
diff changeset
   209
    FundefData.init
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   210
      #> FundefCongs.init
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21255
diff changeset
   211
      #> TerminationRule.init
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   212
      #>  Attrib.add_attributes
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   213
            [("fundef_cong", Attrib.add_del_args cong_add cong_del, "declaration of congruence rule for function definitions")]
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   214
      #> setup_case_cong_hook
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21255
diff changeset
   215
      #> FundefRelation.setup
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   216
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
   217
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
   218
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
   219
21235
674e2731b519 Added datatype hook to declare all case_congs as "fundef_cong" automatically.
krauss
parents: 21211
diff changeset
   220
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   221
(* outer syntax *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   222
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   223
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
   224
20270
3abe7dae681e Function package can now do automatic splits of overlapping datatype patterns
krauss
parents: 19938
diff changeset
   225
3abe7dae681e Function package can now do automatic splits of overlapping datatype patterns
krauss
parents: 19938
diff changeset
   226
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   227
fun or_list1 s = P.enum1 "|" s
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   228
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   229
val opt_sequential = Scan.optional ((P.$$$ "(" |-- P.$$$ "sequential" --| P.$$$ ")") >> K true) false
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   230
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   231
val otherwise = P.$$$ "(" |-- P.$$$ "otherwise" --| P.$$$ ")"
22102
a89ef7144729 moved parts of OuterParse to SpecParse;
wenzelm
parents: 21658
diff changeset
   232
val statement_ow = P.and_list1 (SpecParse.opt_thm_name ":" -- Scan.repeat1 (P.prop -- Scan.optional (otherwise >> K true) false))
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   233
val statements_ow = or_list1 statement_ow
20270
3abe7dae681e Function package can now do automatic splits of overlapping datatype patterns
krauss
parents: 19938
diff changeset
   234
3abe7dae681e Function package can now do automatic splits of overlapping datatype patterns
krauss
parents: 19938
diff changeset
   235
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   236
val functionP =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   237
  OuterSyntax.command "function" "define general recursive functions" K.thy_goal
22102
a89ef7144729 moved parts of OuterParse to SpecParse;
wenzelm
parents: 21658
diff changeset
   238
  ((config_parser default_config -- P.opt_target -- P.fixes --| P.$$$ "where" -- statements_ow)
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20638
diff changeset
   239
     >> (fn (((config, target), fixes), statements) =>
21211
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
   240
            Toplevel.local_theory_to_proof target (add_fundef fixes statements config #> snd)
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
   241
            #> Toplevel.print));
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   242
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   243
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 21000
diff changeset
   244
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   245
val terminationP =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   246
  OuterSyntax.command "termination" "prove termination of a recursive function" K.thy_goal
22102
a89ef7144729 moved parts of OuterParse to SpecParse;
wenzelm
parents: 21658
diff changeset
   247
  (P.opt_target -- Scan.option P.name
21000
54c42dfbcafa Toplevel.local_theory_to_proof: proper target;
wenzelm
parents: 20877
diff changeset
   248
    >> (fn (target, name) =>
54c42dfbcafa Toplevel.local_theory_to_proof: proper target;
wenzelm
parents: 20877
diff changeset
   249
           Toplevel.print o
21211
5370cfbf3070 Preparations for making "lexicographic_order" part of "fun"
krauss
parents: 21051
diff changeset
   250
           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
   251
20338
ecdfc96cf4d0 Added Keywords: "otherwise" and "sequential", needed for function package's
krauss
parents: 20320
diff changeset
   252
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   253
val _ = OuterSyntax.add_parsers [functionP];
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   254
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
   255
val _ = OuterSyntax.add_keywords ["sequential", "otherwise"];
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   256
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   257
end;
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   258
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   259
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19583
diff changeset
   260
end