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