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