| author | haftmann | 
| Thu, 10 Apr 2008 00:46:38 +0200 | |
| changeset 26596 | 07d7d0a6d5fd | 
| parent 26594 | 1c676ae50311 | 
| child 26628 | 63306cb94313 | 
| 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 | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 13 | -> ((bstring * Attrib.src list) * string) list | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 14 | -> FundefCommon.fundef_config | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 15 | -> bool list | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 16 | -> local_theory | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 17 | -> Proof.state | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 18 | |
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 19 | val add_fundef_i: (string * typ option * mixfix) list | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 20 | -> ((bstring * Attrib.src list) * term) list | 
| 22170 | 21 | -> FundefCommon.fundef_config | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 22 | -> bool list | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 23 | -> local_theory | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 24 | -> Proof.state | 
| 21211 
5370cfbf3070
Preparations for making "lexicographic_order" part of "fun"
 krauss parents: 
21051diff
changeset | 25 | |
| 
5370cfbf3070
Preparations for making "lexicographic_order" part of "fun"
 krauss parents: 
21051diff
changeset | 26 | val setup_termination_proof : string option -> local_theory -> Proof.state | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 27 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 28 | val setup : theory -> theory | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 29 | val get_congs : theory -> thm list | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 30 | end | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 31 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 32 | |
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 33 | structure FundefPackage : FUNDEF_PACKAGE = | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 34 | struct | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 35 | |
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
21000diff
changeset | 36 | open FundefLib | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 37 | open FundefCommon | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 38 | |
| 21435 | 39 | val note_theorem = LocalTheory.note Thm.theoremK | 
| 40 | ||
| 22166 | 41 | fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_" | 
| 42 | ||
| 26529 | 43 | fun add_simps fnames post sort extra_qualify label moreatts simps lthy = | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 44 | let | 
| 23819 | 45 | val atts = Attrib.internal (K Simplifier.simp_add) :: moreatts | 
| 46 | val spec = post simps | |
| 47 | |> map (apfst (apsnd (append atts))) | |
| 26529 | 48 | |> map (apfst (apfst extra_qualify)) | 
| 21255 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21240diff
changeset | 49 | |
| 22166 | 50 | val (saved_spec_simps, lthy) = | 
| 23819 | 51 | fold_map note_theorem spec lthy | 
| 52 | ||
| 22498 
62cdd4b3e96b
made function syntax strict, requiring | to separate equations; cleanup
 krauss parents: 
22325diff
changeset | 53 | val saved_simps = flat (map snd saved_spec_simps) | 
| 22166 | 54 | 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 | 55 | |
| 22166 | 56 | fun add_for_f fname simps = | 
| 23819 | 57 | note_theorem ((NameSpace.qualified fname label, []), simps) #> snd | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 58 | in | 
| 22166 | 59 | (saved_simps, | 
| 60 | 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 | 61 | end | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 62 | |
| 26594 
1c676ae50311
fundef_afterqed: removed unused config, added do_print flag;
 wenzelm parents: 
26580diff
changeset | 63 | fun fundef_afterqed do_print fixes post defname cont sort_cont cnames [[proof]] lthy = | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 64 | let | 
| 26529 | 65 |       val FundefResult {fs, R, psimps, trsimps,  simple_pinducts, termination, domintros, cases, ...} = 
 | 
| 22166 | 66 | cont (Goal.close_result proof) | 
| 67 | ||
| 23819 | 68 | val fnames = map (fst o fst) fixes | 
| 21391 
a8809f46bd7f
replaced NameSpace.append by NameSpace.qualified, which handles empty names as expected;
 wenzelm parents: 
21359diff
changeset | 69 | val qualify = NameSpace.qualified defname | 
| 23819 | 70 | val addsmps = add_simps fnames post sort_cont | 
| 21602 | 71 | |
| 72 | val (((psimps', pinducts'), (_, [termination'])), lthy) = | |
| 21255 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21240diff
changeset | 73 | lthy | 
| 26529 | 74 | |> addsmps (NameSpace.qualified "partial") "psimps" [] psimps | 
| 75 | ||> fold_option (snd oo addsmps I "simps" []) trsimps | |
| 22166 | 76 | ||>> note_theorem ((qualify "pinduct", | 
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 77 | [Attrib.internal (K (RuleCases.case_names cnames)), | 
| 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 78 | Attrib.internal (K (RuleCases.consumes 1)), | 
| 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 79 | Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts) | 
| 22166 | 80 | ||>> note_theorem ((qualify "termination", []), [termination]) | 
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 81 | ||> (snd o note_theorem ((qualify "cases", | 
| 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 82 | [Attrib.internal (K (RuleCases.case_names cnames))]), [cases])) | 
| 22166 | 83 | ||> 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: 
21240diff
changeset | 84 | |
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 85 |       val cdata = FundefCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps',
 | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 86 | pinducts=snd pinducts', termination=termination', fs=fs, R=R, defname=defname } | 
| 26594 
1c676ae50311
fundef_afterqed: removed unused config, added do_print flag;
 wenzelm parents: 
26580diff
changeset | 87 | val _ = | 
| 
1c676ae50311
fundef_afterqed: removed unused config, added do_print flag;
 wenzelm parents: 
26580diff
changeset | 88 | if not do_print then () | 
| 
1c676ae50311
fundef_afterqed: removed unused config, added do_print flag;
 wenzelm parents: 
26580diff
changeset | 89 | else Specification.print_consts lthy (K false) (map fst fixes) | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 90 | in | 
| 22623 | 91 | lthy | 
| 25201 | 92 | |> LocalTheory.declaration (add_fundef_data o morph_fundef_data cdata) | 
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 93 | end | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 94 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 95 | |
| 26594 
1c676ae50311
fundef_afterqed: removed unused config, added do_print flag;
 wenzelm parents: 
26580diff
changeset | 96 | fun gen_add_fundef is_external prep fixspec eqnss config flags lthy = | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 97 | let | 
| 25088 
9a13ab12b174
Simultaneous type inference using read_specification
 krauss parents: 
25045diff
changeset | 98 | val ((fixes, spec), ctxt') = prep fixspec (map (fn (n_a, eq) => [(n_a, [eq])]) eqnss) lthy | 
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 99 | val (eqs, post, sort_cont, cnames) = FundefCommon.get_preproc lthy config flags ctxt' fixes spec | 
| 22166 | 100 | |
| 101 | val defname = mk_defname fixes | |
| 102 | ||
| 23819 | 103 | val ((goalstate, cont), lthy) = | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 104 | FundefMutual.prepare_fundef_mutual config defname fixes eqs lthy | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 105 | |
| 26594 
1c676ae50311
fundef_afterqed: removed unused config, added do_print flag;
 wenzelm parents: 
26580diff
changeset | 106 | val afterqed = fundef_afterqed is_external fixes post defname cont sort_cont cnames | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 107 | in | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 108 | lthy | 
| 26594 
1c676ae50311
fundef_afterqed: removed unused config, added do_print flag;
 wenzelm parents: 
26580diff
changeset | 109 | |> is_external ? LocalTheory.set_group (serial_string ()) | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 110 | |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]] | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 111 | |> 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: 
20363diff
changeset | 112 | end | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 113 | |
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 114 | fun total_termination_afterqed data [[totality]] lthy = | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 115 | let | 
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 116 |       val FundefCtxData { add_simps, case_names, psimps, pinducts, defname, ... } = data
 | 
| 21255 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21240diff
changeset | 117 | |
| 22166 | 118 | val totality = Goal.close_result totality | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 119 | |
| 21255 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21240diff
changeset | 120 | 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 | 121 | |
| 22166 | 122 | val tsimps = map remove_domain_condition psimps | 
| 123 | val tinduct = map remove_domain_condition pinducts | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 124 | |
| 21511 
16c62deb1adf
The function package declares the [code] attribute automatically again.
 krauss parents: 
21491diff
changeset | 125 |       val has_guards = exists ((fn (Const ("Trueprop", _) $ _) => false | _ => true) o prop_of) tsimps
 | 
| 24624 
b8383b1bbae3
distinction between regular and default code theorems
 haftmann parents: 
24508diff
changeset | 126 | val allatts = if has_guards then [] else [Attrib.internal (K RecfunCodegen.add_default)] | 
| 21602 | 127 | |
| 21511 
16c62deb1adf
The function package declares the [code] attribute automatically again.
 krauss parents: 
21491diff
changeset | 128 | val qualify = NameSpace.qualified defname; | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 129 | in | 
| 21511 
16c62deb1adf
The function package declares the [code] attribute automatically again.
 krauss parents: 
21491diff
changeset | 130 | lthy | 
| 26529 | 131 | |> add_simps I "simps" allatts tsimps |> snd | 
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 132 | |> note_theorem ((qualify "induct", [Attrib.internal (K (RuleCases.case_names case_names))]), tinduct) |> snd | 
| 20363 
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20338diff
changeset | 133 | end | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 134 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 135 | |
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 136 | fun setup_termination_proof term_opt lthy = | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 137 | let | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 138 | val data = the (case term_opt of | 
| 24508 
c8b82fec6447
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 wenzelm parents: 
24168diff
changeset | 139 | SOME t => import_fundef_data (Syntax.read_term lthy t) (Context.Proof lthy) | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 140 | | NONE => import_last_fundef (Context.Proof lthy)) | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 141 |           handle Option.Option => raise ERROR ("Not a function: " ^ quote (the_default "" term_opt))
 | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 142 | |
| 22166 | 143 |         val FundefCtxData {termination, R, ...} = data
 | 
| 22325 | 144 | val domT = domain_type (fastype_of R) | 
| 145 |         val goal = HOLogic.mk_Trueprop (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 146 | in | 
| 20877 
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
 wenzelm parents: 
20654diff
changeset | 147 | lthy | 
| 22325 | 148 |         |> ProofContext.note_thmss_i "" [(("", [ContextRules.rule_del]), [([allI], [])])] |> snd
 | 
| 149 |         |> ProofContext.note_thmss_i "" [(("", [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
 | |
| 21602 | 150 | |> ProofContext.note_thmss_i "" | 
| 22325 | 151 |           [(("termination", [ContextRules.intro_bang (SOME 0)]),
 | 
| 21602 | 152 | [([Goal.norm_result termination], [])])] |> snd | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 153 | |> Proof.theorem_i NONE (total_termination_afterqed data) [[(goal, [])]] | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 154 | end | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 155 | |
| 26171 
5426a823455c
more precise handling of "group" for termination;
 wenzelm parents: 
26129diff
changeset | 156 | fun termination_cmd term_opt lthy = | 
| 
5426a823455c
more precise handling of "group" for termination;
 wenzelm parents: 
26129diff
changeset | 157 | lthy | 
| 
5426a823455c
more precise handling of "group" for termination;
 wenzelm parents: 
26129diff
changeset | 158 | |> LocalTheory.set_group (serial_string ()) | 
| 
5426a823455c
more precise handling of "group" for termination;
 wenzelm parents: 
26129diff
changeset | 159 | |> setup_termination_proof term_opt; | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 160 | |
| 26129 | 161 | val add_fundef = gen_add_fundef true Specification.read_specification | 
| 162 | val add_fundef_i = gen_add_fundef false Specification.check_specification | |
| 19611 
da2a0014c461
Function Package: Quick-and-dirty-fixed strange "Proved a different theorem bug"
 krauss parents: 
19585diff
changeset | 163 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 164 | |
| 21235 
674e2731b519
Added datatype hook to declare all case_congs as "fundef_cong" automatically.
 krauss parents: 
21211diff
changeset | 165 | (* Datatype hook to declare datatype congs as "fundef_congs" *) | 
| 
674e2731b519
Added datatype hook to declare all case_congs as "fundef_cong" automatically.
 krauss parents: 
21211diff
changeset | 166 | |
| 
674e2731b519
Added datatype hook to declare all case_congs as "fundef_cong" automatically.
 krauss parents: 
21211diff
changeset | 167 | |
| 21602 | 168 | fun add_case_cong n thy = | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
24039diff
changeset | 169 | Context.theory_map (FundefCtxTree.map_fundef_congs (Thm.add_thm | 
| 21235 
674e2731b519
Added datatype hook to declare all case_congs as "fundef_cong" automatically.
 krauss parents: 
21211diff
changeset | 170 | (DatatypePackage.get_datatype thy n |> the | 
| 
674e2731b519
Added datatype hook to declare all case_congs as "fundef_cong" automatically.
 krauss parents: 
21211diff
changeset | 171 | |> #case_cong | 
| 21602 | 172 | |> safe_mk_meta_eq))) | 
| 21235 
674e2731b519
Added datatype hook to declare all case_congs as "fundef_cong" automatically.
 krauss parents: 
21211diff
changeset | 173 | thy | 
| 
674e2731b519
Added datatype hook to declare all case_congs as "fundef_cong" automatically.
 krauss parents: 
21211diff
changeset | 174 | |
| 24626 
85eceef2edc7
introduced generic concepts for theory interpretators
 haftmann parents: 
24624diff
changeset | 175 | val case_cong = fold add_case_cong | 
| 21235 
674e2731b519
Added datatype hook to declare all case_congs as "fundef_cong" automatically.
 krauss parents: 
21211diff
changeset | 176 | |
| 24711 | 177 | val setup_case_cong = DatatypePackage.interpretation case_cong | 
| 24626 
85eceef2edc7
introduced generic concepts for theory interpretators
 haftmann parents: 
24624diff
changeset | 178 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 179 | (* setup *) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 180 | |
| 21602 | 181 | val setup = | 
| 22846 | 182 | Attrib.add_attributes | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
24039diff
changeset | 183 |     [("fundef_cong", Attrib.add_del_args FundefCtxTree.cong_add FundefCtxTree.cong_del,
 | 
| 22846 | 184 | "declaration of congruence rule for function definitions")] | 
| 24626 
85eceef2edc7
introduced generic concepts for theory interpretators
 haftmann parents: 
24624diff
changeset | 185 | #> setup_case_cong | 
| 22846 | 186 | #> FundefRelation.setup | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 187 | |
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
24039diff
changeset | 188 | val get_congs = FundefCtxTree.get_fundef_congs o Context.Theory | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 189 | |
| 21235 
674e2731b519
Added datatype hook to declare all case_congs as "fundef_cong" automatically.
 krauss parents: 
21211diff
changeset | 190 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 191 | (* outer syntax *) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 192 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 193 | 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 | 194 | |
| 25045 
12386aefe9ac
"sequential" is no longer a keyword. It is still used as before, but as a normal
 krauss parents: 
24977diff
changeset | 195 | val _ = OuterSyntax.keywords ["otherwise"]; | 
| 24867 | 196 | |
| 197 | val _ = | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 198 | OuterSyntax.command "function" "define general recursive functions" K.thy_goal | 
| 22498 
62cdd4b3e96b
made function syntax strict, requiring | to separate equations; cleanup
 krauss parents: 
22325diff
changeset | 199 | (fundef_parser default_config | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 200 | >> (fn ((config, fixes), (flags, statements)) => | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 201 | Toplevel.local_theory_to_proof (target_of config) (add_fundef fixes statements config flags) | 
| 21211 
5370cfbf3070
Preparations for making "lexicographic_order" part of "fun"
 krauss parents: 
21051diff
changeset | 202 | #> Toplevel.print)); | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20363diff
changeset | 203 | |
| 24867 | 204 | val _ = | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 205 | OuterSyntax.command "termination" "prove termination of a recursive function" K.thy_goal | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22725diff
changeset | 206 | (P.opt_target -- Scan.option P.term | 
| 21000 | 207 | >> (fn (target, name) => | 
| 208 | Toplevel.print o | |
| 26171 
5426a823455c
more precise handling of "group" for termination;
 wenzelm parents: 
26129diff
changeset | 209 | Toplevel.local_theory_to_proof target (termination_cmd name))); | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 210 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 211 | end; | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 212 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 213 | |
| 19585 | 214 | end |