| author | paulson | 
| Thu, 28 Mar 1996 10:52:59 +0100 | |
| changeset 1623 | 2b8573c1b1c1 | 
| parent 1461 | 6bcb44e4d6e5 | 
| child 1646 | f48fafc18a88 | 
| 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  | 
|
7  | 
Cantor's Theorem; Schroeder-Bernstein Theorem; Composition of homomorphisms...  | 
|
8  | 
*)  | 
|
9  | 
||
10  | 
writeln"ZF/ex/misc";  | 
|
11  | 
||
12  | 
||
13  | 
(*Example 12 (credited to Peter Andrews) from  | 
|
14  | 
W. Bledsoe. A Maximal Method for Set Variables in Automatic Theorem-proving.  | 
|
15  | 
In: J. Hayes and D. Michie and L. Mikulich, eds. Machine Intelligence 9.  | 
|
16  | 
Ellis Horwood, 53-100 (1979). *)  | 
|
17  | 
goal ZF.thy "(ALL F. {x}: F --> {y}:F) --> (ALL A. x:A --> y:A)";
 | 
|
18  | 
by (best_tac ZF_cs 1);  | 
|
19  | 
result();  | 
|
20  | 
||
21  | 
||
22  | 
(*** Composition of homomorphisms is a homomorphism ***)  | 
|
23  | 
||
24  | 
(*Given as a challenge problem in  | 
|
25  | 
R. Boyer et al.,  | 
|
26  | 
Set Theory in First-Order Logic: Clauses for G\"odel's Axioms,  | 
|
27  | 
JAR 2 (1986), 287-327  | 
|
28  | 
*)  | 
|
29  | 
||
| 
7
 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 
lcp 
parents: 
0 
diff
changeset
 | 
30  | 
(*collecting the relevant lemmas*)  | 
| 434 | 31  | 
val hom_ss = ZF_ss addsimps [comp_fun,comp_fun_apply,SigmaI,apply_funtype];  | 
| 0 | 32  | 
|
| 736 | 33  | 
(*This version uses a super application of simp_tac. Needs setloop to help  | 
34  | 
proving conditions of rewrites such as comp_fun_apply;  | 
|
35  | 
rewriting does not instantiate Vars*)  | 
|
| 
7
 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 
lcp 
parents: 
0 
diff
changeset
 | 
36  | 
goal Perm.thy  | 
| 0 | 37  | 
"(ALL A f B g. hom(A,f,B,g) = \  | 
38  | 
\          {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: 
0 
diff
changeset
 | 
39  | 
\ (ALL x:A. ALL y:A. H`(f`<x,y>) = g`<H`x,H`y>)}) --> \  | 
| 0 | 40  | 
\ J : hom(A,f,B,g) & K : hom(B,g,C,h) --> \  | 
41  | 
\ (K O J) : hom(A,f,C,h)";  | 
|
| 736 | 42  | 
by (asm_simp_tac (hom_ss setloop (K (safe_tac FOL_cs))) 1);  | 
| 0 | 43  | 
val comp_homs = result();  | 
44  | 
||
| 
7
 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 
lcp 
parents: 
0 
diff
changeset
 | 
45  | 
(*This version uses meta-level rewriting, safe_tac and asm_simp_tac*)  | 
| 0 | 46  | 
val [hom_def] = goal Perm.thy  | 
47  | 
"(!! A f B g. hom(A,f,B,g) == \  | 
|
48  | 
\          {H: A->B. f:A*A->A & g:B*B->B & \
 | 
|
49  | 
\ (ALL x:A. ALL y:A. H`(f`<x,y>) = g`<H`x,H`y>)}) ==> \  | 
|
50  | 
\ J : hom(A,f,B,g) & K : hom(B,g,C,h) --> \  | 
|
51  | 
\ (K O J) : hom(A,f,C,h)";  | 
|
52  | 
by (rewtac hom_def);  | 
|
53  | 
by (safe_tac ZF_cs);  | 
|
| 
7
 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 
lcp 
parents: 
0 
diff
changeset
 | 
54  | 
by (asm_simp_tac hom_ss 1);  | 
| 
 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 
lcp 
parents: 
0 
diff
changeset
 | 
55  | 
by (asm_simp_tac hom_ss 1);  | 
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
736 
diff
changeset
 | 
56  | 
qed "comp_homs";  | 
| 0 | 57  | 
|
58  | 
||
59  | 
(** A characterization of functions, suggested by Tobias Nipkow **)  | 
|
60  | 
||
| 
695
 
a1586fa1b755
ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
 
lcp 
parents: 
434 
diff
changeset
 | 
61  | 
goalw ZF.thy [Pi_def, function_def]  | 
| 0 | 62  | 
"r: domain(r)->B <-> r <= domain(r)*B & (ALL X. r `` (r -`` X) <= X)";  | 
63  | 
by (safe_tac ZF_cs);  | 
|
| 
695
 
a1586fa1b755
ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
 
lcp 
parents: 
434 
diff
changeset
 | 
64  | 
by (fast_tac ZF_cs 1);  | 
| 0 | 65  | 
by (eres_inst_tac [("x", "{y}")] allE 1);
 | 
66  | 
by (fast_tac ZF_cs 1);  | 
|
67  | 
result();  | 
|
68  | 
||
69  | 
||
70  | 
(**** From D Pastre. Automatic theorem proving in set theory.  | 
|
71  | 
Artificial Intelligence, 10:1--27, 1978.  | 
|
72  | 
These examples require forward reasoning! ****)  | 
|
73  | 
||
74  | 
(*reduce the clauses to units by type checking -- beware of nontermination*)  | 
|
75  | 
fun forw_typechk tyrls [] = []  | 
|
76  | 
| forw_typechk tyrls clauses =  | 
|
77  | 
let val (units, others) = partition (has_fewer_prems 1) clauses  | 
|
78  | 
in gen_union eq_thm (units, forw_typechk tyrls (tyrls RL others))  | 
|
79  | 
end;  | 
|
80  | 
||
81  | 
(*A crude form of forward reasoning*)  | 
|
82  | 
fun forw_iterate tyrls rls facts 0 = facts  | 
|
83  | 
| forw_iterate tyrls rls facts n =  | 
|
84  | 
let val facts' =  | 
|
| 1461 | 85  | 
gen_union eq_thm (forw_typechk (tyrls@facts) (facts RL rls), facts);  | 
| 0 | 86  | 
in forw_iterate tyrls rls facts' (n-1) end;  | 
87  | 
||
88  | 
val pastre_rls =  | 
|
89  | 
[comp_mem_injD1, comp_mem_surjD1, comp_mem_injD2, comp_mem_surjD2];  | 
|
90  | 
||
91  | 
fun pastre_facts (fact1::fact2::fact3::prems) =  | 
|
| 434 | 92  | 
forw_iterate (prems @ [comp_surj, comp_inj, comp_fun])  | 
| 0 | 93  | 
pastre_rls [fact1,fact2,fact3] 4;  | 
94  | 
||
95  | 
val prems = goalw Perm.thy [bij_def]  | 
|
| 1461 | 96  | 
"[| (h O g O f): inj(A,A); \  | 
97  | 
\ (f O h O g): surj(B,B); \  | 
|
98  | 
\ (g O f O h): surj(C,C); \  | 
|
| 0 | 99  | 
\ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)";  | 
100  | 
by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1));  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
736 
diff
changeset
 | 
101  | 
qed "pastre1";  | 
| 0 | 102  | 
|
103  | 
val prems = goalw Perm.thy [bij_def]  | 
|
| 1461 | 104  | 
"[| (h O g O f): surj(A,A); \  | 
105  | 
\ (f O h O g): inj(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: 
736 
diff
changeset
 | 
109  | 
qed "pastre2";  | 
| 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): surj(B,B); \  | 
|
114  | 
\ (g O f O h): inj(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: 
736 
diff
changeset
 | 
117  | 
qed "pastre3";  | 
| 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): inj(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: 
736 
diff
changeset
 | 
125  | 
qed "pastre4";  | 
| 0 | 126  | 
|
127  | 
val prems = goalw Perm.thy [bij_def]  | 
|
| 1461 | 128  | 
"[| (h O g O f): inj(A,A); \  | 
129  | 
\ (f O h O g): surj(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: 
736 
diff
changeset
 | 
133  | 
qed "pastre5";  | 
| 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): inj(B,B); \  | 
|
138  | 
\ (g O f O h): surj(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: 
736 
diff
changeset
 | 
141  | 
qed "pastre6";  | 
| 0 | 142  | 
|
| 
7
 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 
lcp 
parents: 
0 
diff
changeset
 | 
143  | 
(** Yet another example... **)  | 
| 
 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 
lcp 
parents: 
0 
diff
changeset
 | 
144  | 
|
| 
695
 
a1586fa1b755
ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
 
lcp 
parents: 
434 
diff
changeset
 | 
145  | 
goal (merge_theories(Sum.thy,Perm.thy))  | 
| 
7
 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 
lcp 
parents: 
0 
diff
changeset
 | 
146  | 
    "(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: 
0 
diff
changeset
 | 
147  | 
\ : bij(Pow(A+B), Pow(A)*Pow(B))";  | 
| 
1110
 
756aa2e81f6e
Changed some definitions and proofs to use pattern-matching.
 
lcp 
parents: 
782 
diff
changeset
 | 
148  | 
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: 
434 
diff
changeset
 | 
149  | 
lam_bijective 1);  | 
| 
 
a1586fa1b755
ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
 
lcp 
parents: 
434 
diff
changeset
 | 
150  | 
by (TRYALL (etac SigmaE));  | 
| 
 
a1586fa1b755
ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
 
lcp 
parents: 
434 
diff
changeset
 | 
151  | 
by (ALLGOALS (asm_simp_tac ZF_ss));  | 
| 
 
a1586fa1b755
ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
 
lcp 
parents: 
434 
diff
changeset
 | 
152  | 
by (ALLGOALS (fast_tac (sum_cs addSIs [equalityI])));  | 
| 
7
 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 
lcp 
parents: 
0 
diff
changeset
 | 
153  | 
val Pow_bij = result();  | 
| 
 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 
lcp 
parents: 
0 
diff
changeset
 | 
154  | 
|
| 0 | 155  | 
writeln"Reached end of file.";  |