| author | paulson | 
| Fri, 25 Sep 1998 14:05:13 +0200 | |
| changeset 5565 | 301a3a4d3dc7 | 
| parent 5321 | f8848433d240 | 
| child 6153 | bff90585cce5 | 
| 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  | 
||
9  | 
open Sum;  | 
|
10  | 
||
| 744 | 11  | 
(*** Rules for the Part primitive ***)  | 
12  | 
||
| 5067 | 13  | 
Goalw [Part_def]  | 
| 744 | 14  | 
"a : Part(A,h) <-> a:A & (EX y. a=h(y))";  | 
15  | 
by (rtac separation 1);  | 
|
| 760 | 16  | 
qed "Part_iff";  | 
| 744 | 17  | 
|
| 5067 | 18  | 
Goalw [Part_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
19  | 
"[| a : A; a=h(b) |] ==> a : Part(A,h)";  | 
| 2925 | 20  | 
by (Blast_tac 1);  | 
| 760 | 21  | 
qed "Part_eqI";  | 
| 744 | 22  | 
|
23  | 
val PartI = refl RSN (2,Part_eqI);  | 
|
24  | 
||
| 5321 | 25  | 
val major::prems = Goalw [Part_def]  | 
| 744 | 26  | 
"[| a : Part(A,h); !!z. [| a : A; a=h(z) |] ==> P \  | 
27  | 
\ |] ==> P";  | 
|
28  | 
by (rtac (major RS CollectE) 1);  | 
|
29  | 
by (etac exE 1);  | 
|
30  | 
by (REPEAT (ares_tac prems 1));  | 
|
| 760 | 31  | 
qed "PartE";  | 
| 744 | 32  | 
|
| 2925 | 33  | 
AddIs [Part_eqI];  | 
| 2469 | 34  | 
AddSEs [PartE];  | 
35  | 
||
| 5067 | 36  | 
Goalw [Part_def] "Part(A,h) <= A";  | 
| 744 | 37  | 
by (rtac Collect_subset 1);  | 
| 760 | 38  | 
qed "Part_subset";  | 
| 744 | 39  | 
|
40  | 
||
41  | 
(*** Rules for Disjoint Sums ***)  | 
|
42  | 
||
| 0 | 43  | 
val sum_defs = [sum_def,Inl_def,Inr_def,case_def];  | 
44  | 
||
| 5067 | 45  | 
Goalw (bool_def::sum_defs) "Sigma(bool,C) = C(0) + C(1)";  | 
| 2925 | 46  | 
by (Blast_tac 1);  | 
| 760 | 47  | 
qed "Sigma_bool";  | 
| 521 | 48  | 
|
| 0 | 49  | 
(** Introduction rules for the injections **)  | 
50  | 
||
| 5137 | 51  | 
Goalw sum_defs "a : A ==> Inl(a) : A+B";  | 
| 2925 | 52  | 
by (Blast_tac 1);  | 
| 760 | 53  | 
qed "InlI";  | 
| 0 | 54  | 
|
| 5137 | 55  | 
Goalw sum_defs "b : B ==> Inr(b) : A+B";  | 
| 2925 | 56  | 
by (Blast_tac 1);  | 
| 760 | 57  | 
qed "InrI";  | 
| 0 | 58  | 
|
59  | 
(** Elimination rules **)  | 
|
60  | 
||
| 5321 | 61  | 
val major::prems = Goalw sum_defs  | 
| 0 | 62  | 
"[| u: A+B; \  | 
63  | 
\ !!x. [| x:A; u=Inl(x) |] ==> P; \  | 
|
64  | 
\ !!y. [| y:B; u=Inr(y) |] ==> P \  | 
|
65  | 
\ |] ==> P";  | 
|
66  | 
by (rtac (major RS UnE) 1);  | 
|
67  | 
by (REPEAT (rtac refl 1  | 
|
68  | 
ORELSE eresolve_tac (prems@[SigmaE,singletonE,ssubst]) 1));  | 
|
| 760 | 69  | 
qed "sumE";  | 
| 0 | 70  | 
|
71  | 
(** Injection and freeness equivalences, for rewriting **)  | 
|
72  | 
||
| 5067 | 73  | 
Goalw sum_defs "Inl(a)=Inl(b) <-> a=b";  | 
| 2469 | 74  | 
by (Simp_tac 1);  | 
| 760 | 75  | 
qed "Inl_iff";  | 
| 0 | 76  | 
|
| 5067 | 77  | 
Goalw sum_defs "Inr(a)=Inr(b) <-> a=b";  | 
| 2469 | 78  | 
by (Simp_tac 1);  | 
| 760 | 79  | 
qed "Inr_iff";  | 
| 0 | 80  | 
|
| 5067 | 81  | 
Goalw sum_defs "Inl(a)=Inr(b) <-> False";  | 
| 2469 | 82  | 
by (Simp_tac 1);  | 
| 760 | 83  | 
qed "Inl_Inr_iff";  | 
| 0 | 84  | 
|
| 5067 | 85  | 
Goalw sum_defs "Inr(b)=Inl(a) <-> False";  | 
| 2469 | 86  | 
by (Simp_tac 1);  | 
| 760 | 87  | 
qed "Inr_Inl_iff";  | 
| 0 | 88  | 
|
| 5067 | 89  | 
Goalw sum_defs "0+0 = 0";  | 
| 2469 | 90  | 
by (Simp_tac 1);  | 
91  | 
qed "sum_empty";  | 
|
92  | 
||
| 55 | 93  | 
(*Injection and freeness rules*)  | 
94  | 
||
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
773 
diff
changeset
 | 
95  | 
bind_thm ("Inl_inject", (Inl_iff RS iffD1));
 | 
| 
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
773 
diff
changeset
 | 
96  | 
bind_thm ("Inr_inject", (Inr_iff RS iffD1));
 | 
| 
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
773 
diff
changeset
 | 
97  | 
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
 | 
98  | 
bind_thm ("Inr_neq_Inl", (Inr_Inl_iff RS iffD1 RS FalseE));
 | 
| 55 | 99  | 
|
| 2469 | 100  | 
AddSIs [InlI, InrI];  | 
101  | 
AddSEs [sumE, Inl_neq_Inr, Inr_neq_Inl];  | 
|
102  | 
AddSDs [Inl_inject, Inr_inject];  | 
|
103  | 
||
104  | 
Addsimps [InlI, InrI, Inl_iff, Inr_iff, Inl_Inr_iff, Inr_Inl_iff, sum_empty];  | 
|
105  | 
||
| 5137 | 106  | 
Goal "Inl(a): A+B ==> a: A";  | 
| 2925 | 107  | 
by (Blast_tac 1);  | 
| 760 | 108  | 
qed "InlD";  | 
| 55 | 109  | 
|
| 5137 | 110  | 
Goal "Inr(b): A+B ==> b: B";  | 
| 2925 | 111  | 
by (Blast_tac 1);  | 
| 760 | 112  | 
qed "InrD";  | 
| 55 | 113  | 
|
| 5067 | 114  | 
Goal "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))";  | 
| 2925 | 115  | 
by (Blast_tac 1);  | 
| 760 | 116  | 
qed "sum_iff";  | 
| 0 | 117  | 
|
| 5067 | 118  | 
Goal "A+B <= C+D <-> A<=C & B<=D";  | 
| 2925 | 119  | 
by (Blast_tac 1);  | 
| 760 | 120  | 
qed "sum_subset_iff";  | 
| 0 | 121  | 
|
| 5067 | 122  | 
Goal "A+B = C+D <-> A=C & B=D";  | 
| 4091 | 123  | 
by (simp_tac (simpset() addsimps [extension,sum_subset_iff]) 1);  | 
| 2925 | 124  | 
by (Blast_tac 1);  | 
| 760 | 125  | 
qed "sum_equal_iff";  | 
| 0 | 126  | 
|
| 5067 | 127  | 
Goalw [sum_def] "A+A = 2*A";  | 
| 2925 | 128  | 
by (Blast_tac 1);  | 
| 1611 | 129  | 
qed "sum_eq_2_times";  | 
130  | 
||
| 521 | 131  | 
|
| 0 | 132  | 
(*** Eliminator -- case ***)  | 
133  | 
||
| 5067 | 134  | 
Goalw sum_defs "case(c, d, Inl(a)) = c(a)";  | 
| 4091 | 135  | 
by (simp_tac (simpset() addsimps [cond_0]) 1);  | 
| 760 | 136  | 
qed "case_Inl";  | 
| 0 | 137  | 
|
| 5067 | 138  | 
Goalw sum_defs "case(c, d, Inr(b)) = d(b)";  | 
| 4091 | 139  | 
by (simp_tac (simpset() addsimps [cond_1]) 1);  | 
| 760 | 140  | 
qed "case_Inr";  | 
| 0 | 141  | 
|
| 2469 | 142  | 
Addsimps [case_Inl, case_Inr];  | 
143  | 
||
| 5321 | 144  | 
val major::prems = Goal  | 
| 0 | 145  | 
"[| u: A+B; \  | 
146  | 
\ !!x. x: A ==> c(x): C(Inl(x)); \  | 
|
147  | 
\ !!y. y: B ==> d(y): C(Inr(y)) \  | 
|
148  | 
\ |] ==> case(c,d,u) : C(u)";  | 
|
149  | 
by (rtac (major RS sumE) 1);  | 
|
150  | 
by (ALLGOALS (etac ssubst));  | 
|
| 4091 | 151  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps prems)));  | 
| 760 | 152  | 
qed "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";  |