| author | wenzelm | 
| Fri, 26 Oct 2001 14:02:58 +0200 | |
| changeset 11944 | 0594e63e6057 | 
| parent 9907 | 473a6604da94 | 
| child 12199 | 8213fd95acb5 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/func | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1991 University of Cambridge | 
| 5 | ||
| 6 | Functions in Zermelo-Fraenkel Set Theory | |
| 7 | *) | |
| 8 | ||
| 9 | (*** The Pi operator -- dependent function space ***) | |
| 10 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 11 | Goalw [Pi_def] | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 12 | "f: Pi(A,B) <-> function(f) & f<=Sigma(A,B) & A<=domain(f)"; | 
| 2877 | 13 | by (Blast_tac 1); | 
| 760 | 14 | qed "Pi_iff"; | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 15 | |
| 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 16 | (*For upward compatibility with the former definition*) | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 17 | Goalw [Pi_def, function_def] | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 18 | "f: Pi(A,B) <-> f<=Sigma(A,B) & (ALL x:A. EX! y. <x,y>: f)"; | 
| 2877 | 19 | by (Blast_tac 1); | 
| 760 | 20 | qed "Pi_iff_old"; | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 21 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 22 | Goal "f: Pi(A,B) ==> function(f)"; | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 23 | by (asm_full_simp_tac (FOL_ss addsimps [Pi_iff]) 1); | 
| 760 | 24 | qed "fun_is_function"; | 
| 0 | 25 | |
| 26 | (**Two "destruct" rules for Pi **) | |
| 27 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 28 | Goalw [Pi_def] "f: Pi(A,B) ==> f <= Sigma(A,B)"; | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 29 | by (Blast_tac 1); | 
| 760 | 30 | qed "fun_is_rel"; | 
| 0 | 31 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 32 | Goal "[| f: Pi(A,B); a:A |] ==> EX! y. <a,y>: f"; | 
| 4091 | 33 | by (blast_tac (claset() addSDs [Pi_iff_old RS iffD1]) 1); | 
| 760 | 34 | qed "fun_unique_Pair"; | 
| 0 | 35 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 36 | val prems = Goalw [Pi_def] | 
| 0 | 37 | "[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> Pi(A,B) = Pi(A',B')"; | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 38 | by (simp_tac (FOL_ss addsimps prems addcongs [Sigma_cong]) 1); | 
| 760 | 39 | qed "Pi_cong"; | 
| 0 | 40 | |
| 2469 | 41 | (*Sigma_cong, Pi_cong NOT given to Addcongs: they cause | 
| 42 | flex-flex pairs and the "Check your prover" error. Most | |
| 43 | Sigmas and Pis are abbreviated as * or -> *) | |
| 44 | ||
| 485 | 45 | (*Weakening one function type to another; see also Pi_type*) | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 46 | Goalw [Pi_def] "[| f: A->B; B<=D |] ==> f: A->D"; | 
| 2469 | 47 | by (Best_tac 1); | 
| 760 | 48 | qed "fun_weaken_type"; | 
| 0 | 49 | |
| 50 | (*** Function Application ***) | |
| 51 | ||
| 5505 | 52 | Goalw [Pi_def, function_def] "[| <a,b>: f; <a,c>: f; f: Pi(A,B) |] ==> b=c"; | 
| 2877 | 53 | by (Blast_tac 1); | 
| 760 | 54 | qed "apply_equality2"; | 
| 0 | 55 | |
| 5505 | 56 | Goalw [apply_def, function_def] "[| <a,b>: f; function(f) |] ==> f`a = b"; | 
| 57 | by (Blast_tac 1); | |
| 4829 
aa5ea943f8e3
New proof of apply_equality and new thm Pi_image_cons
 paulson parents: 
4512diff
changeset | 58 | qed "function_apply_equality"; | 
| 
aa5ea943f8e3
New proof of apply_equality and new thm Pi_image_cons
 paulson parents: 
4512diff
changeset | 59 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 60 | Goalw [Pi_def] "[| <a,b>: f; f: Pi(A,B) |] ==> f`a = b"; | 
| 4829 
aa5ea943f8e3
New proof of apply_equality and new thm Pi_image_cons
 paulson parents: 
4512diff
changeset | 61 | by (blast_tac (claset() addIs [function_apply_equality]) 1); | 
| 760 | 62 | qed "apply_equality"; | 
| 0 | 63 | |
| 517 | 64 | (*Applying a function outside its domain yields 0*) | 
| 5505 | 65 | Goalw [apply_def] "a ~: domain(f) ==> f`a = 0"; | 
| 517 | 66 | by (rtac the_0 1); | 
| 2877 | 67 | by (Blast_tac 1); | 
| 760 | 68 | qed "apply_0"; | 
| 517 | 69 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 70 | Goal "[| f: Pi(A,B); c: f |] ==> EX x:A. c = <x,f`x>"; | 
| 7499 | 71 | by (ftac fun_is_rel 1); | 
| 4091 | 72 | by (blast_tac (claset() addDs [apply_equality]) 1); | 
| 760 | 73 | qed "Pi_memberD"; | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 74 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 75 | Goal "[| f: Pi(A,B); a:A |] ==> <a,f`a>: f"; | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 76 | by (rtac (fun_unique_Pair RS ex1E) 1); | 
| 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 77 | by (resolve_tac [apply_equality RS ssubst] 3); | 
| 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 78 | by (REPEAT (assume_tac 1)); | 
| 760 | 79 | qed "apply_Pair"; | 
| 0 | 80 | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 81 | (*Conclusion is flexible -- use res_inst_tac or else apply_funtype below!*) | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 82 | Goal "[| f: Pi(A,B); a:A |] ==> f`a : B(a)"; | 
| 0 | 83 | by (rtac (fun_is_rel RS subsetD RS SigmaE2) 1); | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 84 | by (REPEAT (ares_tac [apply_Pair] 1)); | 
| 760 | 85 | qed "apply_type"; | 
| 6153 | 86 | AddTCs [apply_type]; | 
| 0 | 87 | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 88 | (*This version is acceptable to the simplifier*) | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 89 | Goal "[| f: A->B; a:A |] ==> f`a : B"; | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 90 | by (REPEAT (ares_tac [apply_type] 1)); | 
| 760 | 91 | qed "apply_funtype"; | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 92 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 93 | Goal "f: Pi(A,B) ==> <a,b>: f <-> a:A & f`a = b"; | 
| 7499 | 94 | by (ftac fun_is_rel 1); | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 95 | by (blast_tac (claset() addSIs [apply_Pair, apply_equality]) 1); | 
| 760 | 96 | qed "apply_iff"; | 
| 0 | 97 | |
| 98 | (*Refining one Pi type to another*) | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 99 | val pi_prem::prems = Goal | 
| 0 | 100 | "[| f: Pi(A,C); !!x. x:A ==> f`x : B(x) |] ==> f : Pi(A,B)"; | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 101 | by (cut_facts_tac [pi_prem] 1); | 
| 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 102 | by (asm_full_simp_tac (FOL_ss addsimps [Pi_iff]) 1); | 
| 4091 | 103 | by (blast_tac (claset() addIs prems addSDs [pi_prem RS Pi_memberD]) 1); | 
| 760 | 104 | qed "Pi_type"; | 
| 0 | 105 | |
| 8127 
68c6159440f1
new lemmas for Ntree recursor example;  more simprules;  more lemmas borrowed
 paulson parents: 
7499diff
changeset | 106 | (*Such functions arise in non-standard datatypes, ZF/ex/Ntree for instance*) | 
| 
68c6159440f1
new lemmas for Ntree recursor example;  more simprules;  more lemmas borrowed
 paulson parents: 
7499diff
changeset | 107 | Goal "(f : Pi(A, %x. {y:B(x). P(x,y)})) \
 | 
| 
68c6159440f1
new lemmas for Ntree recursor example;  more simprules;  more lemmas borrowed
 paulson parents: 
7499diff
changeset | 108 | \ <-> f : Pi(A,B) & (ALL x: A. P(x, f`x))"; | 
| 
68c6159440f1
new lemmas for Ntree recursor example;  more simprules;  more lemmas borrowed
 paulson parents: 
7499diff
changeset | 109 | by (blast_tac (claset() addIs [Pi_type] addDs [apply_type]) 1); | 
| 
68c6159440f1
new lemmas for Ntree recursor example;  more simprules;  more lemmas borrowed
 paulson parents: 
7499diff
changeset | 110 | qed "Pi_Collect_iff"; | 
| 
68c6159440f1
new lemmas for Ntree recursor example;  more simprules;  more lemmas borrowed
 paulson parents: 
7499diff
changeset | 111 | |
| 9683 | 112 | val [major,minor] = Goal | 
| 113 | "[| f : Pi(A,B); !!x. x:A ==> B(x)<=C(x) |] ==> f : Pi(A,C)"; | |
| 114 | by (fast_tac (claset() addSIs [major RS Pi_type, minor RS subsetD, | |
| 115 | major RS apply_type]) 1); | |
| 116 | qed "Pi_weaken_type"; | |
| 117 | ||
| 0 | 118 | |
| 119 | (** Elimination of membership in a function **) | |
| 120 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 121 | Goal "[| <a,b> : f; f: Pi(A,B) |] ==> a : A"; | 
| 0 | 122 | by (REPEAT (ares_tac [fun_is_rel RS subsetD RS SigmaD1] 1)); | 
| 760 | 123 | qed "domain_type"; | 
| 0 | 124 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 125 | Goal "[| <a,b> : f; f: Pi(A,B) |] ==> b : B(a)"; | 
| 0 | 126 | by (etac (fun_is_rel RS subsetD RS SigmaD2) 1); | 
| 127 | by (assume_tac 1); | |
| 760 | 128 | qed "range_type"; | 
| 0 | 129 | |
| 9173 | 130 | Goal "[| <a,b>: f; f: Pi(A,B) |] ==> a:A & b:B(a) & f`a = b"; | 
| 131 | by (blast_tac (claset() addIs [domain_type, range_type, apply_equality]) 1); | |
| 132 | qed "Pair_mem_PiD"; | |
| 0 | 133 | |
| 134 | (*** Lambda Abstraction ***) | |
| 135 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 136 | Goalw [lam_def] "a:A ==> <a,b(a)> : (lam x:A. b(x))"; | 
| 0 | 137 | by (etac RepFunI 1); | 
| 760 | 138 | qed "lamI"; | 
| 0 | 139 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 140 | val major::prems = Goalw [lam_def] | 
| 0 | 141 | "[| p: (lam x:A. b(x)); !!x.[| x:A; p=<x,b(x)> |] ==> P \ | 
| 142 | \ |] ==> P"; | |
| 143 | by (rtac (major RS RepFunE) 1); | |
| 144 | by (REPEAT (ares_tac prems 1)); | |
| 760 | 145 | qed "lamE"; | 
| 0 | 146 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 147 | Goal "[| <a,c>: (lam x:A. b(x)) |] ==> c = b(a)"; | 
| 0 | 148 | by (REPEAT (eresolve_tac [asm_rl,lamE,Pair_inject,ssubst] 1)); | 
| 760 | 149 | qed "lamD"; | 
| 0 | 150 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 151 | val prems = Goalw [lam_def, Pi_def, function_def] | 
| 3840 | 152 | "[| !!x. x:A ==> b(x): B(x) |] ==> (lam x:A. b(x)) : Pi(A,B)"; | 
| 4091 | 153 | by (blast_tac (claset() addIs prems) 1); | 
| 760 | 154 | qed "lam_type"; | 
| 6153 | 155 | AddTCs [lam_type]; | 
| 0 | 156 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 157 | Goal "(lam x:A. b(x)) : A -> {b(x). x:A}";
 | 
| 0 | 158 | by (REPEAT (ares_tac [refl,lam_type,RepFunI] 1)); | 
| 760 | 159 | qed "lam_funtype"; | 
| 0 | 160 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 161 | Goal "a : A ==> (lam x:A. b(x)) ` a = b(a)"; | 
| 0 | 162 | by (REPEAT (ares_tac [apply_equality,lam_funtype,lamI] 1)); | 
| 760 | 163 | qed "beta"; | 
| 0 | 164 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 165 | Goalw [lam_def] "(lam x:0. b(x)) = 0"; | 
| 2469 | 166 | by (Simp_tac 1); | 
| 167 | qed "lam_empty"; | |
| 168 | ||
| 5156 | 169 | Goalw [lam_def] "domain(Lambda(A,b)) = A"; | 
| 170 | by (Blast_tac 1); | |
| 171 | qed "domain_lam"; | |
| 172 | ||
| 173 | Addsimps [beta, lam_empty, domain_lam]; | |
| 2469 | 174 | |
| 0 | 175 | (*congruence rule for lambda abstraction*) | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 176 | val prems = Goalw [lam_def] | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 177 | "[| A=A'; !!x. x:A' ==> b(x)=b'(x) |] ==> Lambda(A,b) = Lambda(A',b')"; | 
| 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 178 | by (simp_tac (FOL_ss addsimps prems addcongs [RepFun_cong]) 1); | 
| 760 | 179 | qed "lam_cong"; | 
| 0 | 180 | |
| 2469 | 181 | Addcongs [lam_cong]; | 
| 182 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 183 | val [major] = Goal | 
| 0 | 184 | "(!!x. x:A ==> EX! y. Q(x,y)) ==> EX f. ALL x:A. Q(x, f`x)"; | 
| 185 | by (res_inst_tac [("x", "lam x: A. THE y. Q(x,y)")] exI 1);
 | |
| 186 | by (rtac ballI 1); | |
| 2033 | 187 | by (stac beta 1); | 
| 0 | 188 | by (assume_tac 1); | 
| 189 | by (etac (major RS theI) 1); | |
| 760 | 190 | qed "lam_theI"; | 
| 0 | 191 | |
| 9683 | 192 | Goal "[| (lam x:A. f(x)) = (lam x:A. g(x)); a:A |] ==> f(a)=g(a)"; | 
| 193 | by (fast_tac (claset() addSIs [lamI] addEs [equalityE, lamE]) 1); | |
| 194 | qed "lam_eqE"; | |
| 195 | ||
| 196 | ||
| 197 | (*Empty function spaces*) | |
| 198 | Goalw [Pi_def, function_def] "Pi(0,A) = {0}";
 | |
| 199 | by (Blast_tac 1); | |
| 200 | qed "Pi_empty1"; | |
| 201 | Addsimps [Pi_empty1]; | |
| 202 | ||
| 203 | (*The singleton function*) | |
| 204 | Goalw [Pi_def, function_def] "{<a,b>} : {a} -> {b}";
 | |
| 205 | by (Blast_tac 1); | |
| 206 | qed "singleton_fun"; | |
| 207 | Addsimps [singleton_fun]; | |
| 208 | ||
| 209 | Goalw [Pi_def, function_def] "(A->0) = (if A=0 then {0} else 0)";
 | |
| 210 | by (Force_tac 1); | |
| 211 | qed "Pi_empty2"; | |
| 212 | Addsimps [Pi_empty2]; | |
| 213 | ||
| 214 | Goal "(A->X)=0 \\<longleftrightarrow> X=0 & (A \\<noteq> 0)"; | |
| 215 | by Auto_tac; | |
| 216 | by (fast_tac (claset() addSIs [equals0I] addIs [lam_type]) 1); | |
| 217 | qed "fun_space_empty_iff"; | |
| 218 | AddIffs [fun_space_empty_iff]; | |
| 219 | ||
| 0 | 220 | |
| 221 | (** Extensionality **) | |
| 222 | ||
| 223 | (*Semi-extensionality!*) | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 224 | val prems = Goal | 
| 0 | 225 | "[| f : Pi(A,B); g: Pi(C,D); A<=C; \ | 
| 226 | \ !!x. x:A ==> f`x = g`x |] ==> f<=g"; | |
| 227 | by (rtac subsetI 1); | |
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 228 | by (eresolve_tac (prems RL [Pi_memberD RS bexE]) 1); | 
| 0 | 229 | by (etac ssubst 1); | 
| 230 | by (resolve_tac (prems RL [ssubst]) 1); | |
| 231 | by (REPEAT (ares_tac (prems@[apply_Pair,subsetD]) 1)); | |
| 760 | 232 | qed "fun_subset"; | 
| 0 | 233 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 234 | val prems = Goal | 
| 0 | 235 | "[| f : Pi(A,B); g: Pi(A,D); \ | 
| 236 | \ !!x. x:A ==> f`x = g`x |] ==> f=g"; | |
| 237 | by (REPEAT (ares_tac (prems @ (prems RL [sym]) @ | |
| 1461 | 238 | [subset_refl,equalityI,fun_subset]) 1)); | 
| 760 | 239 | qed "fun_extension"; | 
| 0 | 240 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 241 | Goal "f : Pi(A,B) ==> (lam x:A. f`x) = f"; | 
| 0 | 242 | by (rtac fun_extension 1); | 
| 243 | by (REPEAT (ares_tac [lam_type,apply_type,beta] 1)); | |
| 760 | 244 | qed "eta"; | 
| 0 | 245 | |
| 2469 | 246 | Addsimps [eta]; | 
| 247 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 248 | Goal "[| f:Pi(A,B); g:Pi(A,C) |] ==> (ALL a:A. f`a = g`a) <-> f=g"; | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 249 | by (blast_tac (claset() addIs [fun_extension]) 1); | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 250 | qed "fun_extension_iff"; | 
| 5202 | 251 | |
| 252 | (*thanks to Mark Staples*) | |
| 9211 | 253 | Goal "[| f:Pi(A,B); g:Pi(A,C) |] ==> f <= g <-> (f = g)"; | 
| 254 | by (rtac iffI 1); | |
| 255 | by (asm_full_simp_tac ZF_ss 2); | |
| 256 | by (rtac fun_extension 1); | |
| 257 | by (REPEAT (atac 1)); | |
| 258 | by (dtac (apply_Pair RSN (2,subsetD)) 1); | |
| 259 | by (REPEAT (atac 1)); | |
| 260 | by (rtac (apply_equality RS sym) 1); | |
| 261 | by (REPEAT (atac 1)) ; | |
| 262 | qed "fun_subset_eq"; | |
| 5202 | 263 | |
| 264 | ||
| 0 | 265 | (*Every element of Pi(A,B) may be expressed as a lambda abstraction!*) | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 266 | val prems = Goal | 
| 0 | 267 | "[| f: Pi(A,B); \ | 
| 3840 | 268 | \ !!b. [| ALL x:A. b(x):B(x); f = (lam x:A. b(x)) |] ==> P \ | 
| 0 | 269 | \ |] ==> P"; | 
| 270 | by (resolve_tac prems 1); | |
| 271 | by (rtac (eta RS sym) 2); | |
| 272 | by (REPEAT (ares_tac (prems@[ballI,apply_type]) 1)); | |
| 760 | 273 | qed "Pi_lamE"; | 
| 0 | 274 | |
| 275 | ||
| 435 | 276 | (** Images of functions **) | 
| 277 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 278 | Goalw [lam_def] "C <= A ==> (lam x:A. b(x)) `` C = {b(x). x:C}";
 | 
| 2877 | 279 | by (Blast_tac 1); | 
| 760 | 280 | qed "image_lam"; | 
| 435 | 281 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 282 | Goal "[| f : Pi(A,B);  C <= A |] ==> f``C = {f`x. x:C}";
 | 
| 437 | 283 | by (etac (eta RS subst) 1); | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 284 | by (asm_full_simp_tac (simpset() addsimps [image_lam, subset_iff]) 1); | 
| 760 | 285 | qed "image_fun"; | 
| 435 | 286 | |
| 5137 | 287 | Goal "[| f: Pi(A,B); x: A |] ==> f `` cons(x,y) = cons(f`x, f``y)"; | 
| 4829 
aa5ea943f8e3
New proof of apply_equality and new thm Pi_image_cons
 paulson parents: 
4512diff
changeset | 288 | by (blast_tac (claset() addDs [apply_equality, apply_Pair]) 1); | 
| 
aa5ea943f8e3
New proof of apply_equality and new thm Pi_image_cons
 paulson parents: 
4512diff
changeset | 289 | qed "Pi_image_cons"; | 
| 
aa5ea943f8e3
New proof of apply_equality and new thm Pi_image_cons
 paulson parents: 
4512diff
changeset | 290 | |
| 435 | 291 | |
| 0 | 292 | (*** properties of "restrict" ***) | 
| 293 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 294 | Goalw [restrict_def,lam_def] | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 295 | "[| f: Pi(C,B); A<=C |] ==> restrict(f,A) <= f"; | 
| 4091 | 296 | by (blast_tac (claset() addIs [apply_Pair]) 1); | 
| 760 | 297 | qed "restrict_subset"; | 
| 0 | 298 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 299 | val prems = Goalw [restrict_def] | 
| 0 | 300 | "[| !!x. x:A ==> f`x: B(x) |] ==> restrict(f,A) : Pi(A,B)"; | 
| 301 | by (rtac lam_type 1); | |
| 302 | by (eresolve_tac prems 1); | |
| 760 | 303 | qed "restrict_type"; | 
| 0 | 304 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 305 | Goal "[| f: Pi(C,B); A<=C |] ==> restrict(f,A) : Pi(A,B)"; | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 306 | by (blast_tac (claset() addIs [apply_type, restrict_type]) 1); | 
| 760 | 307 | qed "restrict_type2"; | 
| 0 | 308 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 309 | Goalw [restrict_def] "a : A ==> restrict(f,A) ` a = f`a"; | 
| 0 | 310 | by (etac beta 1); | 
| 760 | 311 | qed "restrict"; | 
| 0 | 312 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 313 | Goalw [restrict_def] "restrict(f,0) = 0"; | 
| 2469 | 314 | by (Simp_tac 1); | 
| 315 | qed "restrict_empty"; | |
| 316 | ||
| 317 | Addsimps [restrict, restrict_empty]; | |
| 318 | ||
| 0 | 319 | (*NOT SAFE as a congruence rule for the simplifier! Can cause it to fail!*) | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 320 | val prems = Goalw [restrict_def] | 
| 0 | 321 | "[| A=B; !!x. x:B ==> f`x=g`x |] ==> restrict(f,A) = restrict(g,B)"; | 
| 322 | by (REPEAT (ares_tac (prems@[lam_cong]) 1)); | |
| 760 | 323 | qed "restrict_eqI"; | 
| 0 | 324 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 325 | Goalw [restrict_def, lam_def] "domain(restrict(f,C)) = C"; | 
| 2877 | 326 | by (Blast_tac 1); | 
| 760 | 327 | qed "domain_restrict"; | 
| 8182 | 328 | Addsimps [domain_restrict]; | 
| 0 | 329 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 330 | Goalw [restrict_def] | 
| 0 | 331 | "A<=C ==> restrict(lam x:C. b(x), A) = (lam x:A. b(x))"; | 
| 332 | by (rtac (refl RS lam_cong) 1); | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 333 | by (set_mp_tac 1); | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 334 | by (Asm_simp_tac 1); | 
| 760 | 335 | qed "restrict_lam_eq"; | 
| 0 | 336 | |
| 8182 | 337 | Goal "f : cons(a, b) -> B ==> f = cons(<a, f ` a>, restrict(f, b))"; | 
| 338 | by (rtac equalityI 1); | |
| 339 | by (blast_tac (claset() addIs [apply_Pair, impOfSubs restrict_subset]) 2); | |
| 340 | by (auto_tac (claset() addSDs [Pi_memberD], | |
| 341 | simpset() addsimps [restrict_def, lam_def])); | |
| 342 | qed "fun_cons_restrict_eq"; | |
| 0 | 343 | |
| 344 | ||
| 345 | (*** Unions of functions ***) | |
| 346 | ||
| 347 | (** The Union of a set of COMPATIBLE functions is a function **) | |
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 348 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 349 | Goalw [function_def] | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 350 | "[| ALL x:S. function(x); \ | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 351 | \ ALL x:S. ALL y:S. x<=y | y<=x |] ==> \ | 
| 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 352 | \ function(Union(S))"; | 
| 2925 | 353 | by (fast_tac (ZF_cs addSDs [bspec RS bspec]) 1); | 
| 354 | (*by (Blast_tac 1); takes too long...*) | |
| 760 | 355 | qed "function_Union"; | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 356 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 357 | Goalw [Pi_def] | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 358 | "[| ALL f:S. EX C D. f:C->D; \ | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 359 | \ ALL f:S. ALL y:S. f<=y | y<=f |] ==> \ | 
| 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 360 | \ Union(S) : domain(Union(S)) -> range(Union(S))"; | 
| 2877 | 361 | by (blast_tac (subset_cs addSIs [rel_Union, function_Union]) 1); | 
| 760 | 362 | qed "fun_Union"; | 
| 0 | 363 | |
| 364 | ||
| 365 | (** The Union of 2 disjoint functions is a function **) | |
| 366 | ||
| 9907 | 367 | bind_thms ("Un_rls", [Un_subset_iff, SUM_Un_distrib1, prod_Un_distrib2, 
 | 
| 1461 | 368 | Un_upper1 RSN (2, subset_trans), | 
| 9907 | 369 | Un_upper2 RSN (2, subset_trans)]); | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 370 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 371 | Goal "[| f: A->B; g: C->D; A Int C = 0 |] \ | 
| 2925 | 372 | \ ==> (f Un g) : (A Un C) -> (B Un D)"; | 
| 2877 | 373 | (*Prove the product and domain subgoals using distributive laws*) | 
| 4091 | 374 | by (asm_full_simp_tac (simpset() addsimps [Pi_iff,extension]@Un_rls) 1); | 
| 1461 | 375 | by (rewtac function_def); | 
| 4512 
572440df6aa7
Changed required by new blast_tac (the one that squashes flex-flex pairs)
 paulson parents: 
4091diff
changeset | 376 | by (Blast_tac 1); | 
| 760 | 377 | qed "fun_disjoint_Un"; | 
| 0 | 378 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 379 | Goal "[| a:A; f: A->B; g: C->D; A Int C = 0 |] ==> \ | 
| 0 | 380 | \ (f Un g)`a = f`a"; | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 381 | by (rtac (apply_Pair RS UnI1 RS apply_equality) 1); | 
| 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 382 | by (REPEAT (ares_tac [fun_disjoint_Un] 1)); | 
| 760 | 383 | qed "fun_disjoint_apply1"; | 
| 0 | 384 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 385 | Goal "[| c:C; f: A->B; g: C->D; A Int C = 0 |] ==> \ | 
| 0 | 386 | \ (f Un g)`c = g`c"; | 
| 691 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 387 | by (rtac (apply_Pair RS UnI2 RS apply_equality) 1); | 
| 
b9fc536792bb
ZF/func: tidied many proofs, using new definition of Pi(A,B)
 lcp parents: 
538diff
changeset | 388 | by (REPEAT (ares_tac [fun_disjoint_Un] 1)); | 
| 760 | 389 | qed "fun_disjoint_apply2"; | 
| 0 | 390 | |
| 391 | (** Domain and range of a function/relation **) | |
| 392 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 393 | Goalw [Pi_def] "f : Pi(A,B) ==> domain(f)=A"; | 
| 2877 | 394 | by (Blast_tac 1); | 
| 760 | 395 | qed "domain_of_fun"; | 
| 0 | 396 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 397 | Goal "[| f : Pi(A,B); a: A |] ==> f`a : range(f)"; | 
| 517 | 398 | by (etac (apply_Pair RS rangeI) 1); | 
| 399 | by (assume_tac 1); | |
| 760 | 400 | qed "apply_rangeI"; | 
| 517 | 401 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 402 | Goal "f : Pi(A,B) ==> f : A->range(f)"; | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5321diff
changeset | 403 | by (REPEAT (ares_tac [Pi_type, apply_rangeI] 1)); | 
| 760 | 404 | qed "range_of_fun"; | 
| 0 | 405 | |
| 406 | (*** Extensions of functions ***) | |
| 407 | ||
| 5321 | 408 | Goal "[| f: A->B; c~:A |] ==> cons(<c,b>,f) : cons(c,A) -> cons(b,B)"; | 
| 857 | 409 | by (forward_tac [singleton_fun RS fun_disjoint_Un] 1); | 
| 519 | 410 | by (asm_full_simp_tac (FOL_ss addsimps [cons_eq]) 2); | 
| 2877 | 411 | by (Blast_tac 1); | 
| 760 | 412 | qed "fun_extend"; | 
| 0 | 413 | |
| 5321 | 414 | Goal "[| f: A->B; c~:A; b: B |] ==> cons(<c,b>,f) : cons(c,A) -> B"; | 
| 4091 | 415 | by (blast_tac (claset() addIs [fun_extend RS fun_weaken_type]) 1); | 
| 760 | 416 | qed "fun_extend3"; | 
| 538 | 417 | |
| 5321 | 418 | Goal "[| f: A->B; a:A; c~:A |] ==> cons(<c,b>,f)`a = f`a"; | 
| 0 | 419 | by (rtac (apply_Pair RS consI2 RS apply_equality) 1); | 
| 420 | by (rtac fun_extend 3); | |
| 421 | by (REPEAT (assume_tac 1)); | |
| 760 | 422 | qed "fun_extend_apply1"; | 
| 0 | 423 | |
| 5321 | 424 | Goal "[| f: A->B; c~:A |] ==> cons(<c,b>,f)`c = b"; | 
| 0 | 425 | by (rtac (consI1 RS apply_equality) 1); | 
| 426 | by (rtac fun_extend 1); | |
| 427 | by (REPEAT (assume_tac 1)); | |
| 760 | 428 | qed "fun_extend_apply2"; | 
| 0 | 429 | |
| 2469 | 430 | bind_thm ("singleton_apply", [singletonI, singleton_fun] MRS apply_equality);
 | 
| 431 | Addsimps [singleton_apply]; | |
| 432 | ||
| 538 | 433 | (*For Finite.ML. Inclusion of right into left is easy*) | 
| 5321 | 434 | Goal "c ~: A ==> cons(c,A) -> B = (UN f: A->B. UN b:B. {cons(<c,b>, f)})";
 | 
| 737 
436019ca97d7
cons_fun_eq: modified strange uses of classical reasoner
 lcp parents: 
691diff
changeset | 435 | by (rtac equalityI 1); | 
| 4091 | 436 | by (safe_tac (claset() addSEs [fun_extend3])); | 
| 485 | 437 | (*Inclusion of left into right*) | 
| 438 | by (subgoal_tac "restrict(x, A) : A -> B" 1); | |
| 4091 | 439 | by (blast_tac (claset() addIs [restrict_type2]) 2); | 
| 485 | 440 | by (rtac UN_I 1 THEN assume_tac 1); | 
| 737 
436019ca97d7
cons_fun_eq: modified strange uses of classical reasoner
 lcp parents: 
691diff
changeset | 441 | by (rtac (apply_funtype RS UN_I) 1 THEN REPEAT (ares_tac [consI1] 1)); | 
| 2469 | 442 | by (Simp_tac 1); | 
| 1461 | 443 | by (rtac fun_extension 1 THEN REPEAT (ares_tac [fun_extend] 1)); | 
| 485 | 444 | by (etac consE 1); | 
| 445 | by (ALLGOALS | |
| 4091 | 446 | (asm_simp_tac (simpset() addsimps [restrict, fun_extend_apply1, | 
| 2493 | 447 | fun_extend_apply2]))); | 
| 760 | 448 | qed "cons_fun_eq"; | 
| 485 | 449 |