| author | wenzelm | 
| Fri, 02 Jan 1998 18:47:31 +0100 | |
| changeset 4514 | 78eda600f35d | 
| parent 4477 | b3e5857d8d99 | 
| child 4595 | fa8cee619732 | 
| 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 | ||
| 3000 
7ecb8b6a3d7f
Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
 paulson parents: 
2496diff
changeset | 12 | (*Nice Blast_tac benchmark. Proved in 0.3s; old tactics can't manage it!*) | 
| 
7ecb8b6a3d7f
Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
 paulson parents: 
2496diff
changeset | 13 | goal ZF.thy "!!S. ALL x:S. ALL y:S. x<=y ==> EX z. S <= {z}";
 | 
| 
7ecb8b6a3d7f
Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
 paulson parents: 
2496diff
changeset | 14 | by (Blast_tac 1); | 
| 
7ecb8b6a3d7f
Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
 paulson parents: 
2496diff
changeset | 15 | result(); | 
| 
7ecb8b6a3d7f
Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
 paulson parents: 
2496diff
changeset | 16 | |
| 4322 | 17 | (*variant of the benchmark above*) | 
| 18 | goal ZF.thy "!!S. ALL x:S. Union(S) <= x ==> EX z. S <= {z}";
 | |
| 19 | by (Blast_tac 1); | |
| 20 | result(); | |
| 21 | ||
| 4109 | 22 | context Perm.thy; | 
| 0 | 23 | |
| 24 | (*Example 12 (credited to Peter Andrews) from | |
| 25 | W. Bledsoe. A Maximal Method for Set Variables in Automatic Theorem-proving. | |
| 26 | In: J. Hayes and D. Michie and L. Mikulich, eds. Machine Intelligence 9. | |
| 27 | Ellis Horwood, 53-100 (1979). *) | |
| 2469 | 28 | goal thy "(ALL F. {x}: F --> {y}:F) --> (ALL A. x:A --> y:A)";
 | 
| 29 | by (Best_tac 1); | |
| 0 | 30 | result(); | 
| 31 | ||
| 32 | ||
| 33 | (*** Composition of homomorphisms is a homomorphism ***) | |
| 34 | ||
| 35 | (*Given as a challenge problem in | |
| 36 | R. Boyer et al., | |
| 37 | Set Theory in First-Order Logic: Clauses for G\"odel's Axioms, | |
| 38 | JAR 2 (1986), 287-327 | |
| 39 | *) | |
| 40 | ||
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 41 | (*collecting the relevant lemmas*) | 
| 2469 | 42 | Addsimps [comp_fun, SigmaI, apply_funtype]; | 
| 0 | 43 | |
| 736 | 44 | (*This version uses a super application of simp_tac. Needs setloop to help | 
| 45 | proving conditions of rewrites such as comp_fun_apply; | |
| 46 | 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 | 47 | goal Perm.thy | 
| 0 | 48 | "(ALL A f B g. hom(A,f,B,g) = \ | 
| 49 | \          {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 | 50 | \ (ALL x:A. ALL y:A. H`(f`<x,y>) = g`<H`x,H`y>)}) --> \ | 
| 0 | 51 | \ J : hom(A,f,B,g) & K : hom(B,g,C,h) --> \ | 
| 52 | \ (K O J) : hom(A,f,C,h)"; | |
| 4152 | 53 | by (asm_simp_tac (simpset() setloop (K Safe_tac)) 1); | 
| 0 | 54 | val comp_homs = result(); | 
| 55 | ||
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 56 | (*This version uses meta-level rewriting, safe_tac and asm_simp_tac*) | 
| 0 | 57 | val [hom_def] = goal Perm.thy | 
| 58 | "(!! A f B g. hom(A,f,B,g) == \ | |
| 59 | \          {H: A->B. f:A*A->A & g:B*B->B & \
 | |
| 60 | \ (ALL x:A. ALL y:A. H`(f`<x,y>) = g`<H`x,H`y>)}) ==> \ | |
| 61 | \ J : hom(A,f,B,g) & K : hom(B,g,C,h) --> \ | |
| 62 | \ (K O J) : hom(A,f,C,h)"; | |
| 63 | by (rewtac hom_def); | |
| 4152 | 64 | by Safe_tac; | 
| 2469 | 65 | by (Asm_simp_tac 1); | 
| 66 | by (Asm_simp_tac 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 67 | qed "comp_homs"; | 
| 0 | 68 | |
| 69 | ||
| 70 | (** A characterization of functions, suggested by Tobias Nipkow **) | |
| 71 | ||
| 2469 | 72 | goalw thy [Pi_def, function_def] | 
| 0 | 73 | "r: domain(r)->B <-> r <= domain(r)*B & (ALL X. r `` (r -`` X) <= X)"; | 
| 2469 | 74 | by (Best_tac 1); | 
| 0 | 75 | result(); | 
| 76 | ||
| 77 | ||
| 78 | (**** From D Pastre. Automatic theorem proving in set theory. | |
| 79 | Artificial Intelligence, 10:1--27, 1978. | |
| 80 | These examples require forward reasoning! ****) | |
| 81 | ||
| 82 | (*reduce the clauses to units by type checking -- beware of nontermination*) | |
| 83 | fun forw_typechk tyrls [] = [] | |
| 84 | | forw_typechk tyrls clauses = | |
| 85 | let val (units, others) = partition (has_fewer_prems 1) clauses | |
| 86 | in gen_union eq_thm (units, forw_typechk tyrls (tyrls RL others)) | |
| 87 | end; | |
| 88 | ||
| 89 | (*A crude form of forward reasoning*) | |
| 90 | fun forw_iterate tyrls rls facts 0 = facts | |
| 91 | | forw_iterate tyrls rls facts n = | |
| 92 | let val facts' = | |
| 1461 | 93 | gen_union eq_thm (forw_typechk (tyrls@facts) (facts RL rls), facts); | 
| 0 | 94 | in forw_iterate tyrls rls facts' (n-1) end; | 
| 95 | ||
| 96 | val pastre_rls = | |
| 97 | [comp_mem_injD1, comp_mem_surjD1, comp_mem_injD2, comp_mem_surjD2]; | |
| 98 | ||
| 99 | fun pastre_facts (fact1::fact2::fact3::prems) = | |
| 434 | 100 | forw_iterate (prems @ [comp_surj, comp_inj, comp_fun]) | 
| 0 | 101 | pastre_rls [fact1,fact2,fact3] 4; | 
| 102 | ||
| 103 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 104 | "[| (h O g O f): inj(A,A); \ | 
| 105 | \ (f O h O g): surj(B,B); \ | |
| 106 | \ (g O f O h): surj(C,C); \ | |
| 0 | 107 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 108 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 109 | qed "pastre1"; | 
| 0 | 110 | |
| 111 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 112 | "[| (h O g O f): surj(A,A); \ | 
| 113 | \ (f O h O g): inj(B,B); \ | |
| 114 | \ (g O f O h): surj(C,C); \ | |
| 0 | 115 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 116 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 117 | qed "pastre2"; | 
| 0 | 118 | |
| 119 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 120 | "[| (h O g O f): surj(A,A); \ | 
| 121 | \ (f O h O g): surj(B,B); \ | |
| 122 | \ (g O f O h): inj(C,C); \ | |
| 0 | 123 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 124 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 125 | qed "pastre3"; | 
| 0 | 126 | |
| 127 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 128 | "[| (h O g O f): surj(A,A); \ | 
| 129 | \ (f O h O g): inj(B,B); \ | |
| 130 | \ (g O f O h): inj(C,C); \ | |
| 0 | 131 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 132 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 133 | qed "pastre4"; | 
| 0 | 134 | |
| 135 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 136 | "[| (h O g O f): inj(A,A); \ | 
| 137 | \ (f O h O g): surj(B,B); \ | |
| 138 | \ (g O f O h): inj(C,C); \ | |
| 0 | 139 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 140 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 141 | qed "pastre5"; | 
| 0 | 142 | |
| 143 | val prems = goalw Perm.thy [bij_def] | |
| 1461 | 144 | "[| (h O g O f): inj(A,A); \ | 
| 145 | \ (f O h O g): inj(B,B); \ | |
| 146 | \ (g O f O h): surj(C,C); \ | |
| 0 | 147 | \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; | 
| 148 | by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
736diff
changeset | 149 | qed "pastre6"; | 
| 0 | 150 | |
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 151 | (** Yet another example... **) | 
| 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 152 | |
| 2469 | 153 | goal Perm.thy | 
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 154 |     "(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 | 155 | \ : bij(Pow(A+B), Pow(A)*Pow(B))"; | 
| 1110 
756aa2e81f6e
Changed some definitions and proofs to use pattern-matching.
 lcp parents: 
782diff
changeset | 156 | 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 | 157 | lam_bijective 1); | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4322diff
changeset | 158 | (*Auto_tac no longer proves it*) | 
| 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4322diff
changeset | 159 | by (REPEAT (fast_tac (claset() addss (simpset())) 1)); | 
| 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4322diff
changeset | 160 | qed "Pow_bij"; | 
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 161 | |
| 0 | 162 | writeln"Reached end of file."; |