author | paulson |
Wed, 28 Jun 2000 12:34:08 +0200 | |
changeset 9180 | 3bda56c0d70d |
parent 6153 | bff90585cce5 |
child 9211 | 6236c5285bd8 |
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 |
||
9180 | 11 |
Goal "{a} = {b} <-> a=b"; |
12 |
by (resolve_tac [extension RS iff_trans] 1); |
|
13 |
by (Blast_tac 1) ; |
|
14 |
qed "singleton_eq_iff"; |
|
822
8d5748202c53
Added Krzysztof's theorems singleton_eq_iff, fst_type, snd_type
lcp
parents:
782
diff
changeset
|
15 |
|
9180 | 16 |
Goal "{a,b} = {c,d} <-> (a=c & b=d) | (a=d & b=c)"; |
17 |
by (resolve_tac [extension RS iff_trans] 1); |
|
18 |
by (Blast_tac 1) ; |
|
19 |
qed "doubleton_eq_iff"; |
|
0 | 20 |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
21 |
qed_goalw "Pair_iff" thy [Pair_def] |
0 | 22 |
"<a,b> = <c,d> <-> a=c & b=d" |
4091 | 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 |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
35 |
qed_goalw "Pair_not_0" thy [Pair_def] "<a,b> ~= 0" |
4091 | 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 |
||
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
42 |
qed_goalw "Pair_neq_fst" 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 |
||
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
48 |
qed_goalw "Pair_neq_snd" 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 |
||
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
58 |
qed_goalw "Sigma_iff" thy [Sigma_def] "<a,b>: Sigma(A,B) <-> a:A & b:B(a)" |
2877 | 59 |
(fn _ => [ Blast_tac 1 ]); |
2469 | 60 |
|
61 |
Addsimps [Sigma_iff]; |
|
62 |
||
9180 | 63 |
Goal "[| a:A; b:B(a) |] ==> <a,b> : Sigma(A,B)"; |
64 |
by (Asm_simp_tac 1); |
|
65 |
qed "SigmaI"; |
|
66 |
||
6153 | 67 |
AddTCs [SigmaI]; |
2469 | 68 |
|
69 |
bind_thm ("SigmaD1", Sigma_iff RS iffD1 RS conjunct1); |
|
70 |
bind_thm ("SigmaD2", Sigma_iff RS iffD1 RS conjunct2); |
|
0 | 71 |
|
72 |
(*The general elimination rule*) |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
73 |
qed_goalw "SigmaE" thy [Sigma_def] |
0 | 74 |
"[| c: Sigma(A,B); \ |
75 |
\ !!x y.[| x:A; y:B(x); c=<x,y> |] ==> P \ |
|
76 |
\ |] ==> P" |
|
77 |
(fn major::prems=> |
|
78 |
[ (cut_facts_tac [major] 1), |
|
79 |
(REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); |
|
80 |
||
9180 | 81 |
val [major,minor]= Goal |
0 | 82 |
"[| <a,b> : Sigma(A,B); \ |
83 |
\ [| a:A; b:B(a) |] ==> P \ |
|
9180 | 84 |
\ |] ==> P"; |
85 |
by (rtac minor 1); |
|
86 |
by (rtac (major RS SigmaD1) 1); |
|
87 |
by (rtac (major RS SigmaD2) 1) ; |
|
88 |
qed "SigmaE2"; |
|
0 | 89 |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
90 |
qed_goalw "Sigma_cong" thy [Sigma_def] |
0 | 91 |
"[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \ |
92 |
\ Sigma(A,B) = Sigma(A',B')" |
|
4091 | 93 |
(fn prems=> [ (simp_tac (simpset() addsimps prems) 1) ]); |
2469 | 94 |
|
0 | 95 |
|
2469 | 96 |
(*Sigma_cong, Pi_cong NOT given to Addcongs: they cause |
97 |
flex-flex pairs and the "Check your prover" error. Most |
|
98 |
Sigmas and Pis are abbreviated as * or -> *) |
|
0 | 99 |
|
2469 | 100 |
AddSIs [SigmaI]; |
101 |
AddSEs [SigmaE2, SigmaE]; |
|
0 | 102 |
|
9180 | 103 |
Goal "Sigma(0,B) = 0"; |
104 |
by (Blast_tac 1) ; |
|
105 |
qed "Sigma_empty1"; |
|
2469 | 106 |
|
9180 | 107 |
Goal "A*0 = 0"; |
108 |
by (Blast_tac 1) ; |
|
109 |
qed "Sigma_empty2"; |
|
2469 | 110 |
|
111 |
Addsimps [Sigma_empty1, Sigma_empty2]; |
|
1105 | 112 |
|
5264 | 113 |
Goal "A*B=0 <-> A=0 | B=0"; |
114 |
by (Blast_tac 1); |
|
115 |
qed "Sigma_empty_iff"; |
|
116 |
||
1105 | 117 |
|
118 |
(*** Projections: fst, snd ***) |
|
119 |
||
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
120 |
qed_goalw "fst_conv" thy [fst_def] "fst(<a,b>) = a" |
5505 | 121 |
(fn _=> [ (Blast_tac 1) ]); |
1105 | 122 |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
123 |
qed_goalw "snd_conv" thy [snd_def] "snd(<a,b>) = b" |
5505 | 124 |
(fn _=> [ (Blast_tac 1) ]); |
1105 | 125 |
|
2469 | 126 |
Addsimps [fst_conv,snd_conv]; |
1105 | 127 |
|
9180 | 128 |
Goal "p:Sigma(A,B) ==> fst(p) : A"; |
129 |
by (Auto_tac) ; |
|
130 |
qed "fst_type"; |
|
1105 | 131 |
|
9180 | 132 |
Goal "p:Sigma(A,B) ==> snd(p) : B(fst(p))"; |
133 |
by (Auto_tac) ; |
|
134 |
qed "snd_type"; |
|
1105 | 135 |
|
9180 | 136 |
Goal "a: Sigma(A,B) ==> <fst(a),snd(a)> = a"; |
137 |
by (Auto_tac) ; |
|
138 |
qed "Pair_fst_snd_eq"; |
|
1105 | 139 |
|
0 | 140 |
|
141 |
(*** Eliminator - split ***) |
|
142 |
||
1105 | 143 |
(*A META-equality, so that it applies to higher types as well...*) |
6112 | 144 |
Goalw [split_def] "split(%x y. c(x,y), <a,b>) == c(a,b)"; |
145 |
by (Simp_tac 1); |
|
146 |
qed "split"; |
|
2469 | 147 |
Addsimps [split]; |
148 |
||
9180 | 149 |
val major::prems= Goal |
0 | 150 |
"[| p:Sigma(A,B); \ |
151 |
\ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x,y>) \ |
|
9180 | 152 |
\ |] ==> split(%x y. c(x,y), p) : C(p)"; |
153 |
by (rtac (major RS SigmaE) 1); |
|
154 |
by (asm_simp_tac (simpset() addsimps prems) 1); |
|
155 |
qed "split_type"; |
|
6153 | 156 |
AddTCs [split_type]; |
0 | 157 |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
158 |
Goalw [split_def] |
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
159 |
"u: A*B ==> \ |
435 | 160 |
\ R(split(c,u)) <-> (ALL x:A. ALL y:B. u = <x,y> --> R(c(x,y)))"; |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4091
diff
changeset
|
161 |
by Auto_tac; |
760 | 162 |
qed "expand_split"; |
435 | 163 |
|
164 |
||
0 | 165 |
(*** split for predicates: result type o ***) |
166 |
||
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
167 |
Goalw [split_def] "R(a,b) ==> split(R, <a,b>)"; |
2469 | 168 |
by (Asm_simp_tac 1); |
1105 | 169 |
qed "splitI"; |
0 | 170 |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
171 |
val major::sigma::prems = Goalw [split_def] |
1461 | 172 |
"[| split(R,z); z:Sigma(A,B); \ |
173 |
\ !!x y. [| z = <x,y>; R(x,y) |] ==> P \ |
|
1105 | 174 |
\ |] ==> P"; |
175 |
by (rtac (sigma RS SigmaE) 1); |
|
0 | 176 |
by (cut_facts_tac [major] 1); |
2469 | 177 |
by (REPEAT (ares_tac prems 1)); |
178 |
by (Asm_full_simp_tac 1); |
|
1105 | 179 |
qed "splitE"; |
0 | 180 |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5264
diff
changeset
|
181 |
Goalw [split_def] "split(R,<a,b>) ==> R(a,b)"; |
2469 | 182 |
by (Full_simp_tac 1); |
1105 | 183 |
qed "splitD"; |
0 | 184 |
|
533
7357160bc56a
ZF/pair.ML: moved some definitions here from simpdata.ML
lcp
parents:
437
diff
changeset
|
185 |
|
7357160bc56a
ZF/pair.ML: moved some definitions here from simpdata.ML
lcp
parents:
437
diff
changeset
|
186 |