author | wenzelm |
Fri, 25 Apr 1997 17:50:55 +0200 | |
changeset 3061 | 25b2a895f864 |
parent 2935 | 998cb95fdd43 |
child 3429 | 160f18a686b5 |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/prod |
923 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
923 | 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), |
|
1465 | 19 |
rtac conjI, rtac refl, rtac refl]); |
923 | 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] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
28 |
"[| (a, b) = (a',b'); [| a=a'; b=b' |] ==> R |] ==> R"; |
923 | 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 |
||
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
33 |
AddSEs [Pair_inject]; |
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
34 |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
35 |
goal Prod.thy "((a,b) = (a',b')) = (a=a' & b=b')"; |
2935 | 36 |
by (Blast_tac 1); |
923 | 37 |
qed "Pair_eq"; |
38 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
39 |
goalw Prod.thy [fst_def] "fst((a,b)) = a"; |
2935 | 40 |
by (blast_tac (!claset addIs [select_equality]) 1); |
923 | 41 |
qed "fst_conv"; |
42 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
43 |
goalw Prod.thy [snd_def] "snd((a,b)) = b"; |
2935 | 44 |
by (blast_tac (!claset addIs [select_equality]) 1); |
923 | 45 |
qed "snd_conv"; |
46 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
47 |
goalw Prod.thy [Pair_def] "? x y. p = (x,y)"; |
923 | 48 |
by (rtac (rewrite_rule [Prod_def] Rep_Prod RS CollectE) 1); |
49 |
by (EVERY1[etac exE, etac exE, rtac exI, rtac exI, |
|
1465 | 50 |
rtac (Rep_Prod_inverse RS sym RS trans), etac arg_cong]); |
923 | 51 |
qed "PairE_lemma"; |
52 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
53 |
val [prem] = goal Prod.thy "[| !!x y. p = (x,y) ==> Q |] ==> Q"; |
923 | 54 |
by (rtac (PairE_lemma RS exE) 1); |
55 |
by (REPEAT (eresolve_tac [prem,exE] 1)); |
|
56 |
qed "PairE"; |
|
57 |
||
1301 | 58 |
(* replace parameters of product type by individual component parameters *) |
59 |
local |
|
60 |
fun is_pair (_,Type("*",_)) = true |
|
61 |
| is_pair _ = false; |
|
62 |
||
1727 | 63 |
fun find_pair_param prem = |
64 |
let val params = Logic.strip_params prem |
|
1301 | 65 |
in if exists is_pair params |
1727 | 66 |
then let val params = rev(rename_wrt_term prem params) |
1301 | 67 |
(*as they are printed*) |
68 |
in apsome fst (find_first is_pair params) end |
|
69 |
else None |
|
70 |
end; |
|
71 |
||
72 |
in |
|
73 |
||
1727 | 74 |
val split_all_tac = REPEAT o SUBGOAL (fn (prem,i) => |
75 |
case find_pair_param prem of |
|
1301 | 76 |
None => no_tac |
1727 | 77 |
| Some x => EVERY[res_inst_tac[("p",x)] PairE i, |
78 |
REPEAT(hyp_subst_tac i), prune_params_tac]); |
|
1301 | 79 |
|
80 |
end; |
|
81 |
||
82 |
goal Prod.thy "(!x. P x) = (!a b. P(a,b))"; |
|
2637
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents:
2089
diff
changeset
|
83 |
by (fast_tac (!claset addbefore split_all_tac) 1); |
1301 | 84 |
qed "split_paired_All"; |
85 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
86 |
goalw Prod.thy [split_def] "split c (a,b) = c a b"; |
1485
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
87 |
by (EVERY1[stac fst_conv, stac snd_conv]); |
923 | 88 |
by (rtac refl 1); |
89 |
qed "split"; |
|
90 |
||
1301 | 91 |
Addsimps [fst_conv, snd_conv, split_paired_All, split, Pair_eq]; |
923 | 92 |
|
93 |
goal Prod.thy "(s=t) = (fst(s)=fst(t) & snd(s)=snd(t))"; |
|
94 |
by (res_inst_tac[("p","s")] PairE 1); |
|
95 |
by (res_inst_tac[("p","t")] PairE 1); |
|
1264 | 96 |
by (Asm_simp_tac 1); |
923 | 97 |
qed "Pair_fst_snd_eq"; |
98 |
||
99 |
(*Prevents simplification of c: much faster*) |
|
100 |
qed_goal "split_weak_cong" Prod.thy |
|
101 |
"p=q ==> split c p = split c q" |
|
102 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
|
103 |
||
104 |
(* Do not add as rewrite rule: invalidates some proofs in IMP *) |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
105 |
goal Prod.thy "p = (fst(p),snd(p))"; |
923 | 106 |
by (res_inst_tac [("p","p")] PairE 1); |
1264 | 107 |
by (Asm_simp_tac 1); |
923 | 108 |
qed "surjective_pairing"; |
109 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
110 |
goal Prod.thy "p = split (%x y.(x,y)) p"; |
923 | 111 |
by (res_inst_tac [("p","p")] PairE 1); |
1264 | 112 |
by (Asm_simp_tac 1); |
923 | 113 |
qed "surjective_pairing2"; |
114 |
||
1655 | 115 |
qed_goal "split_eta" Prod.thy "(%(x,y). f(x,y)) = f" |
116 |
(fn _ => [rtac ext 1, split_all_tac 1, rtac split 1]); |
|
117 |
||
923 | 118 |
(*For use with split_tac and the simplifier*) |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
119 |
goal Prod.thy "R(split c p) = (! x y. p = (x,y) --> R(c x y))"; |
923 | 120 |
by (stac surjective_pairing 1); |
121 |
by (stac split 1); |
|
2935 | 122 |
by (Blast_tac 1); |
923 | 123 |
qed "expand_split"; |
124 |
||
125 |
(** split used as a logical connective or set former **) |
|
126 |
||
2935 | 127 |
(*These rules are for use with blast_tac. |
923 | 128 |
Could instead call simp_tac/asm_full_simp_tac using split as rewrite.*) |
129 |
||
1454
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
130 |
goal Prod.thy "!!p. [| !!a b. p=(a,b) ==> c a b |] ==> split c p"; |
1552 | 131 |
by (split_all_tac 1); |
1454
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
132 |
by (Asm_simp_tac 1); |
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
133 |
qed "splitI2"; |
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
134 |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
135 |
goal Prod.thy "!!a b c. c a b ==> split c (a,b)"; |
1264 | 136 |
by (Asm_simp_tac 1); |
923 | 137 |
qed "splitI"; |
138 |
||
139 |
val prems = goalw Prod.thy [split_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
140 |
"[| split c p; !!x y. [| p = (x,y); c x y |] ==> Q |] ==> Q"; |
923 | 141 |
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); |
142 |
qed "splitE"; |
|
143 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
144 |
goal Prod.thy "!!R a b. split R (a,b) ==> R a b"; |
923 | 145 |
by (etac (split RS iffD1) 1); |
146 |
qed "splitD"; |
|
147 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
148 |
goal Prod.thy "!!a b c. z: c a b ==> z: split c (a,b)"; |
1264 | 149 |
by (Asm_simp_tac 1); |
923 | 150 |
qed "mem_splitI"; |
151 |
||
1454
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
152 |
goal Prod.thy "!!p. [| !!a b. p=(a,b) ==> z: c a b |] ==> z: split c p"; |
1552 | 153 |
by (split_all_tac 1); |
1454
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
154 |
by (Asm_simp_tac 1); |
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
155 |
qed "mem_splitI2"; |
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
156 |
|
923 | 157 |
val prems = goalw Prod.thy [split_def] |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
158 |
"[| z: split c p; !!x y. [| p = (x,y); z: c x y |] ==> Q |] ==> Q"; |
923 | 159 |
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); |
160 |
qed "mem_splitE"; |
|
161 |
||
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
162 |
AddSIs [splitI, splitI2, mem_splitI, mem_splitI2]; |
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
163 |
AddSEs [splitE, mem_splitE]; |
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
164 |
|
923 | 165 |
(*** prod_fun -- action of the product functor upon functions ***) |
166 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
167 |
goalw Prod.thy [prod_fun_def] "prod_fun f g (a,b) = (f(a),g(b))"; |
923 | 168 |
by (rtac split 1); |
169 |
qed "prod_fun"; |
|
170 |
||
171 |
goal Prod.thy |
|
172 |
"prod_fun (f1 o f2) (g1 o g2) = ((prod_fun f1 g1) o (prod_fun f2 g2))"; |
|
173 |
by (rtac ext 1); |
|
174 |
by (res_inst_tac [("p","x")] PairE 1); |
|
1264 | 175 |
by (asm_simp_tac (!simpset addsimps [prod_fun,o_def]) 1); |
923 | 176 |
qed "prod_fun_compose"; |
177 |
||
178 |
goal Prod.thy "prod_fun (%x.x) (%y.y) = (%z.z)"; |
|
179 |
by (rtac ext 1); |
|
180 |
by (res_inst_tac [("p","z")] PairE 1); |
|
1264 | 181 |
by (asm_simp_tac (!simpset addsimps [prod_fun]) 1); |
923 | 182 |
qed "prod_fun_ident"; |
183 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
184 |
val prems = goal Prod.thy "(a,b):r ==> (f(a),g(b)) : (prod_fun f g)``r"; |
923 | 185 |
by (rtac image_eqI 1); |
186 |
by (rtac (prod_fun RS sym) 1); |
|
187 |
by (resolve_tac prems 1); |
|
188 |
qed "prod_fun_imageI"; |
|
189 |
||
190 |
val major::prems = goal Prod.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
191 |
"[| c: (prod_fun f g)``r; !!x y. [| c=(f(x),g(y)); (x,y):r |] ==> P \ |
923 | 192 |
\ |] ==> P"; |
193 |
by (rtac (major RS imageE) 1); |
|
194 |
by (res_inst_tac [("p","x")] PairE 1); |
|
195 |
by (resolve_tac prems 1); |
|
2935 | 196 |
by (Blast_tac 2); |
197 |
by (blast_tac (!claset addIs [prod_fun]) 1); |
|
923 | 198 |
qed "prod_fun_imageE"; |
199 |
||
200 |
(*** Disjoint union of a family of sets - Sigma ***) |
|
201 |
||
202 |
qed_goalw "SigmaI" Prod.thy [Sigma_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
203 |
"[| a:A; b:B(a) |] ==> (a,b) : Sigma A B" |
923 | 204 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]); |
205 |
||
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
206 |
AddSIs [SigmaI]; |
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
207 |
|
923 | 208 |
(*The general elimination rule*) |
209 |
qed_goalw "SigmaE" Prod.thy [Sigma_def] |
|
210 |
"[| c: Sigma A B; \ |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
211 |
\ !!x y.[| x:A; y:B(x); c=(x,y) |] ==> P \ |
923 | 212 |
\ |] ==> P" |
213 |
(fn major::prems=> |
|
214 |
[ (cut_facts_tac [major] 1), |
|
215 |
(REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); |
|
216 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
217 |
(** Elimination of (a,b):A*B -- introduces no eigenvariables **) |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
218 |
qed_goal "SigmaD1" Prod.thy "(a,b) : Sigma A B ==> a : A" |
923 | 219 |
(fn [major]=> |
220 |
[ (rtac (major RS SigmaE) 1), |
|
221 |
(REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); |
|
222 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
223 |
qed_goal "SigmaD2" Prod.thy "(a,b) : Sigma A B ==> b : B(a)" |
923 | 224 |
(fn [major]=> |
225 |
[ (rtac (major RS SigmaE) 1), |
|
226 |
(REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); |
|
227 |
||
228 |
qed_goal "SigmaE2" Prod.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
229 |
"[| (a,b) : Sigma A B; \ |
923 | 230 |
\ [| a:A; b:B(a) |] ==> P \ |
231 |
\ |] ==> P" |
|
232 |
(fn [major,minor]=> |
|
233 |
[ (rtac minor 1), |
|
234 |
(rtac (major RS SigmaD1) 1), |
|
235 |
(rtac (major RS SigmaD2) 1) ]); |
|
236 |
||
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
237 |
AddSEs [SigmaE2, SigmaE]; |
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
238 |
|
1515 | 239 |
val prems = goal Prod.thy |
1642 | 240 |
"[| A<=C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D"; |
1515 | 241 |
by (cut_facts_tac prems 1); |
2935 | 242 |
by (blast_tac (!claset addIs (prems RL [subsetD])) 1); |
1515 | 243 |
qed "Sigma_mono"; |
244 |
||
1618 | 245 |
qed_goal "Sigma_empty1" Prod.thy "Sigma {} B = {}" |
2935 | 246 |
(fn _ => [ (Blast_tac 1) ]); |
1618 | 247 |
|
1642 | 248 |
qed_goal "Sigma_empty2" Prod.thy "A Times {} = {}" |
2935 | 249 |
(fn _ => [ (Blast_tac 1) ]); |
1618 | 250 |
|
251 |
Addsimps [Sigma_empty1,Sigma_empty2]; |
|
252 |
||
253 |
goal Prod.thy "((a,b): Sigma A B) = (a:A & b:B(a))"; |
|
2935 | 254 |
by (Blast_tac 1); |
1618 | 255 |
qed "mem_Sigma_iff"; |
256 |
Addsimps [mem_Sigma_iff]; |
|
257 |
||
1515 | 258 |
|
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
259 |
(*Suggested by Pierre Chartier*) |
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
260 |
goal Prod.thy |
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
261 |
"(UN (a,b):(A Times B). E a Times F b) = (UNION A E) Times (UNION B F)"; |
2935 | 262 |
by (Blast_tac 1); |
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
263 |
qed "UNION_Times_distrib"; |
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
264 |
|
923 | 265 |
(*** Domain of a relation ***) |
266 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
267 |
val prems = goalw Prod.thy [image_def] "(a,b) : r ==> a : fst``r"; |
923 | 268 |
by (rtac CollectI 1); |
269 |
by (rtac bexI 1); |
|
270 |
by (rtac (fst_conv RS sym) 1); |
|
271 |
by (resolve_tac prems 1); |
|
272 |
qed "fst_imageI"; |
|
273 |
||
274 |
val major::prems = goal Prod.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
275 |
"[| a : fst``r; !!y.[| (a,y) : r |] ==> P |] ==> P"; |
923 | 276 |
by (rtac (major RS imageE) 1); |
277 |
by (resolve_tac prems 1); |
|
278 |
by (etac ssubst 1); |
|
279 |
by (rtac (surjective_pairing RS subst) 1); |
|
280 |
by (assume_tac 1); |
|
281 |
qed "fst_imageE"; |
|
282 |
||
283 |
(*** Range of a relation ***) |
|
284 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
285 |
val prems = goalw Prod.thy [image_def] "(a,b) : r ==> b : snd``r"; |
923 | 286 |
by (rtac CollectI 1); |
287 |
by (rtac bexI 1); |
|
288 |
by (rtac (snd_conv RS sym) 1); |
|
289 |
by (resolve_tac prems 1); |
|
290 |
qed "snd_imageI"; |
|
291 |
||
292 |
val major::prems = goal Prod.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
293 |
"[| a : snd``r; !!y.[| (y,a) : r |] ==> P |] ==> P"; |
923 | 294 |
by (rtac (major RS imageE) 1); |
295 |
by (resolve_tac prems 1); |
|
296 |
by (etac ssubst 1); |
|
297 |
by (rtac (surjective_pairing RS subst) 1); |
|
298 |
by (assume_tac 1); |
|
299 |
qed "snd_imageE"; |
|
300 |
||
301 |
(** Exhaustion rule for unit -- a degenerate form of induction **) |
|
302 |
||
303 |
goalw Prod.thy [Unity_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
304 |
"u = ()"; |
2886 | 305 |
by (stac (rewrite_rule [unit_def] Rep_unit RS singletonD RS sym) 1); |
2880 | 306 |
by (rtac (Rep_unit_inverse RS sym) 1); |
923 | 307 |
qed "unit_eq"; |
1754
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1746
diff
changeset
|
308 |
|
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1746
diff
changeset
|
309 |
AddIs [fst_imageI, snd_imageI, prod_fun_imageI]; |
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
310 |
AddSEs [fst_imageE, snd_imageE, prod_fun_imageE]; |
923 | 311 |
|
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
312 |
structure Prod_Syntax = |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
313 |
struct |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
314 |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
315 |
val unitT = Type("unit",[]); |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
316 |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
317 |
fun mk_prod (T1,T2) = Type("*", [T1,T2]); |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
318 |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
319 |
(*Maps the type T1*...*Tn to [T1,...,Tn], however nested*) |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
320 |
fun factors (Type("*", [T1,T2])) = factors T1 @ factors T2 |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
321 |
| factors T = [T]; |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
322 |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
323 |
(*Make a correctly typed ordered pair*) |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
324 |
fun mk_Pair (t1,t2) = |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
325 |
let val T1 = fastype_of t1 |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
326 |
and T2 = fastype_of t2 |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
327 |
in Const("Pair", [T1, T2] ---> mk_prod(T1,T2)) $ t1 $ t2 end; |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
328 |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
329 |
fun split_const(Ta,Tb,Tc) = |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
330 |
Const("split", [[Ta,Tb]--->Tc, mk_prod(Ta,Tb)] ---> Tc); |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
331 |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
332 |
(*In ap_split S T u, term u expects separate arguments for the factors of S, |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
333 |
with result type T. The call creates a new term expecting one argument |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
334 |
of type S.*) |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
335 |
fun ap_split (Type("*", [T1,T2])) T3 u = |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
336 |
split_const(T1,T2,T3) $ |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
337 |
Abs("v", T1, |
2031 | 338 |
ap_split T2 T3 |
339 |
((ap_split T1 (factors T2 ---> T3) (incr_boundvars 1 u)) $ |
|
340 |
Bound 0)) |
|
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
341 |
| ap_split T T3 u = u; |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
342 |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
343 |
(*Makes a nested tuple from a list, following the product type structure*) |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
344 |
fun mk_tuple (Type("*", [T1,T2])) tms = |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
345 |
mk_Pair (mk_tuple T1 tms, |
2031 | 346 |
mk_tuple T2 (drop (length (factors T1), tms))) |
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
347 |
| mk_tuple T (t::_) = t; |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
348 |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
349 |
(*Attempts to remove occurrences of split, and pair-valued parameters*) |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
350 |
val remove_split = rewrite_rule [split RS eq_reflection] o |
2031 | 351 |
rule_by_tactic (ALLGOALS split_all_tac); |
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
352 |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
353 |
(*Uncurries any Var of function type in the rule*) |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
354 |
fun split_rule_var (t as Var(v, Type("fun",[T1,T2])), rl) = |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
355 |
let val T' = factors T1 ---> T2 |
2031 | 356 |
val newt = ap_split T1 T2 (Var(v,T')) |
357 |
val cterm = Thm.cterm_of (#sign(rep_thm rl)) |
|
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
358 |
in |
2031 | 359 |
remove_split (instantiate ([], [(cterm t, cterm newt)]) rl) |
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
360 |
end |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
361 |
| split_rule_var (t,rl) = rl; |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
362 |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
363 |
(*Uncurries ALL function variables occurring in a rule's conclusion*) |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
364 |
fun split_rule rl = foldr split_rule_var (term_vars (concl_of rl), rl) |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
365 |
|> standard; |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
366 |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
367 |
end; |