|
923
|
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 |
qed "ProdI";
|
|
|
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 |
qed "Pair_Rep_inject";
|
|
|
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 |
qed "inj_onto_Abs_Prod";
|
|
|
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 |
qed "Pair_inject";
|
|
|
32 |
|
|
|
33 |
goal Prod.thy "(<a,b> = <a',b'>) = (a=a' & b=b')";
|
|
|
34 |
by (fast_tac (set_cs addIs [Pair_inject]) 1);
|
|
|
35 |
qed "Pair_eq";
|
|
|
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 |
qed "fst_conv";
|
|
|
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 |
qed "snd_conv";
|
|
|
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 |
qed "PairE_lemma";
|
|
|
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 |
qed "PairE";
|
|
|
55 |
|
|
|
56 |
goalw Prod.thy [split_def] "split c <a,b> = c a b";
|
|
|
57 |
by (sstac [fst_conv, snd_conv] 1);
|
|
|
58 |
by (rtac refl 1);
|
|
|
59 |
qed "split";
|
|
|
60 |
|
|
|
61 |
val prod_ss = set_ss addsimps [fst_conv, snd_conv, split, Pair_eq];
|
|
|
62 |
|
|
|
63 |
goal Prod.thy "(s=t) = (fst(s)=fst(t) & snd(s)=snd(t))";
|
|
|
64 |
by (res_inst_tac[("p","s")] PairE 1);
|
|
|
65 |
by (res_inst_tac[("p","t")] PairE 1);
|
|
|
66 |
by (asm_simp_tac prod_ss 1);
|
|
|
67 |
qed "Pair_fst_snd_eq";
|
|
|
68 |
|
|
|
69 |
(*Prevents simplification of c: much faster*)
|
|
|
70 |
qed_goal "split_weak_cong" Prod.thy
|
|
|
71 |
"p=q ==> split c p = split c q"
|
|
|
72 |
(fn [prem] => [rtac (prem RS arg_cong) 1]);
|
|
|
73 |
|
|
|
74 |
(* Do not add as rewrite rule: invalidates some proofs in IMP *)
|
|
|
75 |
goal Prod.thy "p = <fst(p),snd(p)>";
|
|
|
76 |
by (res_inst_tac [("p","p")] PairE 1);
|
|
|
77 |
by (asm_simp_tac prod_ss 1);
|
|
|
78 |
qed "surjective_pairing";
|
|
|
79 |
|
|
|
80 |
goal Prod.thy "p = split (%x y.<x,y>) p";
|
|
|
81 |
by (res_inst_tac [("p","p")] PairE 1);
|
|
|
82 |
by (asm_simp_tac prod_ss 1);
|
|
|
83 |
qed "surjective_pairing2";
|
|
|
84 |
|
|
|
85 |
(*For use with split_tac and the simplifier*)
|
|
|
86 |
goal Prod.thy "R(split c p) = (! 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 |
qed "expand_split";
|
|
|
91 |
|
|
|
92 |
(** split used as a logical connective or set former **)
|
|
|
93 |
|
|
|
94 |
(*These rules are for use with fast_tac.
|
|
|
95 |
Could instead call simp_tac/asm_full_simp_tac using split as rewrite.*)
|
|
|
96 |
|
|
|
97 |
goal Prod.thy "!!a b c. c a b ==> split c <a,b>";
|
|
|
98 |
by (asm_simp_tac prod_ss 1);
|
|
|
99 |
qed "splitI";
|
|
|
100 |
|
|
|
101 |
val prems = goalw Prod.thy [split_def]
|
|
|
102 |
"[| split c p; !!x y. [| p = <x,y>; c x y |] ==> Q |] ==> Q";
|
|
|
103 |
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
|
|
|
104 |
qed "splitE";
|
|
|
105 |
|
|
|
106 |
goal Prod.thy "!!R a b. split R <a,b> ==> R a b";
|
|
|
107 |
by (etac (split RS iffD1) 1);
|
|
|
108 |
qed "splitD";
|
|
|
109 |
|
|
|
110 |
goal Prod.thy "!!a b c. z: c a b ==> z: split c <a,b>";
|
|
|
111 |
by (asm_simp_tac prod_ss 1);
|
|
|
112 |
qed "mem_splitI";
|
|
|
113 |
|
|
|
114 |
val prems = goalw Prod.thy [split_def]
|
|
|
115 |
"[| z: split c p; !!x y. [| p = <x,y>; z: c x y |] ==> Q |] ==> Q";
|
|
|
116 |
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
|
|
|
117 |
qed "mem_splitE";
|
|
|
118 |
|
|
|
119 |
(*** prod_fun -- action of the product functor upon functions ***)
|
|
|
120 |
|
|
|
121 |
goalw Prod.thy [prod_fun_def] "prod_fun f g <a,b> = <f(a),g(b)>";
|
|
|
122 |
by (rtac split 1);
|
|
|
123 |
qed "prod_fun";
|
|
|
124 |
|
|
|
125 |
goal Prod.thy
|
|
|
126 |
"prod_fun (f1 o f2) (g1 o g2) = ((prod_fun f1 g1) o (prod_fun f2 g2))";
|
|
|
127 |
by (rtac ext 1);
|
|
|
128 |
by (res_inst_tac [("p","x")] PairE 1);
|
|
|
129 |
by (asm_simp_tac (prod_ss addsimps [prod_fun,o_def]) 1);
|
|
|
130 |
qed "prod_fun_compose";
|
|
|
131 |
|
|
|
132 |
goal Prod.thy "prod_fun (%x.x) (%y.y) = (%z.z)";
|
|
|
133 |
by (rtac ext 1);
|
|
|
134 |
by (res_inst_tac [("p","z")] PairE 1);
|
|
|
135 |
by (asm_simp_tac (prod_ss addsimps [prod_fun]) 1);
|
|
|
136 |
qed "prod_fun_ident";
|
|
|
137 |
|
|
|
138 |
val prems = goal Prod.thy "<a,b>:r ==> <f(a),g(b)> : (prod_fun f g)``r";
|
|
|
139 |
by (rtac image_eqI 1);
|
|
|
140 |
by (rtac (prod_fun RS sym) 1);
|
|
|
141 |
by (resolve_tac prems 1);
|
|
|
142 |
qed "prod_fun_imageI";
|
|
|
143 |
|
|
|
144 |
val major::prems = goal Prod.thy
|
|
|
145 |
"[| c: (prod_fun f g)``r; !!x y. [| c=<f(x),g(y)>; <x,y>:r |] ==> P \
|
|
|
146 |
\ |] ==> P";
|
|
|
147 |
by (rtac (major RS imageE) 1);
|
|
|
148 |
by (res_inst_tac [("p","x")] PairE 1);
|
|
|
149 |
by (resolve_tac prems 1);
|
|
|
150 |
by (fast_tac HOL_cs 2);
|
|
|
151 |
by (fast_tac (HOL_cs addIs [prod_fun]) 1);
|
|
|
152 |
qed "prod_fun_imageE";
|
|
|
153 |
|
|
|
154 |
(*** Disjoint union of a family of sets - Sigma ***)
|
|
|
155 |
|
|
|
156 |
qed_goalw "SigmaI" Prod.thy [Sigma_def]
|
|
|
157 |
"[| a:A; b:B(a) |] ==> <a,b> : Sigma A B"
|
|
|
158 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]);
|
|
|
159 |
|
|
|
160 |
(*The general elimination rule*)
|
|
|
161 |
qed_goalw "SigmaE" Prod.thy [Sigma_def]
|
|
|
162 |
"[| c: Sigma A B; \
|
|
|
163 |
\ !!x y.[| x:A; y:B(x); c=<x,y> |] ==> P \
|
|
|
164 |
\ |] ==> P"
|
|
|
165 |
(fn major::prems=>
|
|
|
166 |
[ (cut_facts_tac [major] 1),
|
|
|
167 |
(REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]);
|
|
|
168 |
|
|
|
169 |
(** Elimination of <a,b>:A*B -- introduces no eigenvariables **)
|
|
|
170 |
qed_goal "SigmaD1" Prod.thy "<a,b> : Sigma A B ==> a : A"
|
|
|
171 |
(fn [major]=>
|
|
|
172 |
[ (rtac (major RS SigmaE) 1),
|
|
|
173 |
(REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
|
|
|
174 |
|
|
|
175 |
qed_goal "SigmaD2" Prod.thy "<a,b> : Sigma A B ==> b : B(a)"
|
|
|
176 |
(fn [major]=>
|
|
|
177 |
[ (rtac (major RS SigmaE) 1),
|
|
|
178 |
(REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
|
|
|
179 |
|
|
|
180 |
qed_goal "SigmaE2" Prod.thy
|
|
|
181 |
"[| <a,b> : Sigma A B; \
|
|
|
182 |
\ [| a:A; b:B(a) |] ==> P \
|
|
|
183 |
\ |] ==> P"
|
|
|
184 |
(fn [major,minor]=>
|
|
|
185 |
[ (rtac minor 1),
|
|
|
186 |
(rtac (major RS SigmaD1) 1),
|
|
|
187 |
(rtac (major RS SigmaD2) 1) ]);
|
|
|
188 |
|
|
|
189 |
(*** Domain of a relation ***)
|
|
|
190 |
|
|
|
191 |
val prems = goalw Prod.thy [image_def] "<a,b> : r ==> a : fst``r";
|
|
|
192 |
by (rtac CollectI 1);
|
|
|
193 |
by (rtac bexI 1);
|
|
|
194 |
by (rtac (fst_conv RS sym) 1);
|
|
|
195 |
by (resolve_tac prems 1);
|
|
|
196 |
qed "fst_imageI";
|
|
|
197 |
|
|
|
198 |
val major::prems = goal Prod.thy
|
|
|
199 |
"[| a : fst``r; !!y.[| <a,y> : r |] ==> P |] ==> P";
|
|
|
200 |
by (rtac (major RS imageE) 1);
|
|
|
201 |
by (resolve_tac prems 1);
|
|
|
202 |
by (etac ssubst 1);
|
|
|
203 |
by (rtac (surjective_pairing RS subst) 1);
|
|
|
204 |
by (assume_tac 1);
|
|
|
205 |
qed "fst_imageE";
|
|
|
206 |
|
|
|
207 |
(*** Range of a relation ***)
|
|
|
208 |
|
|
|
209 |
val prems = goalw Prod.thy [image_def] "<a,b> : r ==> b : snd``r";
|
|
|
210 |
by (rtac CollectI 1);
|
|
|
211 |
by (rtac bexI 1);
|
|
|
212 |
by (rtac (snd_conv RS sym) 1);
|
|
|
213 |
by (resolve_tac prems 1);
|
|
|
214 |
qed "snd_imageI";
|
|
|
215 |
|
|
|
216 |
val major::prems = goal Prod.thy
|
|
|
217 |
"[| a : snd``r; !!y.[| <y,a> : r |] ==> P |] ==> P";
|
|
|
218 |
by (rtac (major RS imageE) 1);
|
|
|
219 |
by (resolve_tac prems 1);
|
|
|
220 |
by (etac ssubst 1);
|
|
|
221 |
by (rtac (surjective_pairing RS subst) 1);
|
|
|
222 |
by (assume_tac 1);
|
|
|
223 |
qed "snd_imageE";
|
|
|
224 |
|
|
|
225 |
(** Exhaustion rule for unit -- a degenerate form of induction **)
|
|
|
226 |
|
|
|
227 |
goalw Prod.thy [Unity_def]
|
|
|
228 |
"u = Unity";
|
|
|
229 |
by (stac (rewrite_rule [Unit_def] Rep_Unit RS CollectD RS sym) 1);
|
|
|
230 |
by (rtac (Rep_Unit_inverse RS sym) 1);
|
|
|
231 |
qed "unit_eq";
|
|
|
232 |
|
|
|
233 |
val prod_cs = set_cs addSIs [SigmaI, mem_splitI]
|
|
|
234 |
addIs [fst_imageI, snd_imageI, prod_fun_imageI]
|
|
|
235 |
addSEs [SigmaE2, SigmaE, mem_splitE,
|
|
|
236 |
fst_imageE, snd_imageE, prod_fun_imageE,
|
|
|
237 |
Pair_inject];
|