| author | krauss | 
| Wed, 21 Jun 2006 11:08:04 +0200 | |
| changeset 19934 | 8190655ea2d4 | 
| parent 19770 | be5c23ebe1eb | 
| child 20270 | 3abe7dae681e | 
| 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: 
19564 
diff
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: 
19564 
diff
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: 
19564 
diff
changeset
 | 
11  | 
("Tools/function_package/mutual.ML")
 | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
12  | 
("Tools/function_package/fundef_package.ML")
 | 
| 
19770
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
13  | 
("Tools/function_package/fundef_datatype.ML")
 | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
14  | 
("Tools/function_package/auto_term.ML")
 | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
15  | 
begin  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
16  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
17  | 
lemma fundef_ex1_existence:  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
18  | 
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
 | 
19  | 
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
 | 
20  | 
shows "(x, f x)\<in>G"  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
21  | 
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
 | 
22  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
23  | 
lemma fundef_ex1_uniqueness:  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
24  | 
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
 | 
25  | 
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
 | 
26  | 
assumes elm: "(x, h x)\<in>G"  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
27  | 
shows "h x = f x"  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
28  | 
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
 | 
29  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
30  | 
lemma fundef_ex1_iff:  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
31  | 
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
 | 
32  | 
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
 | 
33  | 
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
 | 
34  | 
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
 | 
35  | 
by (rule theI', rule ex1)  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
36  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
37  | 
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
 | 
38  | 
by simp  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
40  | 
|
| 
19770
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
41  | 
subsection {* Projections *}
 | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
42  | 
consts  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
43  | 
  lpg::"(('a + 'b) * 'a) set"
 | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
44  | 
  rpg::"(('a + 'b) * 'b) set"
 | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
45  | 
|
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
46  | 
inductive lpg  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
47  | 
intros  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
48  | 
"(Inl x, x) : lpg"  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
49  | 
inductive rpg  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
50  | 
intros  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
51  | 
"(Inr y, y) : rpg"  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
52  | 
definition  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
53  | 
"lproj x = (THE y. (x,y) : lpg)"  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
54  | 
"rproj x = (THE y. (x,y) : rpg)"  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
55  | 
|
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
56  | 
lemma lproj_inl:  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
57  | 
"lproj (Inl x) = x"  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
58  | 
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: 
19564 
diff
changeset
 | 
59  | 
lemma rproj_inr:  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
60  | 
"rproj (Inr x) = x"  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
61  | 
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: 
19564 
diff
changeset
 | 
62  | 
|
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
63  | 
|
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
64  | 
|
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
65  | 
|
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
66  | 
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
 | 
67  | 
use "Tools/function_package/fundef_common.ML"  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
68  | 
use "Tools/function_package/fundef_lib.ML"  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
69  | 
use "Tools/function_package/context_tree.ML"  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
70  | 
use "Tools/function_package/fundef_prep.ML"  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
71  | 
use "Tools/function_package/fundef_proof.ML"  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
72  | 
use "Tools/function_package/termination.ML"  | 
| 
19770
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
73  | 
use "Tools/function_package/mutual.ML"  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
74  | 
use "Tools/function_package/fundef_package.ML"  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
75  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
76  | 
setup FundefPackage.setup  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
77  | 
|
| 
19770
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
78  | 
use "Tools/function_package/fundef_datatype.ML"  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
79  | 
setup FundefDatatype.setup  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
80  | 
|
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
81  | 
use "Tools/function_package/auto_term.ML"  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
82  | 
setup FundefAutoTerm.setup  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
83  | 
|
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
84  | 
|
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
85  | 
lemmas [fundef_cong] =  | 
| 
 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 
krauss 
parents: 
19564 
diff
changeset
 | 
86  | 
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
 | 
87  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
88  | 
|
| 19934 | 89  | 
lemma split_cong[fundef_cong]:  | 
90  | 
"\<lbrakk> \<And>x y. (x, y) = q \<Longrightarrow> f x y = g x y; p = q \<rbrakk>  | 
|
91  | 
\<Longrightarrow> split f p = split g q"  | 
|
92  | 
by (auto simp:split_def)  | 
|
93  | 
||
94  | 
||
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
95  | 
end  |