1478
|
1 |
(* Title: ZF/sum.thy
|
0
|
2 |
ID: $Id$
|
1478
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
0
|
4 |
Copyright 1993 University of Cambridge
|
|
5 |
|
|
6 |
*)
|
|
7 |
|
13356
|
8 |
header{*Disjoint Sums*}
|
|
9 |
|
16417
|
10 |
theory Sum imports Bool equalities begin
|
3923
|
11 |
|
13356
|
12 |
text{*And the "Part" primitive for simultaneous recursive type definitions*}
|
|
13 |
|
3923
|
14 |
global
|
|
15 |
|
13240
|
16 |
constdefs
|
|
17 |
sum :: "[i,i]=>i" (infixr "+" 65)
|
|
18 |
"A+B == {0}*A Un {1}*B"
|
|
19 |
|
|
20 |
Inl :: "i=>i"
|
|
21 |
"Inl(a) == <0,a>"
|
|
22 |
|
|
23 |
Inr :: "i=>i"
|
|
24 |
"Inr(b) == <1,b>"
|
|
25 |
|
|
26 |
"case" :: "[i=>i, i=>i, i]=>i"
|
|
27 |
"case(c,d) == (%<y,z>. cond(y, d(z), c(z)))"
|
|
28 |
|
|
29 |
(*operator for selecting out the various summands*)
|
|
30 |
Part :: "[i,i=>i] => i"
|
|
31 |
"Part(A,h) == {x: A. EX z. x = h(z)}"
|
0
|
32 |
|
3940
|
33 |
local
|
3923
|
34 |
|
13356
|
35 |
subsection{*Rules for the @{term Part} Primitive*}
|
13240
|
36 |
|
|
37 |
lemma Part_iff:
|
|
38 |
"a : Part(A,h) <-> a:A & (EX y. a=h(y))"
|
|
39 |
apply (unfold Part_def)
|
|
40 |
apply (rule separation)
|
|
41 |
done
|
|
42 |
|
|
43 |
lemma Part_eqI [intro]:
|
|
44 |
"[| a : A; a=h(b) |] ==> a : Part(A,h)"
|
13255
|
45 |
by (unfold Part_def, blast)
|
13240
|
46 |
|
|
47 |
lemmas PartI = refl [THEN [2] Part_eqI]
|
|
48 |
|
|
49 |
lemma PartE [elim!]:
|
|
50 |
"[| a : Part(A,h); !!z. [| a : A; a=h(z) |] ==> P
|
|
51 |
|] ==> P"
|
13255
|
52 |
apply (unfold Part_def, blast)
|
13240
|
53 |
done
|
|
54 |
|
|
55 |
lemma Part_subset: "Part(A,h) <= A"
|
|
56 |
apply (unfold Part_def)
|
|
57 |
apply (rule Collect_subset)
|
|
58 |
done
|
|
59 |
|
|
60 |
|
13356
|
61 |
subsection{*Rules for Disjoint Sums*}
|
13240
|
62 |
|
|
63 |
lemmas sum_defs = sum_def Inl_def Inr_def case_def
|
|
64 |
|
|
65 |
lemma Sigma_bool: "Sigma(bool,C) = C(0) + C(1)"
|
13255
|
66 |
by (unfold bool_def sum_def, blast)
|
13240
|
67 |
|
|
68 |
(** Introduction rules for the injections **)
|
|
69 |
|
|
70 |
lemma InlI [intro!,simp,TC]: "a : A ==> Inl(a) : A+B"
|
13255
|
71 |
by (unfold sum_defs, blast)
|
13240
|
72 |
|
|
73 |
lemma InrI [intro!,simp,TC]: "b : B ==> Inr(b) : A+B"
|
13255
|
74 |
by (unfold sum_defs, blast)
|
13240
|
75 |
|
|
76 |
(** Elimination rules **)
|
|
77 |
|
|
78 |
lemma sumE [elim!]:
|
|
79 |
"[| u: A+B;
|
|
80 |
!!x. [| x:A; u=Inl(x) |] ==> P;
|
|
81 |
!!y. [| y:B; u=Inr(y) |] ==> P
|
|
82 |
|] ==> P"
|
13255
|
83 |
by (unfold sum_defs, blast)
|
13240
|
84 |
|
|
85 |
(** Injection and freeness equivalences, for rewriting **)
|
|
86 |
|
|
87 |
lemma Inl_iff [iff]: "Inl(a)=Inl(b) <-> a=b"
|
13255
|
88 |
by (simp add: sum_defs)
|
13240
|
89 |
|
|
90 |
lemma Inr_iff [iff]: "Inr(a)=Inr(b) <-> a=b"
|
13255
|
91 |
by (simp add: sum_defs)
|
13240
|
92 |
|
13823
|
93 |
lemma Inl_Inr_iff [simp]: "Inl(a)=Inr(b) <-> False"
|
13255
|
94 |
by (simp add: sum_defs)
|
13240
|
95 |
|
13823
|
96 |
lemma Inr_Inl_iff [simp]: "Inr(b)=Inl(a) <-> False"
|
13255
|
97 |
by (simp add: sum_defs)
|
13240
|
98 |
|
|
99 |
lemma sum_empty [simp]: "0+0 = 0"
|
13255
|
100 |
by (simp add: sum_defs)
|
13240
|
101 |
|
|
102 |
(*Injection and freeness rules*)
|
|
103 |
|
|
104 |
lemmas Inl_inject = Inl_iff [THEN iffD1, standard]
|
|
105 |
lemmas Inr_inject = Inr_iff [THEN iffD1, standard]
|
13823
|
106 |
lemmas Inl_neq_Inr = Inl_Inr_iff [THEN iffD1, THEN FalseE, elim!]
|
|
107 |
lemmas Inr_neq_Inl = Inr_Inl_iff [THEN iffD1, THEN FalseE, elim!]
|
13240
|
108 |
|
|
109 |
|
|
110 |
lemma InlD: "Inl(a): A+B ==> a: A"
|
13255
|
111 |
by blast
|
13240
|
112 |
|
|
113 |
lemma InrD: "Inr(b): A+B ==> b: B"
|
13255
|
114 |
by blast
|
13240
|
115 |
|
|
116 |
lemma sum_iff: "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))"
|
13255
|
117 |
by blast
|
|
118 |
|
|
119 |
lemma Inl_in_sum_iff [simp]: "(Inl(x) \<in> A+B) <-> (x \<in> A)";
|
|
120 |
by auto
|
|
121 |
|
|
122 |
lemma Inr_in_sum_iff [simp]: "(Inr(y) \<in> A+B) <-> (y \<in> B)";
|
|
123 |
by auto
|
13240
|
124 |
|
|
125 |
lemma sum_subset_iff: "A+B <= C+D <-> A<=C & B<=D"
|
13255
|
126 |
by blast
|
13240
|
127 |
|
|
128 |
lemma sum_equal_iff: "A+B = C+D <-> A=C & B=D"
|
13255
|
129 |
by (simp add: extension sum_subset_iff, blast)
|
13240
|
130 |
|
|
131 |
lemma sum_eq_2_times: "A+A = 2*A"
|
13255
|
132 |
by (simp add: sum_def, blast)
|
13240
|
133 |
|
|
134 |
|
13356
|
135 |
subsection{*The Eliminator: @{term case}*}
|
0
|
136 |
|
13240
|
137 |
lemma case_Inl [simp]: "case(c, d, Inl(a)) = c(a)"
|
13255
|
138 |
by (simp add: sum_defs)
|
13240
|
139 |
|
|
140 |
lemma case_Inr [simp]: "case(c, d, Inr(b)) = d(b)"
|
13255
|
141 |
by (simp add: sum_defs)
|
13240
|
142 |
|
|
143 |
lemma case_type [TC]:
|
|
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)"
|
13255
|
148 |
by auto
|
13240
|
149 |
|
|
150 |
lemma expand_case: "u: A+B ==>
|
|
151 |
R(case(c,d,u)) <->
|
|
152 |
((ALL x:A. u = Inl(x) --> R(c(x))) &
|
|
153 |
(ALL y:B. u = Inr(y) --> R(d(y))))"
|
|
154 |
by auto
|
|
155 |
|
|
156 |
lemma case_cong:
|
|
157 |
"[| z: A+B;
|
|
158 |
!!x. x:A ==> c(x)=c'(x);
|
|
159 |
!!y. y:B ==> d(y)=d'(y)
|
|
160 |
|] ==> case(c,d,z) = case(c',d',z)"
|
13255
|
161 |
by auto
|
13240
|
162 |
|
|
163 |
lemma case_case: "z: A+B ==>
|
13255
|
164 |
|
|
165 |
case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) =
|
13240
|
166 |
case(%x. c(c'(x)), %y. d(d'(y)), z)"
|
|
167 |
by auto
|
|
168 |
|
|
169 |
|
13356
|
170 |
subsection{*More Rules for @{term "Part(A,h)"}*}
|
13240
|
171 |
|
|
172 |
lemma Part_mono: "A<=B ==> Part(A,h)<=Part(B,h)"
|
13255
|
173 |
by blast
|
13240
|
174 |
|
|
175 |
lemma Part_Collect: "Part(Collect(A,P), h) = Collect(Part(A,h), P)"
|
13255
|
176 |
by blast
|
13240
|
177 |
|
|
178 |
lemmas Part_CollectE =
|
|
179 |
Part_Collect [THEN equalityD1, THEN subsetD, THEN CollectE, standard]
|
|
180 |
|
|
181 |
lemma Part_Inl: "Part(A+B,Inl) = {Inl(x). x: A}"
|
13255
|
182 |
by blast
|
13240
|
183 |
|
|
184 |
lemma Part_Inr: "Part(A+B,Inr) = {Inr(y). y: B}"
|
13255
|
185 |
by blast
|
13240
|
186 |
|
|
187 |
lemma PartD1: "a : Part(A,h) ==> a : A"
|
13255
|
188 |
by (simp add: Part_def)
|
13240
|
189 |
|
|
190 |
lemma Part_id: "Part(A,%x. x) = A"
|
13255
|
191 |
by blast
|
13240
|
192 |
|
|
193 |
lemma Part_Inr2: "Part(A+B, %x. Inr(h(x))) = {Inr(y). y: Part(B,h)}"
|
13255
|
194 |
by blast
|
13240
|
195 |
|
|
196 |
lemma Part_sum_equality: "C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C"
|
13255
|
197 |
by blast
|
13240
|
198 |
|
0
|
199 |
end
|