src/HOL/Tools/function_package/fundef_package.ML
author krauss
Wed, 20 Sep 2006 12:05:31 +0200
changeset 20638 241792a4634e
parent 20523 36a59e5d0039
child 20654 d80502f0d701
permissions -rw-r--r--
Removed "induct set" attribute from total induction rules
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
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    12
    val add_fundef :  (string * string option * mixfix) list 
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
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    14
                      -> bool 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    15
                      -> local_theory 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    16
                      -> Proof.state
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    17
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    18
    val add_fundef_i:  (string * typ option * mixfix) list 
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
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    20
                       -> bool 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    21
                       -> local_theory 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    22
                       -> Proof.state
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    23
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    24
    val cong_add: attribute
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    25
    val cong_del: attribute
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
    26
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    27
    val setup : theory -> theory
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
    28
    val get_congs : theory -> thm list
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    29
end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    30
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    31
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    32
structure FundefPackage  =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    33
struct
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    34
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    35
open FundefCommon
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    36
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    37
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    38
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
    39
    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
    40
    in xs ~~ f ys end
20270
3abe7dae681e Function package can now do automatic splits of overlapping datatype patterns
krauss
parents: 19938
diff changeset
    41
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    42
fun restore_spec_structure reps spec =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    43
    (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
    44
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    45
fun with_local_path path f lthy =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    46
    let 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    47
      val restore = Theory.restore_naming (ProofContext.theory_of lthy)
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
    48
    in
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    49
      lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    50
        |> LocalTheory.theory (Theory.add_path path)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    51
        |> f
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    52
        |> LocalTheory.theory restore
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
    53
    end
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
    54
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    55
fun add_simps label moreatts mutual_info fixes psimps spec lthy =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    56
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    57
      val fnames = map (fst o fst) fixes
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    58
      val psimps_by_f = FundefMutual.sort_by_function mutual_info fnames psimps
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    59
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    60
      fun add_for_f fname psimps =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    61
          with_local_path fname
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    62
                          (LocalTheory.note ((label, Attrib.internal Simplifier.simp_add :: moreatts), psimps) #> snd)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    63
    in
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    64
      lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    65
        |> fold_map (fold_map LocalTheory.note) (restore_spec_structure psimps spec) |> snd
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    66
        |> fold2 add_for_f fnames psimps_by_f
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    67
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    68
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    69
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    70
fun fundef_afterqed fixes spec mutual_info defname data [[result]] lthy =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    71
    let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    72
        val fundef_data = FundefMutual.mk_partial_rules_mutual lthy mutual_info data result
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    73
        val FundefMResult {psimps, subset_pinducts, simple_pinducts, termination, domintros, cases, ...} = fundef_data
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    74
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    75
      lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    76
        |> add_simps "psimps" [] mutual_info fixes psimps spec
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    77
        |> with_local_path defname
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    78
                  (LocalTheory.note (("domintros", []), domintros) #> snd
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    79
                   #> LocalTheory.note (("termination", []), [termination]) #> snd
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    80
                   #> LocalTheory.note (("cases", []), [cases]) #> snd
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    81
                   #> LocalTheory.note (("pinduct", [Attrib.internal (InductAttrib.induct_set "")]), simple_pinducts) #> snd)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    82
        |> LocalTheory.theory (Context.theory_map (add_fundef_data defname (fundef_data, mutual_info, (fixes,spec))))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    83
    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
    84
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    85
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    86
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    87
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
    88
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    89
      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
    90
      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
    91
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    92
      val ((fixes, _), ctxt') = prep fixspec [] lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    93
      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
    94
                     |> 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
    95
                     |> map2 (map2 (fn fs => fn (r, thms) => (r, fs ~~ thms))) flags
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    96
                     |> (burrow o burrow_snd o burrow) 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    97
                          (FundefSplit.split_some_equations lthy)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    98
                     |> map (map (apsnd flat))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
    99
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   100
      ((fixes, spec), ctxt')
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   101
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   102
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   103
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   104
fun gen_add_fundef prep_spec fixspec eqnss_flags preprocess lthy =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   105
    let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   106
      val ((fixes, spec), ctxt') = prep_with_flags prep_spec fixspec eqnss_flags preprocess lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   107
      val t_eqns = spec
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   108
                     |> 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
   109
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   110
      val ((mutual_info, name, prep_result as Prep {goal, goalI, ...}), lthy) = 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   111
          FundefMutual.prepare_fundef_mutual fixes t_eqns lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   112
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   113
      val afterqed = fundef_afterqed fixes spec mutual_info name prep_result
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   114
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   115
        lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   116
          |> Proof.theorem_i PureThy.internalK NONE afterqed NONE ("", []) [(("", []), [(goal, [])])]
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   117
          |> Proof.refine (Method.primitive_text (fn _ => goalI)) |> Seq.hd
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   118
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   119
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   120
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   121
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
   122
    let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   123
        val (FundefMResult {psimps, simple_pinducts, ... }, mutual_info, (fixes, stmts)) = data
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   124
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
   125
        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
   126
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   127
        val tsimps = map remove_domain_condition psimps
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
   128
        val tinduct = map remove_domain_condition simple_pinducts
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
   129
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   130
        (* FIXME: How to generate code from (possibly) local contexts 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   131
        val has_guards = exists ((fn (Const ("Trueprop", _) $ _) => false | _ => true) o prop_of) tsimps
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   132
        val allatts = if has_guards then [] else [Attrib.internal (RecfunCodegen.add NONE)]
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   133
        *)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   134
    in
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   135
        lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   136
          |> add_simps "simps" [] mutual_info fixes tsimps stmts
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   137
          |> with_local_path defname
20638
241792a4634e Removed "induct set" attribute from total induction rules
krauss
parents: 20523
diff changeset
   138
                (LocalTheory.note (("induct", []), tinduct) #> snd)
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
   139
    end
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   140
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   141
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   142
fun fundef_setup_termination_proof name_opt lthy =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   143
    let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   144
        val name = the_default (get_last_fundef (Context.Proof lthy)) name_opt
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   145
        val data = the (get_fundef_data name (Context.Proof lthy))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   146
                   handle Option.Option => raise ERROR ("No such function definition: " ^ name)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   147
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   148
        val (res as FundefMResult {termination, ...}, _, _) = data
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   149
        val goal = FundefTermination.mk_total_termination_goal data
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   150
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   151
      lthy 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   152
        |> ProofContext.note_thmss_i [(("termination",
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   153
                                 [ContextRules.intro_query NONE]), [(ProofContext.export_standard lthy lthy [termination], [])])] |> snd
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   154
        |> Proof.theorem_i PureThy.internalK NONE
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   155
                           (total_termination_afterqed name data) NONE ("", [])
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   156
                           [(("", []), [(goal, [])])]
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   157
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   158
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   159
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   160
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
   161
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
   162
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   163
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   164
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   165
(* congruence rules *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   166
19617
7cb4b67d4b97 avoid raw equality on type thm;
wenzelm
parents: 19611
diff changeset
   167
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
   168
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
   169
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
(* setup *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   172
20363
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
   173
val setup = FundefData.init #> FundefCongs.init
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
   174
        #>  Attrib.add_attributes
f34c5dbe74d5 global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents: 20338
diff changeset
   175
                [("fundef_cong", Attrib.add_del_args cong_add cong_del, "declaration of congruence rule for function definitions")]
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   176
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   177
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
   178
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
   179
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19617
diff changeset
   180
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   181
(* outer syntax *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   182
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   183
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
   184
20270
3abe7dae681e Function package can now do automatic splits of overlapping datatype patterns
krauss
parents: 19938
diff changeset
   185
3abe7dae681e Function package can now do automatic splits of overlapping datatype patterns
krauss
parents: 19938
diff changeset
   186
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   187
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
   188
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   189
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
   190
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   191
val otherwise = P.$$$ "(" |-- P.$$$ "otherwise" --| P.$$$ ")"
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   192
val statement_ow = P.and_list1 (P.opt_thm_name ":" -- Scan.repeat1 (P.prop -- Scan.optional (otherwise >> K true) false))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   193
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
   194
3abe7dae681e Function package can now do automatic splits of overlapping datatype patterns
krauss
parents: 19938
diff changeset
   195
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   196
fun local_theory_to_proof f = 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   197
    Toplevel.theory_to_proof (f o LocalTheory.init NONE)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   198
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   199
val functionP =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   200
  OuterSyntax.command "function" "define general recursive functions" K.thy_goal
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   201
  ((opt_sequential -- P.opt_locale_target -- P.fixes --| P.$$$ "where" -- statements_ow)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   202
     >> (fn (((sequential, target), fixes), statements) =>
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   203
            Toplevel.print o local_theory_to_proof (add_fundef fixes statements sequential)));
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   204
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
val terminationP =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   207
  OuterSyntax.command "termination" "prove termination of a recursive function" K.thy_goal
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   208
  (Scan.option P.name 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   209
    >> (fn name => Toplevel.print o local_theory_to_proof (fundef_setup_termination_proof name)));
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   210
20338
ecdfc96cf4d0 Added Keywords: "otherwise" and "sequential", needed for function package's
krauss
parents: 20320
diff changeset
   211
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   212
val _ = OuterSyntax.add_parsers [functionP];
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   213
val _ = OuterSyntax.add_parsers [terminationP];
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20363
diff changeset
   214
val _ = OuterSyntax.add_keywords ["sequential", "otherwise"]; (* currently unused *)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   215
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   216
end;
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   217
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   218
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19583
diff changeset
   219
end