author | wenzelm |
Tue, 16 Oct 2001 00:32:01 +0200 | |
changeset 11790 | 42393a11642d |
parent 9907 | 473a6604da94 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/Sum |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
6 |
Disjoint sums in Zermelo-Fraenkel Set Theory |
|
7 |
*) |
|
8 |
||
744 | 9 |
(*** Rules for the Part primitive ***) |
10 |
||
5067 | 11 |
Goalw [Part_def] |
744 | 12 |
"a : Part(A,h) <-> a:A & (EX y. a=h(y))"; |
13 |
by (rtac separation 1); |
|
760 | 14 |
qed "Part_iff"; |
744 | 15 |
|
5067 | 16 |
Goalw [Part_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
17 |
"[| a : A; a=h(b) |] ==> a : Part(A,h)"; |
2925 | 18 |
by (Blast_tac 1); |
760 | 19 |
qed "Part_eqI"; |
744 | 20 |
|
9907 | 21 |
bind_thm ("PartI", refl RSN (2,Part_eqI)); |
744 | 22 |
|
5321 | 23 |
val major::prems = Goalw [Part_def] |
744 | 24 |
"[| a : Part(A,h); !!z. [| a : A; a=h(z) |] ==> P \ |
25 |
\ |] ==> P"; |
|
26 |
by (rtac (major RS CollectE) 1); |
|
27 |
by (etac exE 1); |
|
28 |
by (REPEAT (ares_tac prems 1)); |
|
760 | 29 |
qed "PartE"; |
744 | 30 |
|
2925 | 31 |
AddIs [Part_eqI]; |
2469 | 32 |
AddSEs [PartE]; |
33 |
||
5067 | 34 |
Goalw [Part_def] "Part(A,h) <= A"; |
744 | 35 |
by (rtac Collect_subset 1); |
760 | 36 |
qed "Part_subset"; |
744 | 37 |
|
38 |
||
39 |
(*** Rules for Disjoint Sums ***) |
|
40 |
||
9907 | 41 |
bind_thms ("sum_defs", [sum_def,Inl_def,Inr_def,case_def]); |
0 | 42 |
|
5067 | 43 |
Goalw (bool_def::sum_defs) "Sigma(bool,C) = C(0) + C(1)"; |
2925 | 44 |
by (Blast_tac 1); |
760 | 45 |
qed "Sigma_bool"; |
521 | 46 |
|
0 | 47 |
(** Introduction rules for the injections **) |
48 |
||
5137 | 49 |
Goalw sum_defs "a : A ==> Inl(a) : A+B"; |
2925 | 50 |
by (Blast_tac 1); |
760 | 51 |
qed "InlI"; |
0 | 52 |
|
5137 | 53 |
Goalw sum_defs "b : B ==> Inr(b) : A+B"; |
2925 | 54 |
by (Blast_tac 1); |
760 | 55 |
qed "InrI"; |
0 | 56 |
|
57 |
(** Elimination rules **) |
|
58 |
||
5321 | 59 |
val major::prems = Goalw sum_defs |
0 | 60 |
"[| u: A+B; \ |
61 |
\ !!x. [| x:A; u=Inl(x) |] ==> P; \ |
|
62 |
\ !!y. [| y:B; u=Inr(y) |] ==> P \ |
|
63 |
\ |] ==> P"; |
|
64 |
by (rtac (major RS UnE) 1); |
|
65 |
by (REPEAT (rtac refl 1 |
|
66 |
ORELSE eresolve_tac (prems@[SigmaE,singletonE,ssubst]) 1)); |
|
760 | 67 |
qed "sumE"; |
0 | 68 |
|
69 |
(** Injection and freeness equivalences, for rewriting **) |
|
70 |
||
5067 | 71 |
Goalw sum_defs "Inl(a)=Inl(b) <-> a=b"; |
2469 | 72 |
by (Simp_tac 1); |
760 | 73 |
qed "Inl_iff"; |
0 | 74 |
|
5067 | 75 |
Goalw sum_defs "Inr(a)=Inr(b) <-> a=b"; |
2469 | 76 |
by (Simp_tac 1); |
760 | 77 |
qed "Inr_iff"; |
0 | 78 |
|
5067 | 79 |
Goalw sum_defs "Inl(a)=Inr(b) <-> False"; |
2469 | 80 |
by (Simp_tac 1); |
760 | 81 |
qed "Inl_Inr_iff"; |
0 | 82 |
|
5067 | 83 |
Goalw sum_defs "Inr(b)=Inl(a) <-> False"; |
2469 | 84 |
by (Simp_tac 1); |
760 | 85 |
qed "Inr_Inl_iff"; |
0 | 86 |
|
5067 | 87 |
Goalw sum_defs "0+0 = 0"; |
2469 | 88 |
by (Simp_tac 1); |
89 |
qed "sum_empty"; |
|
90 |
||
55 | 91 |
(*Injection and freeness rules*) |
92 |
||
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
773
diff
changeset
|
93 |
bind_thm ("Inl_inject", (Inl_iff RS iffD1)); |
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
773
diff
changeset
|
94 |
bind_thm ("Inr_inject", (Inr_iff RS iffD1)); |
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
773
diff
changeset
|
95 |
bind_thm ("Inl_neq_Inr", (Inl_Inr_iff RS iffD1 RS FalseE)); |
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
773
diff
changeset
|
96 |
bind_thm ("Inr_neq_Inl", (Inr_Inl_iff RS iffD1 RS FalseE)); |
55 | 97 |
|
2469 | 98 |
AddSIs [InlI, InrI]; |
99 |
AddSEs [sumE, Inl_neq_Inr, Inr_neq_Inl]; |
|
100 |
AddSDs [Inl_inject, Inr_inject]; |
|
101 |
||
102 |
Addsimps [InlI, InrI, Inl_iff, Inr_iff, Inl_Inr_iff, Inr_Inl_iff, sum_empty]; |
|
6153 | 103 |
AddTCs [InlI, InrI]; |
2469 | 104 |
|
5137 | 105 |
Goal "Inl(a): A+B ==> a: A"; |
2925 | 106 |
by (Blast_tac 1); |
760 | 107 |
qed "InlD"; |
55 | 108 |
|
5137 | 109 |
Goal "Inr(b): A+B ==> b: B"; |
2925 | 110 |
by (Blast_tac 1); |
760 | 111 |
qed "InrD"; |
55 | 112 |
|
5067 | 113 |
Goal "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))"; |
2925 | 114 |
by (Blast_tac 1); |
760 | 115 |
qed "sum_iff"; |
0 | 116 |
|
5067 | 117 |
Goal "A+B <= C+D <-> A<=C & B<=D"; |
2925 | 118 |
by (Blast_tac 1); |
760 | 119 |
qed "sum_subset_iff"; |
0 | 120 |
|
5067 | 121 |
Goal "A+B = C+D <-> A=C & B=D"; |
4091 | 122 |
by (simp_tac (simpset() addsimps [extension,sum_subset_iff]) 1); |
2925 | 123 |
by (Blast_tac 1); |
760 | 124 |
qed "sum_equal_iff"; |
0 | 125 |
|
5067 | 126 |
Goalw [sum_def] "A+A = 2*A"; |
2925 | 127 |
by (Blast_tac 1); |
1611 | 128 |
qed "sum_eq_2_times"; |
129 |
||
521 | 130 |
|
0 | 131 |
(*** Eliminator -- case ***) |
132 |
||
5067 | 133 |
Goalw sum_defs "case(c, d, Inl(a)) = c(a)"; |
8201 | 134 |
by (Simp_tac 1); |
760 | 135 |
qed "case_Inl"; |
0 | 136 |
|
5067 | 137 |
Goalw sum_defs "case(c, d, Inr(b)) = d(b)"; |
8201 | 138 |
by (Simp_tac 1); |
760 | 139 |
qed "case_Inr"; |
0 | 140 |
|
2469 | 141 |
Addsimps [case_Inl, case_Inr]; |
142 |
||
5321 | 143 |
val major::prems = Goal |
0 | 144 |
"[| u: A+B; \ |
145 |
\ !!x. x: A ==> c(x): C(Inl(x)); \ |
|
146 |
\ !!y. y: B ==> d(y): C(Inr(y)) \ |
|
147 |
\ |] ==> case(c,d,u) : C(u)"; |
|
148 |
by (rtac (major RS sumE) 1); |
|
149 |
by (ALLGOALS (etac ssubst)); |
|
4091 | 150 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); |
760 | 151 |
qed "case_type"; |
6153 | 152 |
AddTCs [case_type]; |
0 | 153 |
|
5268 | 154 |
Goal "u: A+B ==> \ |
435 | 155 |
\ R(case(c,d,u)) <-> \ |
156 |
\ ((ALL x:A. u = Inl(x) --> R(c(x))) & \ |
|
157 |
\ (ALL y:B. u = Inr(y) --> R(d(y))))"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4091
diff
changeset
|
158 |
by Auto_tac; |
760 | 159 |
qed "expand_case"; |
435 | 160 |
|
5321 | 161 |
val major::prems = Goal |
1461 | 162 |
"[| z: A+B; \ |
163 |
\ !!x. x:A ==> c(x)=c'(x); \ |
|
164 |
\ !!y. y:B ==> d(y)=d'(y) \ |
|
858 | 165 |
\ |] ==> case(c,d,z) = case(c',d',z)"; |
166 |
by (resolve_tac [major RS sumE] 1); |
|
4091 | 167 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); |
858 | 168 |
qed "case_cong"; |
169 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
170 |
Goal "z: A+B ==> \ |
2469 | 171 |
\ case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) = \ |
172 |
\ case(%x. c(c'(x)), %y. d(d'(y)), z)"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4091
diff
changeset
|
173 |
by Auto_tac; |
858 | 174 |
qed "case_case"; |
175 |
||
773
9f8bcf1a4cff
sum_ss: moved down and added the rewrite rules for "case"
lcp
parents:
760
diff
changeset
|
176 |
|
9f8bcf1a4cff
sum_ss: moved down and added the rewrite rules for "case"
lcp
parents:
760
diff
changeset
|
177 |
(*** More rules for Part(A,h) ***) |
9f8bcf1a4cff
sum_ss: moved down and added the rewrite rules for "case"
lcp
parents:
760
diff
changeset
|
178 |
|
5137 | 179 |
Goal "A<=B ==> Part(A,h)<=Part(B,h)"; |
2925 | 180 |
by (Blast_tac 1); |
760 | 181 |
qed "Part_mono"; |
0 | 182 |
|
5067 | 183 |
Goal "Part(Collect(A,P), h) = Collect(Part(A,h), P)"; |
2925 | 184 |
by (Blast_tac 1); |
760 | 185 |
qed "Part_Collect"; |
744 | 186 |
|
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
187 |
bind_thm ("Part_CollectE", Part_Collect RS equalityD1 RS subsetD RS CollectE); |
744 | 188 |
|
5067 | 189 |
Goal "Part(A+B,Inl) = {Inl(x). x: A}"; |
2925 | 190 |
by (Blast_tac 1); |
760 | 191 |
qed "Part_Inl"; |
0 | 192 |
|
5067 | 193 |
Goal "Part(A+B,Inr) = {Inr(y). y: B}"; |
2925 | 194 |
by (Blast_tac 1); |
760 | 195 |
qed "Part_Inr"; |
0 | 196 |
|
5137 | 197 |
Goalw [Part_def] "a : Part(A,h) ==> a : A"; |
0 | 198 |
by (etac CollectD1 1); |
760 | 199 |
qed "PartD1"; |
0 | 200 |
|
5067 | 201 |
Goal "Part(A,%x. x) = A"; |
2925 | 202 |
by (Blast_tac 1); |
760 | 203 |
qed "Part_id"; |
0 | 204 |
|
5067 | 205 |
Goal "Part(A+B, %x. Inr(h(x))) = {Inr(y). y: Part(B,h)}"; |
2925 | 206 |
by (Blast_tac 1); |
760 | 207 |
qed "Part_Inr2"; |
0 | 208 |
|
5137 | 209 |
Goal "C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C"; |
2925 | 210 |
by (Blast_tac 1); |
760 | 211 |
qed "Part_sum_equality"; |