author | wenzelm |
Wed, 21 Sep 1994 15:40:41 +0200 | |
changeset 145 | a9f7ff3a464c |
parent 128 | 89669c58e506 |
child 171 | 16c4ea954511 |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: HOL/prod |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1991 University of Cambridge |
|
5 |
||
6 |
For prod.thy. Ordered Pairs, the Cartesian product type, the unit type |
|
7 |
*) |
|
8 |
||
9 |
open Prod; |
|
10 |
||
11 |
(*This counts as a non-emptiness result for admitting 'a * 'b as a type*) |
|
12 |
goalw Prod.thy [Prod_def] "Pair_Rep(a,b) : Prod"; |
|
13 |
by (EVERY1 [rtac CollectI, rtac exI, rtac exI, rtac refl]); |
|
14 |
val ProdI = result(); |
|
15 |
||
16 |
val [major] = goalw Prod.thy [Pair_Rep_def] |
|
17 |
"Pair_Rep(a, b) = Pair_Rep(a',b') ==> a=a' & b=b'"; |
|
18 |
by (EVERY1 [rtac (major RS fun_cong RS fun_cong RS subst), |
|
19 |
rtac conjI, rtac refl, rtac refl]); |
|
20 |
val Pair_Rep_inject = result(); |
|
21 |
||
22 |
goal Prod.thy "inj_onto(Abs_Prod,Prod)"; |
|
23 |
by (rtac inj_onto_inverseI 1); |
|
24 |
by (etac Abs_Prod_inverse 1); |
|
25 |
val inj_onto_Abs_Prod = result(); |
|
26 |
||
27 |
val prems = goalw Prod.thy [Pair_def] |
|
28 |
"[| <a, b> = <a',b'>; [| a=a'; b=b' |] ==> R |] ==> R"; |
|
29 |
by (rtac (inj_onto_Abs_Prod RS inj_ontoD RS Pair_Rep_inject RS conjE) 1); |
|
30 |
by (REPEAT (ares_tac (prems@[ProdI]) 1)); |
|
31 |
val Pair_inject = result(); |
|
32 |
||
33 |
goal Prod.thy "(<a,b> = <a',b'>) = (a=a' & b=b')"; |
|
34 |
by (fast_tac (set_cs addIs [Pair_inject]) 1); |
|
35 |
val Pair_eq = result(); |
|
36 |
||
37 |
goalw Prod.thy [fst_def] "fst(<a,b>) = a"; |
|
38 |
by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1); |
|
39 |
val fst_conv = result(); |
|
40 |
||
41 |
goalw Prod.thy [snd_def] "snd(<a,b>) = b"; |
|
42 |
by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1); |
|
43 |
val snd_conv = result(); |
|
44 |
||
45 |
goalw Prod.thy [Pair_def] "? x y. p = <x,y>"; |
|
46 |
by (rtac (rewrite_rule [Prod_def] Rep_Prod RS CollectE) 1); |
|
47 |
by (EVERY1[etac exE, etac exE, rtac exI, rtac exI, |
|
48 |
rtac (Rep_Prod_inverse RS sym RS trans), etac arg_cong]); |
|
49 |
val PairE_lemma = result(); |
|
50 |
||
51 |
val [prem] = goal Prod.thy "[| !!x y. p = <x,y> ==> Q |] ==> Q"; |
|
52 |
by (rtac (PairE_lemma RS exE) 1); |
|
53 |
by (REPEAT (eresolve_tac [prem,exE] 1)); |
|
54 |
val PairE = result(); |
|
55 |
||
112 | 56 |
goalw Prod.thy [split_def] "split(c, <a,b>) = c(a,b)"; |
0 | 57 |
by (sstac [fst_conv, snd_conv] 1); |
58 |
by (rtac refl 1); |
|
59 |
val split = result(); |
|
60 |
||
26 | 61 |
val pair_ss = set_ss addsimps [fst_conv, snd_conv, split, Pair_eq]; |
0 | 62 |
|
31 | 63 |
goal Prod.thy "(s=t) = (fst(s)=fst(t) & snd(s)=snd(t))"; |
73 | 64 |
by (res_inst_tac[("p","s")] PairE 1); |
65 |
by (res_inst_tac[("p","t")] PairE 1); |
|
66 |
by (asm_simp_tac pair_ss 1); |
|
67 |
val Pair_fst_snd_eq = result(); |
|
31 | 68 |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
69 |
(*Prevents simplification of c: much faster*) |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
70 |
val split_weak_cong = prove_goal Prod.thy |
112 | 71 |
"p=q ==> split(c,p) = split(c,q)" |
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
72 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
73 |
|
0 | 74 |
goal Prod.thy "p = <fst(p),snd(p)>"; |
75 |
by (res_inst_tac [("p","p")] PairE 1); |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
76 |
by (asm_simp_tac pair_ss 1); |
0 | 77 |
val surjective_pairing = result(); |
78 |
||
112 | 79 |
goal Prod.thy "p = split(%x y.<x,y>, p)"; |
0 | 80 |
by (res_inst_tac [("p","p")] PairE 1); |
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
81 |
by (asm_simp_tac pair_ss 1); |
0 | 82 |
val surjective_pairing2 = result(); |
83 |
||
112 | 84 |
(*For use with split_tac and the simplifier*) |
85 |
goal Prod.thy "R(split(c,p)) = (! x y. p = <x,y> --> R(c(x,y)))"; |
|
0 | 86 |
by (stac surjective_pairing 1); |
87 |
by (stac split 1); |
|
88 |
by (fast_tac (HOL_cs addSEs [Pair_inject]) 1); |
|
89 |
val expand_split = result(); |
|
90 |
||
112 | 91 |
(** split used as a logical connective or set former **) |
92 |
||
93 |
(*These rules are for use with fast_tac. |
|
94 |
Could instead call simp_tac/asm_full_simp_tac using split as rewrite.*) |
|
95 |
||
96 |
goal Prod.thy "!!a b c. c(a,b) ==> split(c, <a,b>)"; |
|
97 |
by (asm_simp_tac pair_ss 1); |
|
98 |
val splitI = result(); |
|
99 |
||
100 |
val prems = goalw Prod.thy [split_def] |
|
101 |
"[| split(c,p); !!x y. [| p = <x,y>; c(x,y) |] ==> Q |] ==> Q"; |
|
102 |
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); |
|
103 |
val splitE = result(); |
|
104 |
||
128 | 105 |
goal Prod.thy "!!R a b. split(R,<a,b>) ==> R(a,b)"; |
106 |
by (etac (split RS iffD1) 1); |
|
107 |
val splitD = result(); |
|
108 |
||
112 | 109 |
goal Prod.thy "!!a b c. z: c(a,b) ==> z: split(c, <a,b>)"; |
110 |
by (asm_simp_tac pair_ss 1); |
|
111 |
val mem_splitI = result(); |
|
112 |
||
113 |
val prems = goalw Prod.thy [split_def] |
|
114 |
"[| z: split(c,p); !!x y. [| p = <x,y>; z: c(x,y) |] ==> Q |] ==> Q"; |
|
115 |
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); |
|
116 |
val mem_splitE = result(); |
|
117 |
||
0 | 118 |
(*** prod_fun -- action of the product functor upon functions ***) |
119 |
||
120 |
goalw Prod.thy [prod_fun_def] "prod_fun(f,g,<a,b>) = <f(a),g(b)>"; |
|
121 |
by (rtac split 1); |
|
122 |
val prod_fun = result(); |
|
123 |
||
124 |
goal Prod.thy |
|
125 |
"prod_fun(f1 o f2, g1 o g2) = (prod_fun(f1,g1) o prod_fun(f2,g2))"; |
|
126 |
by (rtac ext 1); |
|
127 |
by (res_inst_tac [("p","x")] PairE 1); |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
128 |
by (asm_simp_tac (pair_ss addsimps [prod_fun,o_def]) 1); |
0 | 129 |
val prod_fun_compose = result(); |
130 |
||
131 |
goal Prod.thy "prod_fun(%x.x, %y.y) = (%z.z)"; |
|
132 |
by (rtac ext 1); |
|
133 |
by (res_inst_tac [("p","z")] PairE 1); |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
134 |
by (asm_simp_tac (pair_ss addsimps [prod_fun]) 1); |
0 | 135 |
val prod_fun_ident = result(); |
136 |
||
137 |
val prems = goal Prod.thy "<a,b>:r ==> <f(a),g(b)> : prod_fun(f,g)``r"; |
|
138 |
by (rtac image_eqI 1); |
|
139 |
by (rtac (prod_fun RS sym) 1); |
|
140 |
by (resolve_tac prems 1); |
|
141 |
val prod_fun_imageI = result(); |
|
142 |
||
143 |
val major::prems = goal Prod.thy |
|
144 |
"[| c: prod_fun(f,g)``r; !!x y. [| c=<f(x),g(y)>; <x,y>:r |] ==> P \ |
|
145 |
\ |] ==> P"; |
|
146 |
by (rtac (major RS imageE) 1); |
|
147 |
by (res_inst_tac [("p","x")] PairE 1); |
|
148 |
by (resolve_tac prems 1); |
|
149 |
by (fast_tac HOL_cs 2); |
|
150 |
by (fast_tac (HOL_cs addIs [prod_fun]) 1); |
|
151 |
val prod_fun_imageE = result(); |
|
152 |
||
153 |
(*** Disjoint union of a family of sets - Sigma ***) |
|
154 |
||
155 |
val SigmaI = prove_goalw Prod.thy [Sigma_def] |
|
156 |
"[| a:A; b:B(a) |] ==> <a,b> : Sigma(A,B)" |
|
157 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]); |
|
158 |
||
159 |
(*The general elimination rule*) |
|
160 |
val SigmaE = prove_goalw Prod.thy [Sigma_def] |
|
161 |
"[| c: Sigma(A,B); \ |
|
162 |
\ !!x y.[| x:A; y:B(x); c=<x,y> |] ==> P \ |
|
163 |
\ |] ==> P" |
|
164 |
(fn major::prems=> |
|
165 |
[ (cut_facts_tac [major] 1), |
|
166 |
(REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); |
|
167 |
||
168 |
(** Elimination of <a,b>:A*B -- introduces no eigenvariables **) |
|
169 |
val SigmaD1 = prove_goal Prod.thy "<a,b> : Sigma(A,B) ==> a : A" |
|
170 |
(fn [major]=> |
|
171 |
[ (rtac (major RS SigmaE) 1), |
|
172 |
(REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); |
|
173 |
||
174 |
val SigmaD2 = prove_goal Prod.thy "<a,b> : Sigma(A,B) ==> b : B(a)" |
|
175 |
(fn [major]=> |
|
176 |
[ (rtac (major RS SigmaE) 1), |
|
177 |
(REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); |
|
178 |
||
179 |
val SigmaE2 = prove_goal Prod.thy |
|
180 |
"[| <a,b> : Sigma(A,B); \ |
|
181 |
\ [| a:A; b:B(a) |] ==> P \ |
|
182 |
\ |] ==> P" |
|
183 |
(fn [major,minor]=> |
|
184 |
[ (rtac minor 1), |
|
185 |
(rtac (major RS SigmaD1) 1), |
|
186 |
(rtac (major RS SigmaD2) 1) ]); |
|
187 |
||
188 |
(*** Domain of a relation ***) |
|
189 |
||
190 |
val prems = goalw Prod.thy [image_def] "<a,b> : r ==> a : fst``r"; |
|
191 |
by (rtac CollectI 1); |
|
192 |
by (rtac bexI 1); |
|
193 |
by (rtac (fst_conv RS sym) 1); |
|
194 |
by (resolve_tac prems 1); |
|
195 |
val fst_imageI = result(); |
|
196 |
||
197 |
val major::prems = goal Prod.thy |
|
198 |
"[| a : fst``r; !!y.[| <a,y> : r |] ==> P |] ==> P"; |
|
199 |
by (rtac (major RS imageE) 1); |
|
200 |
by (resolve_tac prems 1); |
|
201 |
by (etac ssubst 1); |
|
202 |
by (rtac (surjective_pairing RS subst) 1); |
|
203 |
by (assume_tac 1); |
|
204 |
val fst_imageE = result(); |
|
205 |
||
206 |
(*** Range of a relation ***) |
|
207 |
||
208 |
val prems = goalw Prod.thy [image_def] "<a,b> : r ==> b : snd``r"; |
|
209 |
by (rtac CollectI 1); |
|
210 |
by (rtac bexI 1); |
|
211 |
by (rtac (snd_conv RS sym) 1); |
|
212 |
by (resolve_tac prems 1); |
|
213 |
val snd_imageI = result(); |
|
214 |
||
215 |
val major::prems = goal Prod.thy |
|
216 |
"[| a : snd``r; !!y.[| <y,a> : r |] ==> P |] ==> P"; |
|
217 |
by (rtac (major RS imageE) 1); |
|
218 |
by (resolve_tac prems 1); |
|
219 |
by (etac ssubst 1); |
|
220 |
by (rtac (surjective_pairing RS subst) 1); |
|
221 |
by (assume_tac 1); |
|
222 |
val snd_imageE = result(); |
|
223 |
||
224 |
(** Exhaustion rule for unit -- a degenerate form of induction **) |
|
225 |
||
226 |
goalw Prod.thy [Unity_def] |
|
227 |
"u = Unity"; |
|
228 |
by (stac (rewrite_rule [Unit_def] Rep_Unit RS CollectD RS sym) 1); |
|
229 |
by (rtac (Rep_Unit_inverse RS sym) 1); |
|
230 |
val unit_eq = result(); |
|
112 | 231 |
|
232 |
val prod_cs = set_cs addSIs [SigmaI, mem_splitI] |
|
233 |
addIs [fst_imageI, snd_imageI, prod_fun_imageI] |
|
234 |
addSEs [SigmaE2, SigmaE, mem_splitE, |
|
235 |
fst_imageE, snd_imageE, prod_fun_imageE, |
|
236 |
Pair_inject]; |
|
237 |
||
238 |
val prod_ss = pair_ss; |