| author | haftmann | 
| Fri, 03 Nov 2006 14:22:38 +0100 | |
| changeset 21152 | e97992896170 | 
| parent 21051 | c49467a9c1e1 | 
| child 21211 | 5370cfbf3070 | 
| permissions | -rw-r--r-- | 
| 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: 
20338diff
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: 
20338diff
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: 
20654diff
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: 
20363diff
changeset | 13 | -> ((bstring * Attrib.src list) * string list) list list | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 14 | -> FundefCommon.fundef_config | 
| 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 15 | -> local_theory | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 16 | -> Proof.state | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 17 | |
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
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: 
20363diff
changeset | 19 | -> ((bstring * Attrib.src list) * term list) list list | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 20 | -> bool | 
| 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 21 | -> local_theory | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
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: 
20338diff
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: 
19617diff
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: 
20363diff
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 | |
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
21000diff
changeset | 35 | open FundefLib | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 36 | open FundefCommon | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 37 | |
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20638diff
changeset | 38 | (*fn (ctxt,s) => Variable.importT_terms (fst (ProofContext.read_termTs ctxt (K false) (K NONE) (K NONE) [] [s])) ctxt;*) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 39 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 40 | 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: 
20363diff
changeset | 41 | let val (xs, ys) = split_list ps | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 42 | in xs ~~ f ys end | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19938diff
changeset | 43 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 44 | fun restore_spec_structure reps spec = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 45 | (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: 
20338diff
changeset | 46 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 47 | 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 | 48 | let | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 49 | val fnames = map (fst o fst) fixes | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 50 | 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 | 51 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 52 | fun add_for_f fname psimps = | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 53 | LocalTheory.note ((NameSpace.append fname 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 | 54 | in | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 55 | lthy | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 56 | |> 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: 
20363diff
changeset | 57 | |> 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 | 58 | end | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 59 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 60 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 61 | 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: 
20363diff
changeset | 62 | let | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 63 | 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: 
20363diff
changeset | 64 |         val FundefMResult {psimps, subset_pinducts, simple_pinducts, termination, domintros, cases, ...} = fundef_data
 | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 65 | val qualify = NameSpace.append defname; | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 66 | in | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 67 | lthy | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 68 | |> add_simps "psimps" [] mutual_info fixes psimps spec | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 69 | |> LocalTheory.note ((qualify "domintros", []), domintros) |> snd | 
| 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 70 | |> LocalTheory.note ((qualify "termination", []), [termination]) |> snd | 
| 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 71 | |> LocalTheory.note ((qualify "cases", []), [cases]) |> snd | 
| 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 72 | |> LocalTheory.note ((qualify "pinduct", [Attrib.internal (InductAttrib.induct_set "")]), simple_pinducts) |> snd | 
| 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 73 | |> LocalTheory.declaration (add_fundef_data defname (fundef_data, mutual_info, (fixes,spec))) | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 74 | end (* FIXME: Add cases for induct and cases thm *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 75 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 76 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 77 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 78 | 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 | 79 | let | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 80 | val eqnss = map (map (apsnd (map fst))) eqnss_flags | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 81 | 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: 
20363diff
changeset | 82 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 83 | val ((fixes, _), ctxt') = prep fixspec [] lthy | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 84 | 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: 
20363diff
changeset | 85 | |> 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: 
20363diff
changeset | 86 | |> map2 (map2 (fn fs => fn (r, thms) => (r, fs ~~ thms))) flags | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 87 | |> (burrow o burrow_snd o burrow) | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 88 | (FundefSplit.split_some_equations lthy) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 89 | |> map (map (apsnd flat)) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 90 | in | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 91 | ((fixes, spec), ctxt') | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 92 | end | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 93 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 94 | |
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20638diff
changeset | 95 | 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: 
20363diff
changeset | 96 | let | 
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20638diff
changeset | 97 |       val FundefConfig {sequential, default, ...} = config
 | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20638diff
changeset | 98 | |
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20638diff
changeset | 99 | val ((fixes, spec), ctxt') = prep_with_flags prep_spec fixspec eqnss_flags sequential lthy | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 100 | val t_eqns = spec | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 101 | |> flat |> map snd |> flat (* flatten external structure *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 102 | |
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 103 |       val ((mutual_info, name, prep_result as Prep {goal, goalI, ...}), lthy) =
 | 
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20638diff
changeset | 104 | FundefMutual.prepare_fundef_mutual fixes t_eqns default lthy | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 105 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 106 | 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: 
20363diff
changeset | 107 | in | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 108 | lthy | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 109 |           |> Proof.theorem_i PureThy.internalK NONE afterqed NONE ("", []) [(("", []), [(goal, [])])]
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 110 | |> Proof.refine (Method.primitive_text (fn _ => goalI)) |> Seq.hd | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 111 | end | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 112 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 113 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 114 | fun total_termination_afterqed defname data [[totality]] lthy = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 115 | let | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 116 |         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 | 117 | |
| 20363 
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20338diff
changeset | 118 | 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 | 119 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 120 | val tsimps = map remove_domain_condition psimps | 
| 20363 
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20338diff
changeset | 121 | val tinduct = map remove_domain_condition simple_pinducts | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 122 | |
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 123 | (* FIXME: How to generate code from (possibly) local contexts | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 124 |         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: 
20363diff
changeset | 125 | 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: 
20363diff
changeset | 126 | *) | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 127 | val qualify = NameSpace.append defname; | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 128 | in | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 129 | lthy | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 130 | |> add_simps "simps" [] mutual_info fixes tsimps stmts | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 131 | |> LocalTheory.note ((qualify "induct", []), tinduct) |> snd | 
| 20363 
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20338diff
changeset | 132 | end | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 133 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 134 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 135 | fun fundef_setup_termination_proof name_opt lthy = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 136 | let | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 137 | 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: 
20363diff
changeset | 138 | 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: 
20363diff
changeset | 139 |                    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: 
20363diff
changeset | 140 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 141 |         val (res as FundefMResult {termination, ...}, _, _) = data
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 142 | val goal = FundefTermination.mk_total_termination_goal data | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 143 | in | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 144 | lthy | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 145 |         |> ProofContext.note_thmss_i [(("termination",
 | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 146 | [ContextRules.intro_query NONE]), [(ProofContext.standard lthy [termination], [])])] |> snd | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 147 | |> Proof.theorem_i PureThy.internalK NONE | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 148 |                            (total_termination_afterqed name data) NONE ("", [])
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 149 |                            [(("", []), [(goal, [])])]
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 150 | end | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 151 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 152 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 153 | val add_fundef = gen_add_fundef Specification.read_specification | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 154 | 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: 
19585diff
changeset | 155 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 156 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 157 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 158 | (* congruence rules *) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 159 | |
| 19617 | 160 | val cong_add = Thm.declaration_attribute (map_fundef_congs o Drule.add_rule o safe_mk_meta_eq); | 
| 161 | 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 | 162 | |
| 
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 | (* setup *) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 165 | |
| 20363 
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20338diff
changeset | 166 | val setup = FundefData.init #> FundefCongs.init | 
| 
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20338diff
changeset | 167 | #> Attrib.add_attributes | 
| 
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20338diff
changeset | 168 |                 [("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 | 169 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 170 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 171 | val get_congs = FundefCommon.get_fundef_congs o Context.Theory | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 172 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 173 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 174 | (* outer syntax *) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 175 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 176 | 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 | 177 | |
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19938diff
changeset | 178 | |
| 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19938diff
changeset | 179 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 180 | fun or_list1 s = P.enum1 "|" s | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 181 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 182 | 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: 
20363diff
changeset | 183 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 184 | val otherwise = P.$$$ "(" |-- P.$$$ "otherwise" --| P.$$$ ")"
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 185 | 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: 
20363diff
changeset | 186 | val statements_ow = or_list1 statement_ow | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19938diff
changeset | 187 | |
| 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19938diff
changeset | 188 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 189 | val functionP = | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 190 | OuterSyntax.command "function" "define general recursive functions" K.thy_goal | 
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
21000diff
changeset | 191 | ((config_parser default_config -- P.opt_locale_target -- P.fixes --| P.$$$ "where" -- statements_ow) | 
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20638diff
changeset | 192 | >> (fn (((config, target), fixes), statements) => | 
| 21000 | 193 | Toplevel.print o | 
| 194 | Toplevel.local_theory_to_proof target (add_fundef fixes statements config))); | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 195 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 196 | |
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
21000diff
changeset | 197 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 198 | val terminationP = | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 199 | OuterSyntax.command "termination" "prove termination of a recursive function" K.thy_goal | 
| 21000 | 200 | (P.opt_locale_target -- Scan.option P.name | 
| 201 | >> (fn (target, name) => | |
| 202 | Toplevel.print o | |
| 203 | Toplevel.local_theory_to_proof target (fundef_setup_termination_proof name))); | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 204 | |
| 20338 
ecdfc96cf4d0
Added Keywords: "otherwise" and "sequential", needed for function package's
 krauss parents: 
20320diff
changeset | 205 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 206 | val _ = OuterSyntax.add_parsers [functionP]; | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 207 | val _ = OuterSyntax.add_parsers [terminationP]; | 
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20638diff
changeset | 208 | val _ = OuterSyntax.add_keywords ["sequential", "otherwise"]; | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 209 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 210 | end; | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 211 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 212 | |
| 19585 | 213 | end |