| author | wenzelm | 
| Tue, 05 Oct 1999 18:26:34 +0200 | |
| changeset 7742 | 01386eb4eab0 | 
| parent 5431 | d50c2783f941 | 
| child 8266 | 4bc79ed1233b | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/ex/misc | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 6 | Miscellaneous examples for Zermelo-Fraenkel Set Theory | |
| 1646 | 7 | Composition of homomorphisms, Pastre's examples, ... | 
| 0 | 8 | *) | 
| 9 | ||
| 10 | writeln"ZF/ex/misc"; | |
| 11 | ||
| 5431 | 12 | (*trivial example of term synthesis: apparently hard for some provers!*) | 
| 13 | Goal "a ~= b ==> a:?X & b ~: ?X"; | |
| 14 | by (Blast_tac 1); | |
| 15 | result(); | |
| 16 | ||
| 3000 
7ecb8b6a3d7f
Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
 paulson parents: 
2496diff
changeset | 17 | (*Nice Blast_tac benchmark. Proved in 0.3s; old tactics can't manage it!*) | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5068diff
changeset | 18 | Goal "ALL x:S. ALL y:S. x<=y ==> EX z. S <= {z}";
 | 
| 3000 
7ecb8b6a3d7f
Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
 paulson parents: 
2496diff
changeset | 19 | by (Blast_tac 1); | 
| 
7ecb8b6a3d7f
Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
 paulson parents: 
2496diff
changeset | 20 | result(); | 
| 
7ecb8b6a3d7f
Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
 paulson parents: 
2496diff
changeset | 21 | |
| 4322 | 22 | (*variant of the benchmark above*) | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5068diff
changeset | 23 | Goal "ALL x:S. Union(S) <= x ==> EX z. S <= {z}";
 | 
| 4322 | 24 | by (Blast_tac 1); | 
| 25 | result(); | |
| 26 | ||
| 4109 | 27 | context Perm.thy; | 
| 0 | 28 | |
| 29 | (*Example 12 (credited to Peter Andrews) from | |
| 30 | W. Bledsoe. A Maximal Method for Set Variables in Automatic Theorem-proving. | |
| 31 | In: J. Hayes and D. Michie and L. Mikulich, eds. Machine Intelligence 9. | |
| 32 | Ellis Horwood, 53-100 (1979). *) | |
| 5068 | 33 | Goal "(ALL F. {x}: F --> {y}:F) --> (ALL A. x:A --> y:A)";
 | 
| 2469 | 34 | by (Best_tac 1); | 
| 0 | 35 | result(); | 
| 36 | ||
| 37 | ||
| 38 | (*** Composition of homomorphisms is a homomorphism ***) | |
| 39 | ||
| 40 | (*Given as a challenge problem in | |
| 41 | R. Boyer et al., | |
| 42 | Set Theory in First-Order Logic: Clauses for G\"odel's Axioms, | |
| 43 | JAR 2 (1986), 287-327 | |
| 44 | *) | |
| 45 | ||
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 46 | (*collecting the relevant lemmas*) | 
| 2469 | 47 | Addsimps [comp_fun, SigmaI, apply_funtype]; | 
| 0 | 48 | |
| 736 | 49 | (*This version uses a super application of simp_tac. Needs setloop to help | 
| 50 | proving conditions of rewrites such as comp_fun_apply; | |
| 51 | rewriting does not instantiate Vars*) | |
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 52 | goal Perm.thy | 
| 0 | 53 | "(ALL A f B g. hom(A,f,B,g) = \ | 
| 54 | \          {H: A->B. f:A*A->A & g:B*B->B & \
 | |
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 55 | \ (ALL x:A. ALL y:A. H`(f`<x,y>) = g`<H`x,H`y>)}) --> \ | 
| 0 | 56 | \ J : hom(A,f,B,g) & K : hom(B,g,C,h) --> \ | 
| 57 | \ (K O J) : hom(A,f,C,h)"; | |
| 4152 | 58 | by (asm_simp_tac (simpset() setloop (K Safe_tac)) 1); | 
| 0 | 59 | val comp_homs = result(); | 
| 60 | ||
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 61 | (*This version uses meta-level rewriting, safe_tac and asm_simp_tac*) | 
| 0 | 62 | val [hom_def] = goal Perm.thy | 
| 63 | "(!! A f B g. hom(A,f,B,g) == \ | |
| 64 | \          {H: A->B. f:A*A->A & g:B*B->B & \
 | |
| 65 | \ (ALL x:A. ALL y:A. H`(f`<x,y>) = g`<H`x,H`y>)}) ==> \ | |
| 66 | \ J : hom(A,f,B,g) & K : hom(B,g,C,h) --> \ | |
| 67 | \ (K O J) : hom(A,f,C,h)"; | |
| 68 | by (rewtac hom_def); | |
| 4152 | 69 | by Safe_tac; | 
| 2469 | 70 | by (Asm_simp_tac 1); | 
| 71 | by (Asm_simp_tac 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 72 | qed "comp_homs"; | 
| 0 | 73 | |
| 74 | ||
| 75 | (** A characterization of functions, suggested by Tobias Nipkow **) | |
| 76 | ||
| 5068 | 77 | Goalw [Pi_def, function_def] | 
| 0 | 78 | "r: domain(r)->B <-> r <= domain(r)*B & (ALL X. r `` (r -`` X) <= X)"; | 
| 2469 | 79 | by (Best_tac 1); | 
| 0 | 80 | result(); | 
| 81 | ||
| 82 | ||
| 83 | (**** From D Pastre. Automatic theorem proving in set theory. | |
| 84 | Artificial Intelligence, 10:1--27, 1978. | |
| 85 | These examples require forward reasoning! ****) | |
| 86 | ||
| 87 | (*reduce the clauses to units by type checking -- beware of nontermination*) | |
| 88 | fun forw_typechk tyrls [] = [] | |
| 89 | | forw_typechk tyrls clauses = | |
| 90 | let val (units, others) = partition (has_fewer_prems 1) clauses | |
| 91 | in gen_union eq_thm (units, forw_typechk tyrls (tyrls RL others)) | |
| 92 | end; | |
| 93 | ||
| 94 | (*A crude form of forward reasoning*) | |
| 95 | fun forw_iterate tyrls rls facts 0 = facts | |
| 96 | | forw_iterate tyrls rls facts n = | |
| 97 | let val facts' = | |
| 1461 | 98 | gen_union eq_thm (forw_typechk (tyrls@facts) (facts RL rls), facts); | 
| 0 | 99 | in forw_iterate tyrls rls facts' (n-1) end; | 
| 100 | ||
| 101 | val pastre_rls = | |
| 102 | [comp_mem_injD1, comp_mem_surjD1, comp_mem_injD2, comp_mem_surjD2]; | |
| 103 | ||
| 104 | fun pastre_facts (fact1::fact2::fact3::prems) = | |
| 434 | 105 | forw_iterate (prems @ [comp_surj, comp_inj, comp_fun]) | 
| 0 | 106 | pastre_rls [fact1,fact2,fact3] 4; | 
| 107 | ||
| 108 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 109 | "[| (h O g O f): inj(A,A); \ | 
| 110 | \ (f O h O g): surj(B,B); \ | |
| 111 | \ (g O f O h): surj(C,C); \ | |
| 0 | 112 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 113 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 114 | qed "pastre1"; | 
| 0 | 115 | |
| 116 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 117 | "[| (h O g O f): surj(A,A); \ | 
| 118 | \ (f O h O g): inj(B,B); \ | |
| 119 | \ (g O f O h): surj(C,C); \ | |
| 0 | 120 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 121 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 122 | qed "pastre2"; | 
| 0 | 123 | |
| 124 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 125 | "[| (h O g O f): surj(A,A); \ | 
| 126 | \ (f O h O g): surj(B,B); \ | |
| 127 | \ (g O f O h): inj(C,C); \ | |
| 0 | 128 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 129 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 130 | qed "pastre3"; | 
| 0 | 131 | |
| 132 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 133 | "[| (h O g O f): surj(A,A); \ | 
| 134 | \ (f O h O g): inj(B,B); \ | |
| 135 | \ (g O f O h): inj(C,C); \ | |
| 0 | 136 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 137 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 138 | qed "pastre4"; | 
| 0 | 139 | |
| 140 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 141 | "[| (h O g O f): inj(A,A); \ | 
| 142 | \ (f O h O g): surj(B,B); \ | |
| 143 | \ (g O f O h): inj(C,C); \ | |
| 0 | 144 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 145 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 146 | qed "pastre5"; | 
| 0 | 147 | |
| 148 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 149 | "[| (h O g O f): inj(A,A); \ | 
| 150 | \ (f O h O g): inj(B,B); \ | |
| 151 | \ (g O f O h): surj(C,C); \ | |
| 0 | 152 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 153 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 154 | qed "pastre6"; | 
| 0 | 155 | |
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 156 | (** Yet another example... **) | 
| 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 157 | |
| 2469 | 158 | goal Perm.thy | 
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 159 |     "(lam Z:Pow(A+B). <{x:A. Inl(x):Z}, {y:B. Inr(y):Z}>) \
 | 
| 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 160 | \ : bij(Pow(A+B), Pow(A)*Pow(B))"; | 
| 1110 
756aa2e81f6e
Changed some definitions and proofs to use pattern-matching.
 lcp parents: 
782diff
changeset | 161 | by (res_inst_tac [("d", "%<X,Y>.{Inl(x).x:X} Un {Inr(y).y:Y}")] 
 | 
| 695 
a1586fa1b755
ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
 lcp parents: 
434diff
changeset | 162 | lam_bijective 1); | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4322diff
changeset | 163 | (*Auto_tac no longer proves it*) | 
| 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4322diff
changeset | 164 | by (REPEAT (fast_tac (claset() addss (simpset())) 1)); | 
| 4595 | 165 | qed "Pow_sum_bij"; | 
| 166 | ||
| 167 | (*As a special case, we have bij(Pow(A*B), A -> Pow B) *) | |
| 168 | goal Perm.thy | |
| 169 |     "(lam r:Pow(Sigma(A,B)). lam x:A. r``{x}) \
 | |
| 170 | \ : bij(Pow(Sigma(A,B)), PROD x:A. Pow(B(x)))"; | |
| 171 | by (res_inst_tac [("d", "%f. UN x:A. UN y:f`x. {<x,y>}")] lam_bijective 1);
 | |
| 172 | by (blast_tac (claset() addDs [apply_type]) 2); | |
| 173 | by (blast_tac (claset() addIs [lam_type]) 1); | |
| 174 | by (ALLGOALS Asm_simp_tac); | |
| 175 | by (Fast_tac 1); | |
| 176 | by (rtac fun_extension 1); | |
| 177 | by (assume_tac 2); | |
| 178 | by (rtac (singletonI RS lam_type) 1); | |
| 179 | by (Asm_simp_tac 1); | |
| 180 | by (Blast_tac 1); | |
| 181 | qed "Pow_Sigma_bij"; | |
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 182 | |
| 0 | 183 | writeln"Reached end of file."; |