| author | wenzelm |
| Fri, 06 Nov 1998 14:04:54 +0100 | |
| changeset 5807 | bd2d9dd34dfd |
| parent 5788 | e3a98a7c0634 |
| child 5810 | 829cae93f132 |
| 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*) |
|
| 5069 | 12 |
Goalw [Prod_def] "Pair_Rep a b : Prod"; |
| 923 | 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 |
||
| 5069 | 22 |
Goal "inj_on Abs_Prod Prod"; |
| 4830 | 23 |
by (rtac inj_on_inverseI 1); |
| 923 | 24 |
by (etac Abs_Prod_inverse 1); |
| 4830 | 25 |
qed "inj_on_Abs_Prod"; |
| 923 | 26 |
|
| 5316 | 27 |
val prems = Goalw [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"; |
| 4830 | 29 |
by (rtac (inj_on_Abs_Prod RS inj_onD RS Pair_Rep_inject RS conjE) 1); |
| 923 | 30 |
by (REPEAT (ares_tac (prems@[ProdI]) 1)); |
31 |
qed "Pair_inject"; |
|
32 |
||
| 5069 | 33 |
Goal "((a,b) = (a',b')) = (a=a' & b=b')"; |
| 4089 | 34 |
by (blast_tac (claset() addSEs [Pair_inject]) 1); |
| 923 | 35 |
qed "Pair_eq"; |
| 3429 | 36 |
AddIffs [Pair_eq]; |
| 923 | 37 |
|
| 5069 | 38 |
Goalw [fst_def] "fst((a,b)) = a"; |
| 4534 | 39 |
by (Blast_tac 1); |
| 923 | 40 |
qed "fst_conv"; |
| 5069 | 41 |
Goalw [snd_def] "snd((a,b)) = b"; |
| 4534 | 42 |
by (Blast_tac 1); |
| 923 | 43 |
qed "snd_conv"; |
| 4534 | 44 |
Addsimps [fst_conv, snd_conv]; |
| 923 | 45 |
|
| 5069 | 46 |
Goalw [Pair_def] "? x y. p = (x,y)"; |
| 923 | 47 |
by (rtac (rewrite_rule [Prod_def] Rep_Prod RS CollectE) 1); |
48 |
by (EVERY1[etac exE, etac exE, rtac exI, rtac exI, |
|
| 1465 | 49 |
rtac (Rep_Prod_inverse RS sym RS trans), etac arg_cong]); |
| 923 | 50 |
qed "PairE_lemma"; |
51 |
||
| 5316 | 52 |
val [prem] = Goal "[| !!x y. p = (x,y) ==> Q |] ==> Q"; |
| 923 | 53 |
by (rtac (PairE_lemma RS exE) 1); |
54 |
by (REPEAT (eresolve_tac [prem,exE] 1)); |
|
55 |
qed "PairE"; |
|
56 |
||
|
4819
ef2e8e2a10e1
improved pair_tac to call prune_params_tac afterwards
oheimb
parents:
4799
diff
changeset
|
57 |
fun pair_tac s = EVERY' [res_inst_tac [("p",s)] PairE, hyp_subst_tac,
|
|
ef2e8e2a10e1
improved pair_tac to call prune_params_tac afterwards
oheimb
parents:
4799
diff
changeset
|
58 |
K prune_params_tac]; |
| 4134 | 59 |
|
| 4828 | 60 |
(* Do not add as rewrite rule: invalidates some proofs in IMP *) |
| 5069 | 61 |
Goal "p = (fst(p),snd(p))"; |
| 4828 | 62 |
by (pair_tac "p" 1); |
63 |
by (Asm_simp_tac 1); |
|
64 |
qed "surjective_pairing"; |
|
65 |
||
66 |
val surj_pair = prove_goal Prod.thy "? x y. z = (x, y)" (K [ |
|
67 |
rtac exI 1, rtac exI 1, rtac surjective_pairing 1]); |
|
68 |
Addsimps [surj_pair]; |
|
69 |
||
70 |
||
| 5699 | 71 |
bind_thm ("split_paired_all",
|
72 |
SplitPairedAll.rule (standard (surjective_pairing RS eq_reflection))); |
|
| 4828 | 73 |
(* |
74 |
Addsimps [split_paired_all] does not work with simplifier |
|
75 |
because it also affects premises in congrence rules, |
|
76 |
where is can lead to premises of the form !!a b. ... = ?P(a,b) |
|
77 |
which cannot be solved by reflexivity. |
|
78 |
*) |
|
79 |
||
| 1301 | 80 |
(* replace parameters of product type by individual component parameters *) |
81 |
local |
|
|
4819
ef2e8e2a10e1
improved pair_tac to call prune_params_tac afterwards
oheimb
parents:
4799
diff
changeset
|
82 |
fun is_pair (_,Type("*",_)) = true
|
|
ef2e8e2a10e1
improved pair_tac to call prune_params_tac afterwards
oheimb
parents:
4799
diff
changeset
|
83 |
| is_pair _ = false; |
| 4828 | 84 |
fun exists_paired_all prem = exists is_pair (Logic.strip_params prem); |
85 |
val split_tac = full_simp_tac (HOL_basic_ss addsimps [split_paired_all]); |
|
| 1301 | 86 |
in |
| 4828 | 87 |
val split_all_tac = SUBGOAL (fn (prem,i) => |
88 |
if exists_paired_all prem then split_tac i else no_tac); |
|
| 1301 | 89 |
end; |
90 |
||
| 4828 | 91 |
claset_ref() := claset() addSWrapper ("split_all_tac",
|
92 |
fn tac2 => split_all_tac ORELSE' tac2); |
|
|
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3429
diff
changeset
|
93 |
|
| 5069 | 94 |
Goal "(!x. P x) = (!a b. P(a,b))"; |
| 4650 | 95 |
by (Fast_tac 1); |
| 1301 | 96 |
qed "split_paired_All"; |
|
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3429
diff
changeset
|
97 |
Addsimps [split_paired_All]; |
|
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3429
diff
changeset
|
98 |
(* AddIffs is not a good idea because it makes Blast_tac loop *) |
|
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3429
diff
changeset
|
99 |
|
|
5715
5fc697ad232b
Added theorem prod_induct (needed for rep_datatype).
berghofe
parents:
5699
diff
changeset
|
100 |
bind_thm ("prod_induct",
|
|
5fc697ad232b
Added theorem prod_induct (needed for rep_datatype).
berghofe
parents:
5699
diff
changeset
|
101 |
allI RS (allI RS (split_paired_All RS iffD2)) RS spec); |
|
5fc697ad232b
Added theorem prod_induct (needed for rep_datatype).
berghofe
parents:
5699
diff
changeset
|
102 |
|
| 5069 | 103 |
Goal "(? x. P x) = (? a b. P(a,b))"; |
| 4650 | 104 |
by (Fast_tac 1); |
|
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3429
diff
changeset
|
105 |
qed "split_paired_Ex"; |
| 4534 | 106 |
Addsimps [split_paired_Ex]; |
| 1301 | 107 |
|
| 5069 | 108 |
Goalw [split_def] "split c (a,b) = c a b"; |
| 4534 | 109 |
by (Simp_tac 1); |
| 923 | 110 |
qed "split"; |
| 4534 | 111 |
Addsimps [split]; |
| 923 | 112 |
|
| 5069 | 113 |
Goal "split Pair p = p"; |
| 4828 | 114 |
by (pair_tac "p" 1); |
115 |
by (Simp_tac 1); |
|
116 |
qed "split_Pair"; |
|
117 |
(*unused: val surjective_pairing2 = split_Pair RS sym;*) |
|
118 |
||
| 5069 | 119 |
Goal "!!s t. (s=t) = (fst(s)=fst(t) & snd(s)=snd(t))"; |
| 4828 | 120 |
by (split_all_tac 1); |
| 1264 | 121 |
by (Asm_simp_tac 1); |
| 923 | 122 |
qed "Pair_fst_snd_eq"; |
123 |
||
124 |
(*Prevents simplification of c: much faster*) |
|
125 |
qed_goal "split_weak_cong" Prod.thy |
|
126 |
"p=q ==> split c p = split c q" |
|
127 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
|
128 |
||
| 1655 | 129 |
qed_goal "split_eta" Prod.thy "(%(x,y). f(x,y)) = f" |
| 4534 | 130 |
(K [rtac ext 1, split_all_tac 1, rtac split 1]); |
| 1655 | 131 |
|
| 4989 | 132 |
qed_goal "cond_split_eta" Prod.thy |
133 |
"!!f. (!!x y. f x y = g(x,y)) ==> (%(x,y). f x y) = g" |
|
134 |
(K [asm_simp_tac (simpset() addsimps [split_eta]) 1]); |
|
| 5294 | 135 |
|
136 |
||
137 |
(*simplification procedure for cond_split_eta. |
|
138 |
using split_eta a rewrite rule is not general enough, and using |
|
139 |
cond_split_eta directly would render some existing proofs very inefficient*) |
|
140 |
local |
|
141 |
val split_eta_pat = Thm.read_cterm (sign_of thy) |
|
142 |
("split (%x. split (%y. f x y))", HOLogic.termTVar);
|
|
| 5553 | 143 |
val split_eta_meta_eq = standard (mk_meta_eq cond_split_eta); |
| 5294 | 144 |
fun Pair_pat 0 (Bound 0) = true |
145 |
| Pair_pat k (Const ("Pair", _) $ Bound m $ t) =
|
|
146 |
m = k andalso Pair_pat (k-1) t |
|
147 |
| Pair_pat _ _ = false; |
|
148 |
fun split_pat k (Abs (_, _, f $ |
|
149 |
(Const ("Pair",_) $ Bound m $
|
|
150 |
(Const ("Pair",_) $ Bound n $ t)))) =
|
|
151 |
if m = k andalso n = k-1 andalso Pair_pat (k-2) t |
|
152 |
then Some f else None |
|
153 |
| split_pat k (Const ("split", _) $ Abs (_, _, t)) = split_pat (k+1) t
|
|
154 |
| split_pat k _ = None; |
|
| 5361 | 155 |
fun proc sg _ (s as |
| 5294 | 156 |
(Const ("split", _) $ Abs (_, _,
|
157 |
(Const ("split", _) $ Abs (_, _, t))))) = (case split_pat 2 t of
|
|
158 |
Some f => (let val fvar = Free ("f", fastype_of f);
|
|
159 |
fun atom_fun t = if t = f then fvar else atom t |
|
160 |
and atom (Abs (x, T, t)) = Abs (x, T,atom_fun t) |
|
161 |
| atom (t $ u) = atom_fun t $ atom_fun u |
|
162 |
| atom x = x; |
|
| 5361 | 163 |
val ct = cterm_of sg (HOLogic.mk_Trueprop |
| 5294 | 164 |
(HOLogic.mk_eq (atom_fun s,fvar))); |
165 |
val ss = HOL_basic_ss addsimps [cond_split_eta]; |
|
| 5553 | 166 |
in Some (mk_meta_eq (prove_goalw_cterm [] ct (K [simp_tac ss 1]))) end) |
| 5294 | 167 |
| None => None) |
168 |
| proc _ _ _ = None; |
|
169 |
||
170 |
in |
|
171 |
val split_eta_proc = Simplifier.mk_simproc "split_eta" [split_eta_pat] proc; |
|
172 |
end; |
|
173 |
||
174 |
Addsimprocs [split_eta_proc]; |
|
175 |
||
| 4989 | 176 |
|
|
4819
ef2e8e2a10e1
improved pair_tac to call prune_params_tac afterwards
oheimb
parents:
4799
diff
changeset
|
177 |
qed_goal "split_beta" Prod.thy "(%(x,y). P x y) z = P (fst z) (snd z)" |
| 4534 | 178 |
(K [stac surjective_pairing 1, stac split 1, rtac refl 1]); |
| 4134 | 179 |
|
| 923 | 180 |
(*For use with split_tac and the simplifier*) |
| 5069 | 181 |
Goal "R (split c p) = (! x y. p = (x,y) --> R (c x y))"; |
| 923 | 182 |
by (stac surjective_pairing 1); |
183 |
by (stac split 1); |
|
| 2935 | 184 |
by (Blast_tac 1); |
| 4830 | 185 |
qed "split_split"; |
| 923 | 186 |
|
|
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3429
diff
changeset
|
187 |
(* could be done after split_tac has been speeded up significantly: |
| 4830 | 188 |
simpset_ref() := simpset() addsplits [split_split]; |
|
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3429
diff
changeset
|
189 |
precompute the constants involved and don't do anything unless |
|
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3429
diff
changeset
|
190 |
the current goal contains one of those constants |
|
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3429
diff
changeset
|
191 |
*) |
|
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3429
diff
changeset
|
192 |
|
| 5069 | 193 |
Goal "R (split c p) = (~(? x y. p = (x,y) & (~R (c x y))))"; |
| 4830 | 194 |
by (stac split_split 1); |
| 4435 | 195 |
by (Simp_tac 1); |
196 |
qed "expand_split_asm"; |
|
197 |
||
| 923 | 198 |
(** split used as a logical connective or set former **) |
199 |
||
| 2935 | 200 |
(*These rules are for use with blast_tac. |
| 923 | 201 |
Could instead call simp_tac/asm_full_simp_tac using split as rewrite.*) |
202 |
||
| 5069 | 203 |
Goal "!!p. [| !!a b. p=(a,b) ==> c a b |] ==> split c p"; |
| 1552 | 204 |
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
|
205 |
by (Asm_simp_tac 1); |
|
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
206 |
qed "splitI2"; |
|
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
207 |
|
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5132
diff
changeset
|
208 |
Goal "c a b ==> split c (a,b)"; |
| 1264 | 209 |
by (Asm_simp_tac 1); |
| 923 | 210 |
qed "splitI"; |
211 |
||
| 5316 | 212 |
val prems = Goalw [split_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
213 |
"[| split c p; !!x y. [| p = (x,y); c x y |] ==> Q |] ==> Q"; |
| 923 | 214 |
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); |
215 |
qed "splitE"; |
|
216 |
||
| 4134 | 217 |
val splitE2 = prove_goal Prod.thy |
218 |
"[|Q (split P z); !!x y. [|z = (x, y); Q (P x y)|] ==> R|] ==> R" (fn prems => [ |
|
219 |
REPEAT (resolve_tac (prems@[surjective_pairing]) 1), |
|
220 |
rtac (split_beta RS subst) 1, |
|
221 |
rtac (hd prems) 1]); |
|
222 |
||
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5132
diff
changeset
|
223 |
Goal "split R (a,b) ==> R a b"; |
| 923 | 224 |
by (etac (split RS iffD1) 1); |
225 |
qed "splitD"; |
|
226 |
||
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5132
diff
changeset
|
227 |
Goal "z: c a b ==> z: split c (a,b)"; |
| 1264 | 228 |
by (Asm_simp_tac 1); |
| 923 | 229 |
qed "mem_splitI"; |
230 |
||
| 5069 | 231 |
Goal "!!p. [| !!a b. p=(a,b) ==> z: c a b |] ==> z: split c p"; |
| 1552 | 232 |
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
|
233 |
by (Asm_simp_tac 1); |
|
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
234 |
qed "mem_splitI2"; |
|
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1301
diff
changeset
|
235 |
|
| 5316 | 236 |
val prems = Goalw [split_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
237 |
"[| z: split c p; !!x y. [| p = (x,y); z: c x y |] ==> Q |] ==> Q"; |
| 923 | 238 |
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); |
239 |
qed "mem_splitE"; |
|
240 |
||
|
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
241 |
AddSIs [splitI, splitI2, mem_splitI, mem_splitI2]; |
|
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
242 |
AddSEs [splitE, mem_splitE]; |
|
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
243 |
|
| 4534 | 244 |
(* allows simplifications of nested splits in case of independent predicates *) |
| 5069 | 245 |
Goal "(%(a,b). P & Q a b) = (%ab. P & split Q ab)"; |
| 4534 | 246 |
by (rtac ext 1); |
247 |
by (Blast_tac 1); |
|
248 |
qed "split_part"; |
|
249 |
Addsimps [split_part]; |
|
250 |
||
| 5069 | 251 |
Goal "(@(x',y'). x = x' & y = y') = (x,y)"; |
| 4534 | 252 |
by (Blast_tac 1); |
253 |
qed "Eps_split_eq"; |
|
254 |
Addsimps [Eps_split_eq]; |
|
255 |
(* |
|
256 |
the following would be slightly more general, |
|
257 |
but cannot be used as rewrite rule: |
|
258 |
### Cannot add premise as rewrite rule because it contains (type) unknowns: |
|
259 |
### ?y = .x |
|
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5132
diff
changeset
|
260 |
Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)"; |
| 4534 | 261 |
by (rtac select_equality 1); |
262 |
by ( Simp_tac 1); |
|
263 |
by (split_all_tac 1); |
|
264 |
by (Asm_full_simp_tac 1); |
|
265 |
qed "Eps_split_eq"; |
|
266 |
*) |
|
267 |
||
| 923 | 268 |
(*** prod_fun -- action of the product functor upon functions ***) |
269 |
||
| 5069 | 270 |
Goalw [prod_fun_def] "prod_fun f g (a,b) = (f(a),g(b))"; |
| 923 | 271 |
by (rtac split 1); |
272 |
qed "prod_fun"; |
|
| 4521 | 273 |
Addsimps [prod_fun]; |
| 923 | 274 |
|
| 5278 | 275 |
Goal "prod_fun (f1 o f2) (g1 o g2) = ((prod_fun f1 g1) o (prod_fun f2 g2))"; |
| 923 | 276 |
by (rtac ext 1); |
| 4828 | 277 |
by (pair_tac "x" 1); |
| 4521 | 278 |
by (Asm_simp_tac 1); |
| 923 | 279 |
qed "prod_fun_compose"; |
280 |
||
| 5069 | 281 |
Goal "prod_fun (%x. x) (%y. y) = (%z. z)"; |
| 923 | 282 |
by (rtac ext 1); |
| 4828 | 283 |
by (pair_tac "z" 1); |
| 4521 | 284 |
by (Asm_simp_tac 1); |
| 923 | 285 |
qed "prod_fun_ident"; |
| 4521 | 286 |
Addsimps [prod_fun_ident]; |
| 923 | 287 |
|
| 5316 | 288 |
Goal "(a,b):r ==> (f(a),g(b)) : (prod_fun f g)``r"; |
| 923 | 289 |
by (rtac image_eqI 1); |
290 |
by (rtac (prod_fun RS sym) 1); |
|
| 5316 | 291 |
by (assume_tac 1); |
| 923 | 292 |
qed "prod_fun_imageI"; |
293 |
||
| 5316 | 294 |
val major::prems = Goal |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
295 |
"[| c: (prod_fun f g)``r; !!x y. [| c=(f(x),g(y)); (x,y):r |] ==> P \ |
| 923 | 296 |
\ |] ==> P"; |
297 |
by (rtac (major RS imageE) 1); |
|
298 |
by (res_inst_tac [("p","x")] PairE 1);
|
|
299 |
by (resolve_tac prems 1); |
|
| 2935 | 300 |
by (Blast_tac 2); |
| 4089 | 301 |
by (blast_tac (claset() addIs [prod_fun]) 1); |
| 923 | 302 |
qed "prod_fun_imageE"; |
303 |
||
| 5788 | 304 |
AddIs [prod_fun_imageI]; |
305 |
AddSEs [prod_fun_imageE]; |
|
306 |
||
| 4521 | 307 |
|
| 923 | 308 |
(*** Disjoint union of a family of sets - Sigma ***) |
309 |
||
310 |
qed_goalw "SigmaI" Prod.thy [Sigma_def] |
|
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
311 |
"[| a:A; b:B(a) |] ==> (a,b) : Sigma A B" |
| 923 | 312 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]); |
313 |
||
|
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
314 |
AddSIs [SigmaI]; |
|
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
315 |
|
| 923 | 316 |
(*The general elimination rule*) |
317 |
qed_goalw "SigmaE" Prod.thy [Sigma_def] |
|
318 |
"[| c: Sigma A B; \ |
|
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
319 |
\ !!x y.[| x:A; y:B(x); c=(x,y) |] ==> P \ |
| 923 | 320 |
\ |] ==> P" |
321 |
(fn major::prems=> |
|
322 |
[ (cut_facts_tac [major] 1), |
|
323 |
(REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); |
|
324 |
||
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
325 |
(** Elimination of (a,b):A*B -- introduces no eigenvariables **) |
|
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
326 |
qed_goal "SigmaD1" Prod.thy "(a,b) : Sigma A B ==> a : A" |
| 923 | 327 |
(fn [major]=> |
328 |
[ (rtac (major RS SigmaE) 1), |
|
329 |
(REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); |
|
330 |
||
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
331 |
qed_goal "SigmaD2" Prod.thy "(a,b) : Sigma A B ==> b : B(a)" |
| 923 | 332 |
(fn [major]=> |
333 |
[ (rtac (major RS SigmaE) 1), |
|
334 |
(REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); |
|
335 |
||
336 |
qed_goal "SigmaE2" Prod.thy |
|
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
337 |
"[| (a,b) : Sigma A B; \ |
| 923 | 338 |
\ [| a:A; b:B(a) |] ==> P \ |
339 |
\ |] ==> P" |
|
340 |
(fn [major,minor]=> |
|
341 |
[ (rtac minor 1), |
|
342 |
(rtac (major RS SigmaD1) 1), |
|
343 |
(rtac (major RS SigmaD2) 1) ]); |
|
344 |
||
|
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
345 |
AddSEs [SigmaE2, SigmaE]; |
|
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
346 |
|
| 5316 | 347 |
val prems = Goal |
| 1642 | 348 |
"[| A<=C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D"; |
| 1515 | 349 |
by (cut_facts_tac prems 1); |
| 4089 | 350 |
by (blast_tac (claset() addIs (prems RL [subsetD])) 1); |
| 1515 | 351 |
qed "Sigma_mono"; |
352 |
||
| 1618 | 353 |
qed_goal "Sigma_empty1" Prod.thy "Sigma {} B = {}"
|
| 2935 | 354 |
(fn _ => [ (Blast_tac 1) ]); |
| 1618 | 355 |
|
| 1642 | 356 |
qed_goal "Sigma_empty2" Prod.thy "A Times {} = {}"
|
| 2935 | 357 |
(fn _ => [ (Blast_tac 1) ]); |
| 1618 | 358 |
|
359 |
Addsimps [Sigma_empty1,Sigma_empty2]; |
|
360 |
||
| 5069 | 361 |
Goal "((a,b): Sigma A B) = (a:A & b:B(a))"; |
| 2935 | 362 |
by (Blast_tac 1); |
| 1618 | 363 |
qed "mem_Sigma_iff"; |
|
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3429
diff
changeset
|
364 |
AddIffs [mem_Sigma_iff]; |
| 1618 | 365 |
|
| 4534 | 366 |
val Collect_split = prove_goal Prod.thy |
| 4134 | 367 |
"{(a,b). P a & Q b} = Collect P Times Collect Q" (K [Blast_tac 1]);
|
| 4534 | 368 |
Addsimps [Collect_split]; |
| 1515 | 369 |
|
|
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
370 |
(*Suggested by Pierre Chartier*) |
| 5278 | 371 |
Goal "(UN (a,b):(A Times B). E a Times F b) = (UNION A E) Times (UNION B F)"; |
| 2935 | 372 |
by (Blast_tac 1); |
|
2856
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
373 |
qed "UNION_Times_distrib"; |
|
cdb908486a96
Reorganization of how classical rules are installed
paulson
parents:
2637
diff
changeset
|
374 |
|
| 5083 | 375 |
|
| 923 | 376 |
(** Exhaustion rule for unit -- a degenerate form of induction **) |
377 |
||
| 5069 | 378 |
Goalw [Unity_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
379 |
"u = ()"; |
| 2886 | 380 |
by (stac (rewrite_rule [unit_def] Rep_unit RS singletonD RS sym) 1); |
| 2880 | 381 |
by (rtac (Rep_unit_inverse RS sym) 1); |
| 923 | 382 |
qed "unit_eq"; |
|
1754
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1746
diff
changeset
|
383 |
|
|
5088
e4aa78d1312f
New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents:
5083
diff
changeset
|
384 |
(*simplification procedure for unit_eq. |
|
e4aa78d1312f
New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents:
5083
diff
changeset
|
385 |
Cannot use this rule directly -- it loops!*) |
| 5083 | 386 |
local |
387 |
val unit_pat = Thm.cterm_of (sign_of thy) (Free ("x", HOLogic.unitT));
|
|
| 5553 | 388 |
val unit_meta_eq = standard (mk_meta_eq unit_eq); |
| 5083 | 389 |
fun proc _ _ t = |
390 |
if HOLogic.is_unit t then None |
|
391 |
else Some unit_meta_eq; |
|
392 |
in |
|
393 |
val unit_eq_proc = Simplifier.mk_simproc "unit_eq" [unit_pat] proc; |
|
394 |
end; |
|
395 |
||
396 |
Addsimprocs [unit_eq_proc]; |
|
397 |
||
398 |
||
| 5761 | 399 |
Goal "P () ==> P x"; |
400 |
by (Simp_tac 1); |
|
401 |
qed "unit_induct"; |
|
402 |
||
403 |
||
|
5088
e4aa78d1312f
New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents:
5083
diff
changeset
|
404 |
(*This rewrite counters the effect of unit_eq_proc on (%u::unit. f u), |
|
e4aa78d1312f
New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents:
5083
diff
changeset
|
405 |
replacing it by f rather than by %u.f(). *) |
|
e4aa78d1312f
New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents:
5083
diff
changeset
|
406 |
Goal "(%u::unit. f()) = f"; |
|
e4aa78d1312f
New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents:
5083
diff
changeset
|
407 |
by (rtac ext 1); |
|
e4aa78d1312f
New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents:
5083
diff
changeset
|
408 |
by (Simp_tac 1); |
|
e4aa78d1312f
New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents:
5083
diff
changeset
|
409 |
qed "unit_abs_eta_conv"; |
|
e4aa78d1312f
New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents:
5083
diff
changeset
|
410 |
Addsimps [unit_abs_eta_conv]; |
|
e4aa78d1312f
New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents:
5083
diff
changeset
|
411 |
|
|
e4aa78d1312f
New rewrite unit_abs_eta_conv to compensate for unit_eq_proc
paulson
parents:
5083
diff
changeset
|
412 |
|
|
5096
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
413 |
(*Attempts to remove occurrences of split, and pair-valued parameters*) |
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
414 |
val remove_split = rewrite_rule [split RS eq_reflection] o |
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
415 |
rule_by_tactic (TRYALL split_all_tac); |
|
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
416 |
|
|
5096
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
417 |
local |
|
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
418 |
|
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
419 |
(*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
|
420 |
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
|
421 |
of type S.*) |
|
5096
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
422 |
fun ap_split (Type ("*", [T1, T2])) T3 u =
|
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
423 |
HOLogic.split_const (T1, T2, T3) $ |
|
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
424 |
Abs("v", T1,
|
| 2031 | 425 |
ap_split T2 T3 |
|
5096
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
426 |
((ap_split T1 (HOLogic.prodT_factors T2 ---> T3) (incr_boundvars 1 u)) $ |
| 2031 | 427 |
Bound 0)) |
|
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
428 |
| ap_split T T3 u = u; |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
429 |
|
|
5096
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
430 |
(*Curries any Var of function type in the rule*) |
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
431 |
fun split_rule_var' (t as Var (v, Type ("fun", [T1, T2])), rl) =
|
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
432 |
let val T' = HOLogic.prodT_factors T1 ---> T2 |
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
433 |
val newt = ap_split T1 T2 (Var (v, T')) |
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
434 |
val cterm = Thm.cterm_of (#sign (rep_thm rl)) |
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
435 |
in |
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
436 |
instantiate ([], [(cterm t, cterm newt)]) rl |
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
437 |
end |
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
438 |
| split_rule_var' (t, rl) = rl; |
|
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
439 |
|
|
5096
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
440 |
in |
|
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
441 |
|
|
5096
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
442 |
val split_rule_var = standard o remove_split o split_rule_var'; |
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
443 |
|
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
444 |
(*Curries ALL function variables occurring in a rule's conclusion*) |
|
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
5088
diff
changeset
|
445 |
fun split_rule rl = remove_split (foldr split_rule_var' (term_vars (concl_of rl), rl)) |
|
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
446 |
|> standard; |
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
447 |
|
|
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1727
diff
changeset
|
448 |
end; |