| author | wenzelm | 
| Fri, 24 Oct 1997 17:18:49 +0200 | |
| changeset 3998 | 6ec8d42082f1 | 
| parent 3840 | e0baea4d485a | 
| child 4091 | 771b1f6422a8 | 
| permissions | -rw-r--r-- | 
| 1461 | 1  | 
(* Title: ZF/pair  | 
| 0 | 2  | 
ID: $Id$  | 
| 1461 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 0 | 4  | 
Copyright 1992 University of Cambridge  | 
5  | 
||
6  | 
Ordered pairs in Zermelo-Fraenkel Set Theory  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
(** Lemmas for showing that <a,b> uniquely determines a and b **)  | 
|
10  | 
||
| 2877 | 11  | 
qed_goal "singleton_eq_iff" ZF.thy  | 
| 
822
 
8d5748202c53
Added Krzysztof's theorems singleton_eq_iff, fst_type, snd_type
 
lcp 
parents: 
782 
diff
changeset
 | 
12  | 
    "{a} = {b} <-> a=b"
 | 
| 
 
8d5748202c53
Added Krzysztof's theorems singleton_eq_iff, fst_type, snd_type
 
lcp 
parents: 
782 
diff
changeset
 | 
13  | 
(fn _=> [ (resolve_tac [extension RS iff_trans] 1),  | 
| 2877 | 14  | 
(Blast_tac 1) ]);  | 
| 
822
 
8d5748202c53
Added Krzysztof's theorems singleton_eq_iff, fst_type, snd_type
 
lcp 
parents: 
782 
diff
changeset
 | 
15  | 
|
| 2877 | 16  | 
qed_goal "doubleton_eq_iff" ZF.thy  | 
| 0 | 17  | 
    "{a,b} = {c,d} <-> (a=c & b=d) | (a=d & b=c)"
 | 
18  | 
(fn _=> [ (resolve_tac [extension RS iff_trans] 1),  | 
|
| 2877 | 19  | 
(Blast_tac 1) ]);  | 
| 0 | 20  | 
|
| 2877 | 21  | 
qed_goalw "Pair_iff" ZF.thy [Pair_def]  | 
| 0 | 22  | 
"<a,b> = <c,d> <-> a=c & b=d"  | 
| 2469 | 23  | 
(fn _=> [ (simp_tac (!simpset addsimps [doubleton_eq_iff]) 1),  | 
| 2877 | 24  | 
(Blast_tac 1) ]);  | 
| 0 | 25  | 
|
| 2469 | 26  | 
Addsimps [Pair_iff];  | 
27  | 
||
28  | 
bind_thm ("Pair_inject", Pair_iff RS iffD1 RS conjE);
 | 
|
| 0 | 29  | 
|
| 2469 | 30  | 
AddSEs [Pair_inject];  | 
31  | 
||
32  | 
bind_thm ("Pair_inject1", Pair_iff RS iffD1 RS conjunct1);
 | 
|
33  | 
bind_thm ("Pair_inject2", Pair_iff RS iffD1 RS conjunct2);
 | 
|
| 0 | 34  | 
|
| 2877 | 35  | 
qed_goalw "Pair_not_0" ZF.thy [Pair_def] "<a,b> ~= 0"  | 
36  | 
(fn _ => [ (blast_tac (!claset addEs [equalityE]) 1) ]);  | 
|
| 0 | 37  | 
|
| 2469 | 38  | 
bind_thm ("Pair_neq_0", Pair_not_0 RS notE);
 | 
| 0 | 39  | 
|
| 2469 | 40  | 
AddSEs [Pair_neq_0, sym RS Pair_neq_0];  | 
41  | 
||
| 2877 | 42  | 
qed_goalw "Pair_neq_fst" ZF.thy [Pair_def] "<a,b>=a ==> P"  | 
| 0 | 43  | 
(fn [major]=>  | 
| 437 | 44  | 
[ (rtac (consI1 RS mem_asym RS FalseE) 1),  | 
| 0 | 45  | 
(rtac (major RS subst) 1),  | 
46  | 
(rtac consI1 1) ]);  | 
|
47  | 
||
| 2877 | 48  | 
qed_goalw "Pair_neq_snd" ZF.thy [Pair_def] "<a,b>=b ==> P"  | 
| 0 | 49  | 
(fn [major]=>  | 
| 437 | 50  | 
[ (rtac (consI1 RS consI2 RS mem_asym RS FalseE) 1),  | 
| 0 | 51  | 
(rtac (major RS subst) 1),  | 
52  | 
(rtac (consI1 RS consI2) 1) ]);  | 
|
53  | 
||
54  | 
||
55  | 
(*** Sigma: Disjoint union of a family of sets  | 
|
56  | 
Generalizes Cartesian product ***)  | 
|
57  | 
||
| 2877 | 58  | 
qed_goalw "Sigma_iff" ZF.thy [Sigma_def] "<a,b>: Sigma(A,B) <-> a:A & b:B(a)"  | 
59  | 
(fn _ => [ Blast_tac 1 ]);  | 
|
| 2469 | 60  | 
|
61  | 
Addsimps [Sigma_iff];  | 
|
62  | 
||
| 2877 | 63  | 
qed_goal "SigmaI" ZF.thy  | 
| 2469 | 64  | 
"!!a b. [| a:A; b:B(a) |] ==> <a,b> : Sigma(A,B)"  | 
65  | 
(fn _ => [ Asm_simp_tac 1 ]);  | 
|
66  | 
||
67  | 
bind_thm ("SigmaD1", Sigma_iff RS iffD1 RS conjunct1);
 | 
|
68  | 
bind_thm ("SigmaD2", Sigma_iff RS iffD1 RS conjunct2);
 | 
|
| 0 | 69  | 
|
70  | 
(*The general elimination rule*)  | 
|
| 2877 | 71  | 
qed_goalw "SigmaE" ZF.thy [Sigma_def]  | 
| 0 | 72  | 
"[| c: Sigma(A,B); \  | 
73  | 
\ !!x y.[| x:A; y:B(x); c=<x,y> |] ==> P \  | 
|
74  | 
\ |] ==> P"  | 
|
75  | 
(fn major::prems=>  | 
|
76  | 
[ (cut_facts_tac [major] 1),  | 
|
77  | 
(REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]);  | 
|
78  | 
||
| 2877 | 79  | 
qed_goal "SigmaE2" ZF.thy  | 
| 0 | 80  | 
"[| <a,b> : Sigma(A,B); \  | 
81  | 
\ [| a:A; b:B(a) |] ==> P \  | 
|
82  | 
\ |] ==> P"  | 
|
83  | 
(fn [major,minor]=>  | 
|
84  | 
[ (rtac minor 1),  | 
|
85  | 
(rtac (major RS SigmaD1) 1),  | 
|
86  | 
(rtac (major RS SigmaD2) 1) ]);  | 
|
87  | 
||
| 2877 | 88  | 
qed_goalw "Sigma_cong" ZF.thy [Sigma_def]  | 
| 0 | 89  | 
"[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \  | 
90  | 
\ Sigma(A,B) = Sigma(A',B')"  | 
|
| 2469 | 91  | 
(fn prems=> [ (simp_tac (!simpset addsimps prems) 1) ]);  | 
92  | 
||
| 0 | 93  | 
|
| 2469 | 94  | 
(*Sigma_cong, Pi_cong NOT given to Addcongs: they cause  | 
95  | 
flex-flex pairs and the "Check your prover" error. Most  | 
|
96  | 
Sigmas and Pis are abbreviated as * or -> *)  | 
|
| 0 | 97  | 
|
| 2469 | 98  | 
AddSIs [SigmaI];  | 
99  | 
AddSEs [SigmaE2, SigmaE];  | 
|
| 0 | 100  | 
|
| 2877 | 101  | 
qed_goal "Sigma_empty1" ZF.thy "Sigma(0,B) = 0"  | 
102  | 
(fn _ => [ (Blast_tac 1) ]);  | 
|
| 2469 | 103  | 
|
| 2877 | 104  | 
qed_goal "Sigma_empty2" ZF.thy "A*0 = 0"  | 
105  | 
(fn _ => [ (Blast_tac 1) ]);  | 
|
| 2469 | 106  | 
|
107  | 
Addsimps [Sigma_empty1, Sigma_empty2];  | 
|
| 1105 | 108  | 
|
109  | 
||
110  | 
(*** Projections: fst, snd ***)  | 
|
111  | 
||
| 2877 | 112  | 
qed_goalw "fst_conv" ZF.thy [fst_def] "fst(<a,b>) = a"  | 
113  | 
(fn _=> [ (blast_tac (!claset addIs [the_equality]) 1) ]);  | 
|
| 1105 | 114  | 
|
| 2877 | 115  | 
qed_goalw "snd_conv" ZF.thy [snd_def] "snd(<a,b>) = b"  | 
116  | 
(fn _=> [ (blast_tac (!claset addIs [the_equality]) 1) ]);  | 
|
| 1105 | 117  | 
|
| 2469 | 118  | 
Addsimps [fst_conv,snd_conv];  | 
| 1105 | 119  | 
|
| 2877 | 120  | 
qed_goal "fst_type" ZF.thy "!!p. p:Sigma(A,B) ==> fst(p) : A"  | 
| 2469 | 121  | 
(fn _=> [ Auto_tac() ]);  | 
| 1105 | 122  | 
|
| 2877 | 123  | 
qed_goal "snd_type" ZF.thy "!!p. p:Sigma(A,B) ==> snd(p) : B(fst(p))"  | 
| 2469 | 124  | 
(fn _=> [ Auto_tac() ]);  | 
| 1105 | 125  | 
|
| 2877 | 126  | 
qed_goal "Pair_fst_snd_eq" ZF.thy  | 
| 2469 | 127  | 
"!!a A B. a: Sigma(A,B) ==> <fst(a),snd(a)> = a"  | 
128  | 
(fn _=> [ Auto_tac() ]);  | 
|
| 1105 | 129  | 
|
| 0 | 130  | 
|
131  | 
(*** Eliminator - split ***)  | 
|
132  | 
||
| 1105 | 133  | 
(*A META-equality, so that it applies to higher types as well...*)  | 
| 3840 | 134  | 
qed_goalw "split" ZF.thy [split_def] "split(%x y. c(x,y), <a,b>) == c(a,b)"  | 
| 2469 | 135  | 
(fn _ => [ (Simp_tac 1),  | 
| 1105 | 136  | 
(rtac reflexive_thm 1) ]);  | 
| 0 | 137  | 
|
| 2469 | 138  | 
Addsimps [split];  | 
139  | 
||
| 2877 | 140  | 
qed_goal "split_type" ZF.thy  | 
| 0 | 141  | 
"[| p:Sigma(A,B); \  | 
142  | 
\ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x,y>) \  | 
|
| 3840 | 143  | 
\ |] ==> split(%x y. c(x,y), p) : C(p)"  | 
| 0 | 144  | 
(fn major::prems=>  | 
145  | 
[ (rtac (major RS SigmaE) 1),  | 
|
| 2469 | 146  | 
(asm_simp_tac (!simpset addsimps prems) 1) ]);  | 
| 0 | 147  | 
|
| 2877 | 148  | 
goalw ZF.thy [split_def]  | 
| 435 | 149  | 
"!!u. u: A*B ==> \  | 
150  | 
\ R(split(c,u)) <-> (ALL x:A. ALL y:B. u = <x,y> --> R(c(x,y)))";  | 
|
| 2469 | 151  | 
by (Auto_tac());  | 
| 760 | 152  | 
qed "expand_split";  | 
| 435 | 153  | 
|
154  | 
||
| 0 | 155  | 
(*** split for predicates: result type o ***)  | 
156  | 
||
| 2877 | 157  | 
goalw ZF.thy [split_def] "!!R a b. R(a,b) ==> split(R, <a,b>)";  | 
| 2469 | 158  | 
by (Asm_simp_tac 1);  | 
| 1105 | 159  | 
qed "splitI";  | 
| 0 | 160  | 
|
| 2877 | 161  | 
val major::sigma::prems = goalw ZF.thy [split_def]  | 
| 1461 | 162  | 
"[| split(R,z); z:Sigma(A,B); \  | 
163  | 
\ !!x y. [| z = <x,y>; R(x,y) |] ==> P \  | 
|
| 1105 | 164  | 
\ |] ==> P";  | 
165  | 
by (rtac (sigma RS SigmaE) 1);  | 
|
| 0 | 166  | 
by (cut_facts_tac [major] 1);  | 
| 2469 | 167  | 
by (REPEAT (ares_tac prems 1));  | 
168  | 
by (Asm_full_simp_tac 1);  | 
|
| 1105 | 169  | 
qed "splitE";  | 
| 0 | 170  | 
|
| 2877 | 171  | 
goalw ZF.thy [split_def] "!!R a b. split(R,<a,b>) ==> R(a,b)";  | 
| 2469 | 172  | 
by (Full_simp_tac 1);  | 
| 1105 | 173  | 
qed "splitD";  | 
| 0 | 174  | 
|
| 
533
 
7357160bc56a
ZF/pair.ML: moved some definitions here from simpdata.ML
 
lcp 
parents: 
437 
diff
changeset
 | 
175  | 
|
| 
 
7357160bc56a
ZF/pair.ML: moved some definitions here from simpdata.ML
 
lcp 
parents: 
437 
diff
changeset
 | 
176  |