| author | webertj | 
| Tue, 01 Aug 2006 14:58:43 +0200 | |
| changeset 20276 | d94dc40673b1 | 
| parent 20270 | 3abe7dae681e | 
| child 20324 | 71d63a30cc96 | 
| permissions | -rw-r--r-- | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 1 | theory FunDef | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 2 | imports Accessible_Part Datatype Recdef | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 3 | uses | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 4 | ("Tools/function_package/sum_tools.ML")
 | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 5 | ("Tools/function_package/fundef_common.ML")
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 6 | ("Tools/function_package/fundef_lib.ML")
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 7 | ("Tools/function_package/context_tree.ML")
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 8 | ("Tools/function_package/fundef_prep.ML")
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 9 | ("Tools/function_package/fundef_proof.ML")
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 10 | ("Tools/function_package/termination.ML")
 | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 11 | ("Tools/function_package/mutual.ML")
 | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19934diff
changeset | 12 | ("Tools/function_package/pattern_split.ML")
 | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 13 | ("Tools/function_package/fundef_package.ML")
 | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 14 | ("Tools/function_package/fundef_datatype.ML")
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 15 | ("Tools/function_package/auto_term.ML")
 | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 16 | begin | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 17 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 18 | lemma fundef_ex1_existence: | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 19 | assumes f_def: "\<And>x. f x \<equiv> THE y. (x,y)\<in>G" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 20 | 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 | 21 | shows "(x, f x)\<in>G" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 22 | by (simp only:f_def, rule theI', rule ex1) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 23 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 24 | lemma fundef_ex1_uniqueness: | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 25 | assumes f_def: "\<And>x. f x \<equiv> THE y. (x,y)\<in>G" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 26 | 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 | 27 | assumes elm: "(x, h x)\<in>G" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 28 | shows "h x = f x" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 29 | by (simp only:f_def, rule the1_equality[symmetric], rule ex1, rule elm) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 30 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 31 | lemma fundef_ex1_iff: | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 32 | assumes f_def: "\<And>x. f x \<equiv> THE y. (x,y)\<in>G" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 33 | 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 | 34 | shows "((x, y)\<in>G) = (f x = y)" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 35 | apply (auto simp:ex1 f_def the1_equality) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 36 | by (rule theI', rule ex1) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 37 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 38 | lemma True_implies: "(True \<Longrightarrow> PROP P) \<equiv> PROP P" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 39 | by simp | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 40 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 41 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 42 | subsection {* Projections *}
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 43 | consts | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 44 |   lpg::"(('a + 'b) * 'a) set"
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 45 |   rpg::"(('a + 'b) * 'b) set"
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 46 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 47 | inductive lpg | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 48 | intros | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 49 | "(Inl x, x) : lpg" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 50 | inductive rpg | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 51 | intros | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 52 | "(Inr y, y) : rpg" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 53 | definition | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 54 | "lproj x = (THE y. (x,y) : lpg)" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 55 | "rproj x = (THE y. (x,y) : rpg)" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 56 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 57 | lemma lproj_inl: | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 58 | "lproj (Inl x) = x" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 59 | 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 | 60 | lemma rproj_inr: | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 61 | "rproj (Inr x) = x" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 62 | 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 | 63 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 64 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 65 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 66 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 67 | 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 | 68 | use "Tools/function_package/fundef_common.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 69 | use "Tools/function_package/fundef_lib.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 70 | use "Tools/function_package/context_tree.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 71 | use "Tools/function_package/fundef_prep.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 72 | use "Tools/function_package/fundef_proof.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 73 | use "Tools/function_package/termination.ML" | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 74 | use "Tools/function_package/mutual.ML" | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19934diff
changeset | 75 | 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 | 76 | use "Tools/function_package/fundef_package.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 77 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 78 | setup FundefPackage.setup | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 79 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 80 | use "Tools/function_package/fundef_datatype.ML" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 81 | setup FundefDatatype.setup | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 82 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 83 | use "Tools/function_package/auto_term.ML" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 84 | setup FundefAutoTerm.setup | 
| 
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 | lemmas [fundef_cong] = | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 88 | 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 | 89 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 90 | |
| 19934 | 91 | lemma split_cong[fundef_cong]: | 
| 92 | "\<lbrakk> \<And>x y. (x, y) = q \<Longrightarrow> f x y = g x y; p = q \<rbrakk> | |
| 93 | \<Longrightarrow> split f p = split g q" | |
| 94 | by (auto simp:split_def) | |
| 95 | ||
| 96 | ||
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 97 | end |