| author | wenzelm | 
| Thu, 01 Oct 2009 18:10:41 +0200 | |
| changeset 32815 | 1a5e364584ae | 
| parent 32740 | 9dd0a2f83429 | 
| child 32966 | 5b21661fe618 | 
| permissions | -rw-r--r-- | 
| 31775 | 1 | (* Title: HOL/Tools/Function/fundef_common.ML | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 2 | Author: Alexander Krauss, TU Muenchen | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 3 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 4 | A package for general recursive function definitions. | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 5 | Common definitions and other infrastructure. | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 6 | *) | 
| 
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 | structure FundefCommon = | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 9 | struct | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 10 | |
| 23215 | 11 | local open FundefLib in | 
| 12 | ||
| 22498 
62cdd4b3e96b
made function syntax strict, requiring | to separate equations; cleanup
 krauss parents: 
22279diff
changeset | 13 | (* Profiling *) | 
| 32740 | 14 | val profile = Unsynchronized.ref false; | 
| 21255 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 15 | |
| 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 16 | fun PROFILE msg = if !profile then timeap_msg msg else I | 
| 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 17 | |
| 22498 
62cdd4b3e96b
made function syntax strict, requiring | to separate equations; cleanup
 krauss parents: 
22279diff
changeset | 18 | |
| 32603 | 19 | val acc_const_name = @{const_name accp}
 | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 20 | fun mk_acc domT R = | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 21 | Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 22 | |
| 21319 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21255diff
changeset | 23 | val function_name = suffix "C" | 
| 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21255diff
changeset | 24 | val graph_name = suffix "_graph" | 
| 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21255diff
changeset | 25 | val rel_name = suffix "_rel" | 
| 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21255diff
changeset | 26 | val dom_name = suffix "_dom" | 
| 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21255diff
changeset | 27 | |
| 28287 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 28 | (* Termination rules *) | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 29 | |
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 30 | structure TerminationRule = GenericDataFun | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 31 | ( | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 32 | type T = thm list | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 33 | val empty = [] | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 34 | val extend = I | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 35 | fun merge _ = Thm.merge_thms | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 36 | ); | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 37 | |
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 38 | val get_termination_rules = TerminationRule.get | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 39 | val store_termination_rule = TerminationRule.map o cons | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 40 | val apply_termination_rule = resolve_tac o get_termination_rules o Context.Proof | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 41 | |
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 42 | |
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 43 | (* Function definition result data *) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 44 | |
| 19583 
c5fa77b03442
function-package: Changed record usage to make sml/nj happy...
 krauss parents: 
19564diff
changeset | 45 | datatype fundef_result = | 
| 
c5fa77b03442
function-package: Changed record usage to make sml/nj happy...
 krauss parents: 
19564diff
changeset | 46 | FundefResult of | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 47 |      {
 | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 48 | fs: term list, | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 49 | G: term, | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 50 | R: term, | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 51 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 52 | psimps : thm list, | 
| 22166 | 53 | trsimps : thm list option, | 
| 54 | ||
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 55 | simple_pinducts : thm list, | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 56 | cases : thm, | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 57 | termination : thm, | 
| 22166 | 58 | domintros : thm list option | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 59 | } | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 60 | |
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 61 | |
| 21255 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 62 | datatype fundef_context_data = | 
| 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 63 | FundefCtxData of | 
| 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 64 |      {
 | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 65 | defname : string, | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 66 | |
| 23819 | 67 | (* contains no logical entities: invariant under morphisms *) | 
| 30790 | 68 | add_simps : (binding -> binding) -> string -> Attrib.src list -> thm list | 
| 26529 | 69 | -> local_theory -> thm list * local_theory, | 
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 70 | case_names : string list, | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19617diff
changeset | 71 | |
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 72 | fs : term list, | 
| 21255 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 73 | R : term, | 
| 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 74 | |
| 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 75 | psimps: thm list, | 
| 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 76 | pinducts: thm list, | 
| 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 77 | termination: thm | 
| 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 78 | } | 
| 
617fdb08abe9
added profiling code, improved handling of proof terms, generation of domain
 krauss parents: 
21237diff
changeset | 79 | |
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 80 | fun morph_fundef_data (FundefCtxData {add_simps, case_names, fs, R, 
 | 
| 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 81 | psimps, pinducts, termination, defname}) phi = | 
| 22623 | 82 | let | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 83 | val term = Morphism.term phi val thm = Morphism.thm phi val fact = Morphism.fact phi | 
| 30223 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 wenzelm parents: 
29922diff
changeset | 84 | val name = Binding.name_of o Morphism.binding phi o Binding.name | 
| 22623 | 85 | in | 
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 86 |       FundefCtxData { add_simps = add_simps, case_names = case_names,
 | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 87 | fs = map term fs, R = term R, psimps = fact psimps, | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 88 | pinducts = fact pinducts, termination = thm termination, | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 89 | defname = name defname } | 
| 22623 | 90 | end | 
| 91 | ||
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 92 | structure FundefData = GenericDataFun | 
| 22846 | 93 | ( | 
| 30560 | 94 | type T = (term * fundef_context_data) Item_Net.T; | 
| 95 | val empty = Item_Net.init | |
| 22760 | 96 | (op aconv o pairself fst : (term * fundef_context_data) * (term * fundef_context_data) -> bool) | 
| 97 | fst; | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 98 | val copy = I; | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 99 | val extend = I; | 
| 30560 | 100 | fun merge _ (tab1, tab2) = Item_Net.merge (tab1, tab2) | 
| 22846 | 101 | ); | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 102 | |
| 30492 | 103 | val get_fundef = FundefData.get o Context.Proof; | 
| 104 | ||
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 105 | |
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 106 | (* Generally useful?? *) | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 107 | fun lift_morphism thy f = | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 108 | let | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 109 | val term = Drule.term_rule thy f | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 110 | in | 
| 28883 | 111 | Morphism.thm_morphism f $> Morphism.term_morphism term | 
| 112 | $> Morphism.typ_morphism (Logic.type_map term) | |
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 113 | end | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 114 | |
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 115 | fun import_fundef_data t ctxt = | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 116 | let | 
| 30492 | 117 | val thy = ProofContext.theory_of ctxt | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 118 | val ct = cterm_of thy t | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 119 | val inst_morph = lift_morphism thy o Thm.instantiate | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 120 | |
| 25201 | 121 | fun match (trm, data) = | 
| 122 | SOME (morph_fundef_data data (inst_morph (Thm.match (cterm_of thy trm, ct)))) | |
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 123 | handle Pattern.MATCH => NONE | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 124 | in | 
| 30560 | 125 | get_first match (Item_Net.retrieve (get_fundef ctxt) t) | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 126 | end | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 127 | |
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 128 | fun import_last_fundef ctxt = | 
| 30560 | 129 | case Item_Net.content (get_fundef ctxt) of | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 130 | [] => NONE | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 131 | | (t, data) :: _ => | 
| 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 132 | let | 
| 30492 | 133 | val ([t'], ctxt') = Variable.import_terms true [t] ctxt | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 134 | in | 
| 30492 | 135 | import_fundef_data t' ctxt' | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 136 | end | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 137 | |
| 30560 | 138 | val all_fundef_data = Item_Net.content o get_fundef | 
| 21319 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21255diff
changeset | 139 | |
| 28287 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 140 | fun add_fundef_data (data as FundefCtxData {fs, termination, ...}) =
 | 
| 30560 | 141 | FundefData.map (fold (fn f => Item_Net.insert (f, data)) fs) | 
| 28287 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 142 | #> store_termination_rule termination | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 143 | |
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 144 | |
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 145 | (* Simp rules for termination proofs *) | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 146 | |
| 31902 | 147 | structure Termination_Simps = Named_Thms | 
| 26989 | 148 | ( | 
| 26749 
397a1aeede7d
* New attribute "termination_simp": Simp rules for termination proofs
 krauss parents: 
26748diff
changeset | 149 | val name = "termination_simp" | 
| 
397a1aeede7d
* New attribute "termination_simp": Simp rules for termination proofs
 krauss parents: 
26748diff
changeset | 150 | val description = "Simplification rule for termination proofs" | 
| 
397a1aeede7d
* New attribute "termination_simp": Simp rules for termination proofs
 krauss parents: 
26748diff
changeset | 151 | ); | 
| 
397a1aeede7d
* New attribute "termination_simp": Simp rules for termination proofs
 krauss parents: 
26748diff
changeset | 152 | |
| 28287 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 153 | |
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 154 | (* Default Termination Prover *) | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 155 | |
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 156 | structure TerminationProver = GenericDataFun | 
| 22846 | 157 | ( | 
| 30510 
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
 wenzelm parents: 
30492diff
changeset | 158 | type T = Proof.context -> Proof.method | 
| 28287 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 159 | val empty = (fn _ => error "Termination prover not configured") | 
| 22846 | 160 | val extend = I | 
| 28287 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 161 | fun merge _ (a,b) = b (* FIXME *) | 
| 22846 | 162 | ); | 
| 21319 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21255diff
changeset | 163 | |
| 28287 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 164 | val set_termination_prover = TerminationProver.put | 
| 30492 | 165 | val get_termination_prover = TerminationProver.get o Context.Proof | 
| 22733 
0b14bb35be90
definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
 krauss parents: 
22623diff
changeset | 166 | |
| 21319 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21255diff
changeset | 167 | |
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 168 | (* Configuration management *) | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 169 | datatype fundef_opt | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 170 | = Sequential | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 171 | | Default of string | 
| 21319 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21255diff
changeset | 172 | | DomIntros | 
| 22166 | 173 | | Tailrec | 
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 174 | |
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 175 | datatype fundef_config | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 176 | = FundefConfig of | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 177 |    {
 | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 178 | sequential: bool, | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 179 | default: string, | 
| 22166 | 180 | domintros: bool, | 
| 181 | tailrec: bool | |
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 182 | } | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 183 | |
| 26989 | 184 | fun apply_opt Sequential (FundefConfig {sequential, default, domintros,tailrec}) = 
 | 
| 185 |     FundefConfig {sequential=true, default=default, domintros=domintros, tailrec=tailrec}
 | |
| 186 |   | apply_opt (Default d) (FundefConfig {sequential, default, domintros,tailrec}) = 
 | |
| 187 |     FundefConfig {sequential=sequential, default=d, domintros=domintros, tailrec=tailrec}
 | |
| 188 |   | apply_opt DomIntros (FundefConfig {sequential, default, domintros,tailrec}) =
 | |
| 189 |     FundefConfig {sequential=sequential, default=default, domintros=true,tailrec=tailrec}
 | |
| 190 |   | apply_opt Tailrec (FundefConfig {sequential, default, domintros,tailrec}) =
 | |
| 191 |     FundefConfig {sequential=sequential, default=default, domintros=domintros,tailrec=true}
 | |
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20523diff
changeset | 192 | |
| 26989 | 193 | val default_config = | 
| 28883 | 194 |   FundefConfig { sequential=false, default="%x. undefined" (*FIXME dynamic scoping*), 
 | 
| 195 | domintros=false, tailrec=false } | |
| 23819 | 196 | |
| 197 | ||
| 28287 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 198 | (* Analyzing function equations *) | 
| 23189 | 199 | |
| 24170 | 200 | fun split_def ctxt geq = | 
| 23189 | 201 | let | 
| 24920 | 202 | fun input_error msg = cat_lines [msg, Syntax.string_of_term ctxt geq] | 
| 28287 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 203 | val qs = Term.strip_qnt_vars "all" geq | 
| 
c86fa4e0aedb
termination prover for "fun" can be configured using context data.
 krauss parents: 
28083diff
changeset | 204 | val imp = Term.strip_qnt_body "all" geq | 
| 26529 | 205 | val (gs, eq) = Logic.strip_horn imp | 
| 23189 | 206 | |
| 207 | val (f_args, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq) | |
| 24170 | 208 | handle TERM _ => error (input_error "Not an equation") | 
| 23189 | 209 | |
| 210 | val (head, args) = strip_comb f_args | |
| 211 | ||
| 212 | val fname = fst (dest_Free head) | |
| 24170 | 213 | handle TERM _ => error (input_error "Head symbol must not be a bound variable") | 
| 23189 | 214 | in | 
| 215 | (fname, qs, gs, args, rhs) | |
| 216 | end | |
| 217 | ||
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 218 | (* Check for all sorts of errors in the input *) | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 219 | fun check_defs ctxt fixes eqs = | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 220 | let | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 221 | val fnames = map (fst o fst) fixes | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 222 | |
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 223 | fun check geq = | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 224 | let | 
| 28883 | 225 | fun input_error msg = error (cat_lines [msg, Syntax.string_of_term ctxt geq]) | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 226 | |
| 24170 | 227 | val fqgar as (fname, qs, gs, args, rhs) = split_def ctxt geq | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 228 | |
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 229 | val _ = fname mem fnames | 
| 28883 | 230 | orelse input_error | 
| 231 |                              ("Head symbol of left hand side must be " 
 | |
| 232 | ^ plural "" "one out of " fnames ^ commas_quote fnames) | |
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 233 | |
| 30832 
521f7d801da1
explicitly check that at least one argument is present to avoid low-level exception
 krauss parents: 
30791diff
changeset | 234 | val _ = length args > 0 orelse input_error "Function has no arguments:" | 
| 
521f7d801da1
explicitly check that at least one argument is present to avoid low-level exception
 krauss parents: 
30791diff
changeset | 235 | |
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 236 | fun add_bvs t is = add_loose_bnos (t, 0, is) | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 237 | val rvs = (add_bvs rhs [] \\ fold add_bvs args []) | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 238 | |> map (fst o nth (rev qs)) | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 239 | |
| 28883 | 240 | val _ = null rvs orelse input_error | 
| 241 |                         ("Variable" ^ plural " " "s " rvs ^ commas_quote rvs
 | |
| 242 | ^ " occur" ^ plural "s" "" rvs ^ " on right hand side only:") | |
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 243 | |
| 28883 | 244 | val _ = forall (not o Term.exists_subterm | 
| 29922 
60a304bc5a07
reject defined function in patterns with errmsg, e.g. f (f x) = x
 krauss parents: 
29006diff
changeset | 245 | (fn Free (n, _) => n mem fnames | _ => false)) (gs @ args) | 
| 
60a304bc5a07
reject defined function in patterns with errmsg, e.g. f (f x) = x
 krauss parents: 
29006diff
changeset | 246 | orelse input_error "Defined function may not occur in premises or arguments" | 
| 24171 
25381ce95316
Issue a warning, when "function" encounters variables occuring in function position,
 krauss parents: 
24170diff
changeset | 247 | |
| 
25381ce95316
Issue a warning, when "function" encounters variables occuring in function position,
 krauss parents: 
24170diff
changeset | 248 | val freeargs = map (fn t => subst_bounds (rev (map Free qs), t)) args | 
| 
25381ce95316
Issue a warning, when "function" encounters variables occuring in function position,
 krauss parents: 
24170diff
changeset | 249 | val funvars = filter (fn q => exists (exists_subterm (fn (Free q') $ _ => q = q' | _ => false)) freeargs) qs | 
| 
25381ce95316
Issue a warning, when "function" encounters variables occuring in function position,
 krauss parents: 
24170diff
changeset | 250 | val _ = null funvars | 
| 28883 | 251 | orelse (warning (cat_lines | 
| 252 | ["Bound variable" ^ plural " " "s " funvars | |
| 253 | ^ commas_quote (map fst funvars) ^ | |
| 254 | " occur" ^ plural "s" "" funvars ^ " in function position.", | |
| 255 | "Misspelled constructor???"]); true) | |
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 256 | in | 
| 30906 
3c7a76e79898
simplify computation and consistency checks of argument counts in the input
 krauss parents: 
30832diff
changeset | 257 | (fname, length args) | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 258 | end | 
| 30906 
3c7a76e79898
simplify computation and consistency checks of argument counts in the input
 krauss parents: 
30832diff
changeset | 259 | |
| 
3c7a76e79898
simplify computation and consistency checks of argument counts in the input
 krauss parents: 
30832diff
changeset | 260 | val _ = AList.group (op =) (map check eqs) | 
| 
3c7a76e79898
simplify computation and consistency checks of argument counts in the input
 krauss parents: 
30832diff
changeset | 261 | |> map (fn (fname, ars) => | 
| 
3c7a76e79898
simplify computation and consistency checks of argument counts in the input
 krauss parents: 
30832diff
changeset | 262 | length (distinct (op =) ars) = 1 | 
| 
3c7a76e79898
simplify computation and consistency checks of argument counts in the input
 krauss parents: 
30832diff
changeset | 263 |              orelse error ("Function " ^ quote fname ^
 | 
| 
3c7a76e79898
simplify computation and consistency checks of argument counts in the input
 krauss parents: 
30832diff
changeset | 264 | " has different numbers of arguments in different equations")) | 
| 28884 | 265 | |
| 24170 | 266 | fun check_sorts ((fname, fT), _) = | 
| 267 | Sorts.of_sort (Sign.classes_of (ProofContext.theory_of ctxt)) (fT, HOLogic.typeS) | |
| 28883 | 268 | orelse error (cat_lines | 
| 269 | ["Type of " ^ quote fname ^ " is not of sort " ^ quote "type" ^ ":", | |
| 270 | setmp show_sorts true (Syntax.string_of_typ ctxt) fT]) | |
| 24170 | 271 | |
| 272 | val _ = map check_sorts fixes | |
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 273 | in | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 274 | () | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 275 | end | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 276 | |
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 277 | (* Preprocessors *) | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 278 | |
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 279 | type fixes = ((string * typ) * mixfix) list | 
| 30790 | 280 | type 'a spec = (Attrib.binding * 'a list) list | 
| 30787 
5b7a5a05c7aa
remove "otherwise" feature from function package which was never really documented or used
 krauss parents: 
30560diff
changeset | 281 | type preproc = fundef_config -> Proof.context -> fixes -> term spec | 
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 282 | -> (term list * (thm list -> thm spec) * (thm list -> thm list list) * string list) | 
| 23819 | 283 | |
| 28883 | 284 | val fname_of = fst o dest_Free o fst o strip_comb o fst | 
| 285 | o HOLogic.dest_eq o HOLogic.dest_Trueprop o Logic.strip_imp_concl o snd o dest_all_all | |
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 286 | |
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 287 | fun mk_case_names i "" k = mk_case_names i (string_of_int (i + 1)) k | 
| 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 288 | | mk_case_names _ n 0 = [] | 
| 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 289 | | mk_case_names _ n 1 = [n] | 
| 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 290 | | mk_case_names _ n k = map (fn i => n ^ "_" ^ string_of_int i) (1 upto k) | 
| 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 291 | |
| 30787 
5b7a5a05c7aa
remove "otherwise" feature from function package which was never really documented or used
 krauss parents: 
30560diff
changeset | 292 | fun empty_preproc check _ ctxt fixes spec = | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 293 | let | 
| 30790 | 294 | val (bnds, tss) = split_list spec | 
| 23819 | 295 | val ts = flat tss | 
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 296 | val _ = check ctxt fixes ts | 
| 23819 | 297 | val fnames = map (fst o fst) fixes | 
| 298 | val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) ts | |
| 299 | ||
| 28883 | 300 | fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) | 
| 301 | (indices ~~ xs) | |
| 23819 | 302 | |> map (map snd) | 
| 25222 
78943ac46f6d
fun/function: generate case names for induction rules
 krauss parents: 
25201diff
changeset | 303 | |
| 25361 
1aa441e48496
function package: using the names of the equations as case names turned out to be impractical => disabled
 krauss parents: 
25222diff
changeset | 304 | (* using theorem names for case name currently disabled *) | 
| 30790 | 305 | val cnames = map_index (fn (i, _) => mk_case_names i "" 1) bnds |> flat | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 306 | in | 
| 30790 | 307 | (ts, curry op ~~ bnds o Library.unflat tss, sort, cnames) | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 308 | end | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 309 | |
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 310 | structure Preprocessor = GenericDataFun | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 311 | ( | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 312 | type T = preproc | 
| 23206 | 313 | val empty : T = empty_preproc check_defs | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 314 | val extend = I | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 315 | fun merge _ (a, _) = a | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 316 | ); | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 317 | |
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 318 | val get_preproc = Preprocessor.get o Context.Proof | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 319 | val set_preproc = Preprocessor.map o K | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 320 | |
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 321 | |
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 322 | |
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 323 | local | 
| 25558 | 324 | structure P = OuterParse and K = OuterKeyword | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 325 | |
| 28883 | 326 | val option_parser = | 
| 327 | P.group "option" ((P.reserved "sequential" >> K Sequential) | |
| 328 | || ((P.reserved "default" |-- P.term) >> Default) | |
| 329 | || (P.reserved "domintros" >> K DomIntros) | |
| 330 | || (P.reserved "tailrec" >> K Tailrec)) | |
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 331 | |
| 28883 | 332 | fun config_parser default = | 
| 333 |       (Scan.optional (P.$$$ "(" |-- P.!!! (P.list1 option_parser) --| P.$$$ ")") [])
 | |
| 334 | >> (fn opts => fold apply_opt opts default) | |
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 335 | in | 
| 28883 | 336 | fun fundef_parser default_cfg = | 
| 30791 | 337 | config_parser default_cfg -- P.fixes -- SpecParse.where_alt_specs | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 338 | end | 
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 339 | |
| 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
23189diff
changeset | 340 | |
| 23215 | 341 | end | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 342 | end | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 343 |