|
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 |
|
56 goalw Prod.thy [split_def] "split(<a,b>, c) = c(a,b)"; |
|
57 by (sstac [fst_conv, snd_conv] 1); |
|
58 by (rtac refl 1); |
|
59 val split = result(); |
|
60 |
|
61 (*FIXME: split's congruence rule should only simplifies the pair*) |
|
62 val pair_ss = set_ss addsimps [fst_conv, snd_conv, split]; |
|
63 |
|
64 goal Prod.thy "p = <fst(p),snd(p)>"; |
|
65 by (res_inst_tac [("p","p")] PairE 1); |
|
66 by(asm_simp_tac pair_ss 1); |
|
67 val surjective_pairing = result(); |
|
68 |
|
69 goal Prod.thy "p = split(p, %x y.<x,y>)"; |
|
70 by (res_inst_tac [("p","p")] PairE 1); |
|
71 by(asm_simp_tac pair_ss 1); |
|
72 val surjective_pairing2 = result(); |
|
73 |
|
74 (** split used as a logical connective, with result type bool **) |
|
75 |
|
76 val prems = goal Prod.thy "c(a,b) ==> split(<a,b>, c)"; |
|
77 by (stac split 1); |
|
78 by (resolve_tac prems 1); |
|
79 val splitI = result(); |
|
80 |
|
81 val prems = goalw Prod.thy [split_def] |
|
82 "[| split(p,c); !!x y. [| p = <x,y>; c(x,y) |] ==> Q |] ==> Q"; |
|
83 by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); |
|
84 val splitE = result(); |
|
85 |
|
86 goal Prod.thy "R(split(p,c)) = (! x y. p = <x,y> --> R(c(x,y)))"; |
|
87 by (stac surjective_pairing 1); |
|
88 by (stac split 1); |
|
89 by (fast_tac (HOL_cs addSEs [Pair_inject]) 1); |
|
90 val expand_split = result(); |
|
91 |
|
92 (*** prod_fun -- action of the product functor upon functions ***) |
|
93 |
|
94 goalw Prod.thy [prod_fun_def] "prod_fun(f,g,<a,b>) = <f(a),g(b)>"; |
|
95 by (rtac split 1); |
|
96 val prod_fun = result(); |
|
97 |
|
98 goal Prod.thy |
|
99 "prod_fun(f1 o f2, g1 o g2) = (prod_fun(f1,g1) o prod_fun(f2,g2))"; |
|
100 by (rtac ext 1); |
|
101 by (res_inst_tac [("p","x")] PairE 1); |
|
102 by(asm_simp_tac (pair_ss addsimps [prod_fun,o_def]) 1); |
|
103 val prod_fun_compose = result(); |
|
104 |
|
105 goal Prod.thy "prod_fun(%x.x, %y.y) = (%z.z)"; |
|
106 by (rtac ext 1); |
|
107 by (res_inst_tac [("p","z")] PairE 1); |
|
108 by(asm_simp_tac (pair_ss addsimps [prod_fun]) 1); |
|
109 val prod_fun_ident = result(); |
|
110 |
|
111 val prems = goal Prod.thy "<a,b>:r ==> <f(a),g(b)> : prod_fun(f,g)``r"; |
|
112 by (rtac image_eqI 1); |
|
113 by (rtac (prod_fun RS sym) 1); |
|
114 by (resolve_tac prems 1); |
|
115 val prod_fun_imageI = result(); |
|
116 |
|
117 val major::prems = goal Prod.thy |
|
118 "[| c: prod_fun(f,g)``r; !!x y. [| c=<f(x),g(y)>; <x,y>:r |] ==> P \ |
|
119 \ |] ==> P"; |
|
120 by (rtac (major RS imageE) 1); |
|
121 by (res_inst_tac [("p","x")] PairE 1); |
|
122 by (resolve_tac prems 1); |
|
123 by (fast_tac HOL_cs 2); |
|
124 by (fast_tac (HOL_cs addIs [prod_fun]) 1); |
|
125 val prod_fun_imageE = result(); |
|
126 |
|
127 (*** Disjoint union of a family of sets - Sigma ***) |
|
128 |
|
129 val SigmaI = prove_goalw Prod.thy [Sigma_def] |
|
130 "[| a:A; b:B(a) |] ==> <a,b> : Sigma(A,B)" |
|
131 (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]); |
|
132 |
|
133 (*The general elimination rule*) |
|
134 val SigmaE = prove_goalw Prod.thy [Sigma_def] |
|
135 "[| c: Sigma(A,B); \ |
|
136 \ !!x y.[| x:A; y:B(x); c=<x,y> |] ==> P \ |
|
137 \ |] ==> P" |
|
138 (fn major::prems=> |
|
139 [ (cut_facts_tac [major] 1), |
|
140 (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); |
|
141 |
|
142 (** Elimination of <a,b>:A*B -- introduces no eigenvariables **) |
|
143 val SigmaD1 = prove_goal Prod.thy "<a,b> : Sigma(A,B) ==> a : A" |
|
144 (fn [major]=> |
|
145 [ (rtac (major RS SigmaE) 1), |
|
146 (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); |
|
147 |
|
148 val SigmaD2 = prove_goal Prod.thy "<a,b> : Sigma(A,B) ==> b : B(a)" |
|
149 (fn [major]=> |
|
150 [ (rtac (major RS SigmaE) 1), |
|
151 (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); |
|
152 |
|
153 val SigmaE2 = prove_goal Prod.thy |
|
154 "[| <a,b> : Sigma(A,B); \ |
|
155 \ [| a:A; b:B(a) |] ==> P \ |
|
156 \ |] ==> P" |
|
157 (fn [major,minor]=> |
|
158 [ (rtac minor 1), |
|
159 (rtac (major RS SigmaD1) 1), |
|
160 (rtac (major RS SigmaD2) 1) ]); |
|
161 |
|
162 (*** Domain of a relation ***) |
|
163 |
|
164 val prems = goalw Prod.thy [image_def] "<a,b> : r ==> a : fst``r"; |
|
165 by (rtac CollectI 1); |
|
166 by (rtac bexI 1); |
|
167 by (rtac (fst_conv RS sym) 1); |
|
168 by (resolve_tac prems 1); |
|
169 val fst_imageI = result(); |
|
170 |
|
171 val major::prems = goal Prod.thy |
|
172 "[| a : fst``r; !!y.[| <a,y> : r |] ==> P |] ==> P"; |
|
173 by (rtac (major RS imageE) 1); |
|
174 by (resolve_tac prems 1); |
|
175 by (etac ssubst 1); |
|
176 by (rtac (surjective_pairing RS subst) 1); |
|
177 by (assume_tac 1); |
|
178 val fst_imageE = result(); |
|
179 |
|
180 (*** Range of a relation ***) |
|
181 |
|
182 val prems = goalw Prod.thy [image_def] "<a,b> : r ==> b : snd``r"; |
|
183 by (rtac CollectI 1); |
|
184 by (rtac bexI 1); |
|
185 by (rtac (snd_conv RS sym) 1); |
|
186 by (resolve_tac prems 1); |
|
187 val snd_imageI = result(); |
|
188 |
|
189 val major::prems = goal Prod.thy |
|
190 "[| a : snd``r; !!y.[| <y,a> : r |] ==> P |] ==> P"; |
|
191 by (rtac (major RS imageE) 1); |
|
192 by (resolve_tac prems 1); |
|
193 by (etac ssubst 1); |
|
194 by (rtac (surjective_pairing RS subst) 1); |
|
195 by (assume_tac 1); |
|
196 val snd_imageE = result(); |
|
197 |
|
198 (** Exhaustion rule for unit -- a degenerate form of induction **) |
|
199 |
|
200 goalw Prod.thy [Unity_def] |
|
201 "u = Unity"; |
|
202 by (stac (rewrite_rule [Unit_def] Rep_Unit RS CollectD RS sym) 1); |
|
203 by (rtac (Rep_Unit_inverse RS sym) 1); |
|
204 val unit_eq = result(); |