author | paulson |
Thu, 19 Aug 1999 15:11:12 +0200 | |
changeset 7283 | 5cfe2944910a |
parent 7254 | fc7f95f293da |
child 7293 | 959e060f4a2f |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/Sum.ML |
923 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
923 | 4 |
Copyright 1991 University of Cambridge |
5 |
||
5316 | 6 |
The disjoint sum of two types |
923 | 7 |
*) |
8 |
||
9 |
(** Inl_Rep and Inr_Rep: Representations of the constructors **) |
|
10 |
||
11 |
(*This counts as a non-emptiness result for admitting 'a+'b as a type*) |
|
5069 | 12 |
Goalw [Sum_def] "Inl_Rep(a) : Sum"; |
923 | 13 |
by (EVERY1 [rtac CollectI, rtac disjI1, rtac exI, rtac refl]); |
14 |
qed "Inl_RepI"; |
|
15 |
||
5069 | 16 |
Goalw [Sum_def] "Inr_Rep(b) : Sum"; |
923 | 17 |
by (EVERY1 [rtac CollectI, rtac disjI2, rtac exI, rtac refl]); |
18 |
qed "Inr_RepI"; |
|
19 |
||
5069 | 20 |
Goal "inj_on Abs_Sum Sum"; |
4830 | 21 |
by (rtac inj_on_inverseI 1); |
923 | 22 |
by (etac Abs_Sum_inverse 1); |
4830 | 23 |
qed "inj_on_Abs_Sum"; |
923 | 24 |
|
25 |
(** Distinctness of Inl and Inr **) |
|
26 |
||
5069 | 27 |
Goalw [Inl_Rep_def, Inr_Rep_def] "Inl_Rep(a) ~= Inr_Rep(b)"; |
923 | 28 |
by (EVERY1 [rtac notI, |
1465 | 29 |
etac (fun_cong RS fun_cong RS fun_cong RS iffE), |
30 |
rtac (notE RS ccontr), etac (mp RS conjunct2), |
|
31 |
REPEAT o (ares_tac [refl,conjI]) ]); |
|
923 | 32 |
qed "Inl_Rep_not_Inr_Rep"; |
33 |
||
5069 | 34 |
Goalw [Inl_def,Inr_def] "Inl(a) ~= Inr(b)"; |
4830 | 35 |
by (rtac (inj_on_Abs_Sum RS inj_on_contraD) 1); |
923 | 36 |
by (rtac Inl_Rep_not_Inr_Rep 1); |
37 |
by (rtac Inl_RepI 1); |
|
38 |
by (rtac Inr_RepI 1); |
|
39 |
qed "Inl_not_Inr"; |
|
40 |
||
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1761
diff
changeset
|
41 |
bind_thm ("Inr_not_Inl", Inl_not_Inr RS not_sym); |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1761
diff
changeset
|
42 |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1761
diff
changeset
|
43 |
AddIffs [Inl_not_Inr, Inr_not_Inl]; |
923 | 44 |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1761
diff
changeset
|
45 |
bind_thm ("Inl_neq_Inr", Inl_not_Inr RS notE); |
923 | 46 |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1761
diff
changeset
|
47 |
val Inr_neq_Inl = sym RS Inl_neq_Inr; |
923 | 48 |
|
49 |
||
50 |
(** Injectiveness of Inl and Inr **) |
|
51 |
||
5316 | 52 |
Goalw [Inl_Rep_def] "Inl_Rep(a) = Inl_Rep(c) ==> a=c"; |
53 |
by (etac (fun_cong RS fun_cong RS fun_cong RS iffE) 1); |
|
2891 | 54 |
by (Blast_tac 1); |
923 | 55 |
qed "Inl_Rep_inject"; |
56 |
||
5316 | 57 |
Goalw [Inr_Rep_def] "Inr_Rep(b) = Inr_Rep(d) ==> b=d"; |
58 |
by (etac (fun_cong RS fun_cong RS fun_cong RS iffE) 1); |
|
2891 | 59 |
by (Blast_tac 1); |
923 | 60 |
qed "Inr_Rep_inject"; |
61 |
||
5069 | 62 |
Goalw [Inl_def] "inj(Inl)"; |
923 | 63 |
by (rtac injI 1); |
4830 | 64 |
by (etac (inj_on_Abs_Sum RS inj_onD RS Inl_Rep_inject) 1); |
923 | 65 |
by (rtac Inl_RepI 1); |
66 |
by (rtac Inl_RepI 1); |
|
67 |
qed "inj_Inl"; |
|
68 |
val Inl_inject = inj_Inl RS injD; |
|
69 |
||
5069 | 70 |
Goalw [Inr_def] "inj(Inr)"; |
923 | 71 |
by (rtac injI 1); |
4830 | 72 |
by (etac (inj_on_Abs_Sum RS inj_onD RS Inr_Rep_inject) 1); |
923 | 73 |
by (rtac Inr_RepI 1); |
74 |
by (rtac Inr_RepI 1); |
|
75 |
qed "inj_Inr"; |
|
76 |
val Inr_inject = inj_Inr RS injD; |
|
77 |
||
5069 | 78 |
Goal "(Inl(x)=Inl(y)) = (x=y)"; |
4089 | 79 |
by (blast_tac (claset() addSDs [Inl_inject]) 1); |
923 | 80 |
qed "Inl_eq"; |
81 |
||
5069 | 82 |
Goal "(Inr(x)=Inr(y)) = (x=y)"; |
4089 | 83 |
by (blast_tac (claset() addSDs [Inr_inject]) 1); |
923 | 84 |
qed "Inr_eq"; |
85 |
||
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1761
diff
changeset
|
86 |
AddIffs [Inl_eq, Inr_eq]; |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1761
diff
changeset
|
87 |
|
923 | 88 |
(*** Rules for the disjoint sum of two SETS ***) |
89 |
||
90 |
(** Introduction rules for the injections **) |
|
91 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
92 |
Goalw [sum_def] "a : A ==> Inl(a) : A Plus B"; |
2891 | 93 |
by (Blast_tac 1); |
923 | 94 |
qed "InlI"; |
95 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
96 |
Goalw [sum_def] "b : B ==> Inr(b) : A Plus B"; |
2891 | 97 |
by (Blast_tac 1); |
923 | 98 |
qed "InrI"; |
99 |
||
100 |
(** Elimination rules **) |
|
101 |
||
5316 | 102 |
val major::prems = Goalw [sum_def] |
2212 | 103 |
"[| u: A Plus B; \ |
923 | 104 |
\ !!x. [| x:A; u=Inl(x) |] ==> P; \ |
105 |
\ !!y. [| y:B; u=Inr(y) |] ==> P \ |
|
106 |
\ |] ==> P"; |
|
107 |
by (rtac (major RS UnE) 1); |
|
108 |
by (REPEAT (rtac refl 1 |
|
109 |
ORELSE eresolve_tac (prems@[imageE,ssubst]) 1)); |
|
2212 | 110 |
qed "PlusE"; |
923 | 111 |
|
112 |
||
1760
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1515
diff
changeset
|
113 |
AddSIs [InlI, InrI]; |
2212 | 114 |
AddSEs [PlusE]; |
1760
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1515
diff
changeset
|
115 |
|
923 | 116 |
|
117 |
(** sum_case -- the selection operator for sums **) |
|
118 |
||
7254
fc7f95f293da
Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents:
7087
diff
changeset
|
119 |
Goalw [sum_case_def] "basic_sum_case f g (Inl x) = f(x)"; |
4535 | 120 |
by (Blast_tac 1); |
923 | 121 |
qed "sum_case_Inl"; |
122 |
||
7254
fc7f95f293da
Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents:
7087
diff
changeset
|
123 |
Goalw [sum_case_def] "basic_sum_case f g (Inr x) = g(x)"; |
4535 | 124 |
by (Blast_tac 1); |
923 | 125 |
qed "sum_case_Inr"; |
126 |
||
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1761
diff
changeset
|
127 |
Addsimps [sum_case_Inl, sum_case_Inr]; |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1761
diff
changeset
|
128 |
|
923 | 129 |
(** Exhaustion rule for sums -- a degenerate form of induction **) |
130 |
||
5316 | 131 |
val prems = Goalw [Inl_def,Inr_def] |
923 | 132 |
"[| !!x::'a. s = Inl(x) ==> P; !!y::'b. s = Inr(y) ==> P \ |
133 |
\ |] ==> P"; |
|
134 |
by (rtac (rewrite_rule [Sum_def] Rep_Sum RS CollectE) 1); |
|
135 |
by (REPEAT (eresolve_tac [disjE,exE] 1 |
|
136 |
ORELSE EVERY1 [resolve_tac prems, |
|
1465 | 137 |
etac subst, |
138 |
rtac (Rep_Sum_inverse RS sym)])); |
|
923 | 139 |
qed "sumE"; |
140 |
||
5316 | 141 |
val prems = Goal "[| !!x. P (Inl x); !!x. P (Inr x) |] ==> P x"; |
5183 | 142 |
by (res_inst_tac [("s","x")] sumE 1); |
143 |
by (ALLGOALS (hyp_subst_tac THEN' (resolve_tac prems))); |
|
144 |
qed "sum_induct"; |
|
145 |
||
7254
fc7f95f293da
Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents:
7087
diff
changeset
|
146 |
Goal "basic_sum_case (%x::'a. f(Inl x)) (%y::'b. f(Inr y)) s = f(s)"; |
923 | 147 |
by (EVERY1 [res_inst_tac [("s","s")] sumE, |
1465 | 148 |
etac ssubst, rtac sum_case_Inl, |
149 |
etac ssubst, rtac sum_case_Inr]); |
|
923 | 150 |
qed "surjective_sum"; |
151 |
||
7254
fc7f95f293da
Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents:
7087
diff
changeset
|
152 |
Goal "R(basic_sum_case f g s) = \ |
923 | 153 |
\ ((! x. s = Inl(x) --> R(f(x))) & (! y. s = Inr(y) --> R(g(y))))"; |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1761
diff
changeset
|
154 |
by (res_inst_tac [("s","s")] sumE 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4089
diff
changeset
|
155 |
by Auto_tac; |
4830 | 156 |
qed "split_sum_case"; |
923 | 157 |
|
7254
fc7f95f293da
Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents:
7087
diff
changeset
|
158 |
Goal "P (basic_sum_case f g s) = \ |
7031 | 159 |
\ (~ ((? x. s = Inl x & ~P (f x)) | (? y. s = Inr y & ~P (g y))))"; |
160 |
by (stac split_sum_case 1); |
|
161 |
by (Blast_tac 1); |
|
162 |
qed "split_sum_case_asm"; |
|
4988 | 163 |
|
923 | 164 |
(*Prevents simplification of f and g: much faster*) |
7254
fc7f95f293da
Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents:
7087
diff
changeset
|
165 |
Goal "s=t ==> basic_sum_case f g s = basic_sum_case f g t"; |
7031 | 166 |
by (etac arg_cong 1); |
167 |
qed "sum_case_weak_cong"; |
|
923 | 168 |
|
7087 | 169 |
val [p1,p2] = Goal |
7254
fc7f95f293da
Renamed sum_case to basic_sum_case and removed translations for sum_case
berghofe
parents:
7087
diff
changeset
|
170 |
"[| basic_sum_case f1 f2 = basic_sum_case g1 g2; \ |
7087 | 171 |
\ [| f1 = g1; f2 = g2 |] ==> P |] \ |
172 |
\ ==> P"; |
|
173 |
by (rtac p2 1); |
|
174 |
by (rtac ext 1); |
|
175 |
by (cut_inst_tac [("x","Inl x")] (p1 RS fun_cong) 1); |
|
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5316
diff
changeset
|
176 |
by (Asm_full_simp_tac 1); |
7087 | 177 |
by (rtac ext 1); |
178 |
by (cut_inst_tac [("x","Inr x")] (p1 RS fun_cong) 1); |
|
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5316
diff
changeset
|
179 |
by (Asm_full_simp_tac 1); |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5316
diff
changeset
|
180 |
qed "sum_case_inject"; |
923 | 181 |
|
182 |
||
183 |
(** Rules for the Part primitive **) |
|
184 |
||
5148
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
185 |
Goalw [Part_def] "[| a : A; a=h(b) |] ==> a : Part A h"; |
2891 | 186 |
by (Blast_tac 1); |
923 | 187 |
qed "Part_eqI"; |
188 |
||
189 |
val PartI = refl RSN (2,Part_eqI); |
|
190 |
||
5316 | 191 |
val major::prems = Goalw [Part_def] |
923 | 192 |
"[| a : Part A h; !!z. [| a : A; a=h(z) |] ==> P \ |
193 |
\ |] ==> P"; |
|
194 |
by (rtac (major RS IntE) 1); |
|
195 |
by (etac CollectE 1); |
|
196 |
by (etac exE 1); |
|
197 |
by (REPEAT (ares_tac prems 1)); |
|
198 |
qed "PartE"; |
|
199 |
||
2891 | 200 |
AddIs [Part_eqI]; |
201 |
AddSEs [PartE]; |
|
202 |
||
5069 | 203 |
Goalw [Part_def] "Part A h <= A"; |
923 | 204 |
by (rtac Int_lower1 1); |
205 |
qed "Part_subset"; |
|
206 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
207 |
Goal "A<=B ==> Part A h <= Part B h"; |
2922 | 208 |
by (Blast_tac 1); |
923 | 209 |
qed "Part_mono"; |
210 |
||
1515 | 211 |
val basic_monos = basic_monos @ [Part_mono]; |
212 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
213 |
Goalw [Part_def] "a : Part A h ==> a : A"; |
923 | 214 |
by (etac IntD1 1); |
215 |
qed "PartD1"; |
|
216 |
||
5069 | 217 |
Goal "Part A (%x. x) = A"; |
2891 | 218 |
by (Blast_tac 1); |
923 | 219 |
qed "Part_id"; |
220 |
||
5069 | 221 |
Goal "Part (A Int B) h = (Part A h) Int (Part B h)"; |
2922 | 222 |
by (Blast_tac 1); |
1188
0443e4dc8511
Added Part_Int and Part_Collect for inductive definitions
lcp
parents:
923
diff
changeset
|
223 |
qed "Part_Int"; |
0443e4dc8511
Added Part_Int and Part_Collect for inductive definitions
lcp
parents:
923
diff
changeset
|
224 |
|
0443e4dc8511
Added Part_Int and Part_Collect for inductive definitions
lcp
parents:
923
diff
changeset
|
225 |
(*For inductive definitions*) |
5069 | 226 |
Goal "Part (A Int {x. P x}) h = (Part A h) Int {x. P x}"; |
2922 | 227 |
by (Blast_tac 1); |
1188
0443e4dc8511
Added Part_Int and Part_Collect for inductive definitions
lcp
parents:
923
diff
changeset
|
228 |
qed "Part_Collect"; |