| author | wenzelm | 
| Mon, 18 Sep 2006 19:12:44 +0200 | |
| changeset 20572 | 2b88de40da57 | 
| parent 20536 | f088edff8af8 | 
| child 20654 | d80502f0d701 | 
| permissions | -rw-r--r-- | 
| 20324 | 1 | (* Title: HOL/FunDef.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Alexander Krauss, TU Muenchen | |
| 4 | ||
| 5 | A package for general recursive function definitions. | |
| 6 | *) | |
| 7 | ||
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 8 | theory FunDef | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 9 | imports Accessible_Part Datatype Recdef | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 10 | uses | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 11 | ("Tools/function_package/sum_tools.ML")
 | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 12 | ("Tools/function_package/fundef_common.ML")
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 13 | ("Tools/function_package/fundef_lib.ML")
 | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20324diff
changeset | 14 | ("Tools/function_package/inductive_wrap.ML")
 | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 15 | ("Tools/function_package/context_tree.ML")
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 16 | ("Tools/function_package/fundef_prep.ML")
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 17 | ("Tools/function_package/fundef_proof.ML")
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 18 | ("Tools/function_package/termination.ML")
 | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 19 | ("Tools/function_package/mutual.ML")
 | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19934diff
changeset | 20 | ("Tools/function_package/pattern_split.ML")
 | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 21 | ("Tools/function_package/fundef_package.ML")
 | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 22 | ("Tools/function_package/fundef_datatype.ML")
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 23 | ("Tools/function_package/auto_term.ML")
 | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 24 | begin | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 25 | |
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 26 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 27 | definition | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 28 |   THE_default :: "'a \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a"
 | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 29 | "THE_default d P = (if (\<exists>!x. P x) then (THE x. P x) else d)" | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 30 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 31 | lemma THE_defaultI': "\<exists>!x. P x \<Longrightarrow> P (THE_default d P)" | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 32 | by (simp add:theI' THE_default_def) | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 33 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 34 | lemma THE_default1_equality: | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 35 | "\<lbrakk>\<exists>!x. P x; P a\<rbrakk> \<Longrightarrow> THE_default d P = a" | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 36 | by (simp add:the1_equality THE_default_def) | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 37 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 38 | lemma THE_default_none: | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 39 | "\<not>(\<exists>!x. P x) \<Longrightarrow> THE_default d P = d" | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 40 | by (simp add:THE_default_def) | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 41 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 42 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 43 | lemma fundef_ex1_existence: | 
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 44 | assumes f_def: "f \<equiv> \<lambda>x. THE_default d (\<lambda>y. (x,y)\<in>G)" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 45 | assumes ex1: "\<exists>!y. (x,y)\<in>G" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 46 | shows "(x, f x)\<in>G" | 
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 47 | by (simp only:f_def, rule THE_defaultI', rule ex1) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 48 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 49 | lemma fundef_ex1_uniqueness: | 
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 50 | assumes f_def: "f \<equiv> \<lambda>x. THE_default d (\<lambda>y. (x,y)\<in>G)" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 51 | assumes ex1: "\<exists>!y. (x,y)\<in>G" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 52 | assumes elm: "(x, h x)\<in>G" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 53 | shows "h x = f x" | 
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 54 | by (simp only:f_def, rule THE_default1_equality[symmetric], rule ex1, rule elm) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 55 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 56 | lemma fundef_ex1_iff: | 
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 57 | assumes f_def: "f \<equiv> \<lambda>x. THE_default d (\<lambda>y. (x,y)\<in>G)" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 58 | assumes ex1: "\<exists>!y. (x,y)\<in>G" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 59 | shows "((x, y)\<in>G) = (f x = y)" | 
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 60 | apply (auto simp:ex1 f_def THE_default1_equality) | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 61 | by (rule THE_defaultI', rule ex1) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 62 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 63 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 64 | subsection {* Projections *}
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 65 | consts | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 66 |   lpg::"(('a + 'b) * 'a) set"
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 67 |   rpg::"(('a + 'b) * 'b) set"
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 68 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 69 | inductive lpg | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 70 | intros | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 71 | "(Inl x, x) : lpg" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 72 | inductive rpg | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 73 | intros | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 74 | "(Inr y, y) : rpg" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 75 | definition | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 76 | "lproj x = (THE y. (x,y) : lpg)" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 77 | "rproj x = (THE y. (x,y) : rpg)" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 78 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 79 | lemma lproj_inl: | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 80 | "lproj (Inl x) = x" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 81 | by (auto simp:lproj_def intro: the_equality lpg.intros elim: lpg.cases) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 82 | lemma rproj_inr: | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 83 | "rproj (Inr x) = x" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 84 | by (auto simp:rproj_def intro: the_equality rpg.intros elim: rpg.cases) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 85 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 86 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 87 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 88 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 89 | use "Tools/function_package/sum_tools.ML" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 90 | use "Tools/function_package/fundef_common.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 91 | use "Tools/function_package/fundef_lib.ML" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20324diff
changeset | 92 | use "Tools/function_package/inductive_wrap.ML" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 93 | use "Tools/function_package/context_tree.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 94 | use "Tools/function_package/fundef_prep.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 95 | use "Tools/function_package/fundef_proof.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 96 | use "Tools/function_package/termination.ML" | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 97 | use "Tools/function_package/mutual.ML" | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19934diff
changeset | 98 | use "Tools/function_package/pattern_split.ML" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 99 | use "Tools/function_package/fundef_package.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 100 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 101 | setup FundefPackage.setup | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 102 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 103 | use "Tools/function_package/fundef_datatype.ML" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 104 | setup FundefDatatype.setup | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 105 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 106 | use "Tools/function_package/auto_term.ML" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 107 | setup FundefAutoTerm.setup | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 108 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 109 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 110 | lemmas [fundef_cong] = | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 111 | let_cong if_cong image_cong INT_cong UN_cong bex_cong ball_cong imp_cong | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 112 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 113 | |
| 19934 | 114 | lemma split_cong[fundef_cong]: | 
| 115 | "\<lbrakk> \<And>x y. (x, y) = q \<Longrightarrow> f x y = g x y; p = q \<rbrakk> | |
| 116 | \<Longrightarrow> split f p = split g q" | |
| 117 | by (auto simp:split_def) | |
| 118 | ||
| 119 | ||
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 120 | end |