| author | wenzelm | 
| Wed, 05 Nov 1997 19:39:34 +0100 | |
| changeset 4176 | 84a0bfbd74e5 | 
| parent 4134 | 5c6cb2a25816 | 
| child 4192 | c38ab5af38b5 | 
| 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: 
923diff
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 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 33 | goal Prod.thy "((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 | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 38 | goalw Prod.thy [fst_def] "fst((a,b)) = a"; | 
| 4089 | 39 | by (blast_tac (claset() addIs [select_equality]) 1); | 
| 923 | 40 | qed "fst_conv"; | 
| 41 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 42 | goalw Prod.thy [snd_def] "snd((a,b)) = b"; | 
| 4089 | 43 | by (blast_tac (claset() addIs [select_equality]) 1); | 
| 923 | 44 | qed "snd_conv"; | 
| 45 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 46 | goalw Prod.thy [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 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 52 | val [prem] = goal Prod.thy "[| !!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 | ||
| 4134 | 57 | fun pair_tac s = res_inst_tac [("p",s)] PairE THEN' hyp_subst_tac;
 | 
| 58 | ||
| 1301 | 59 | (* replace parameters of product type by individual component parameters *) | 
| 60 | local | |
| 61 | fun is_pair (_,Type("*",_)) = true
 | |
| 62 | | is_pair _ = false; | |
| 63 | ||
| 1727 | 64 | fun find_pair_param prem = | 
| 65 | let val params = Logic.strip_params prem | |
| 1301 | 66 | in if exists is_pair params | 
| 1727 | 67 | then let val params = rev(rename_wrt_term prem params) | 
| 1301 | 68 | (*as they are printed*) | 
| 69 | in apsome fst (find_first is_pair params) end | |
| 70 | else None | |
| 71 | end; | |
| 72 | ||
| 73 | in | |
| 74 | ||
| 1727 | 75 | val split_all_tac = REPEAT o SUBGOAL (fn (prem,i) => | 
| 76 | case find_pair_param prem of | |
| 1301 | 77 | None => no_tac | 
| 1727 | 78 |   | Some x => EVERY[res_inst_tac[("p",x)] PairE i,
 | 
| 79 | REPEAT(hyp_subst_tac i), prune_params_tac]); | |
| 1301 | 80 | |
| 81 | end; | |
| 82 | ||
| 3568 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 83 | (* Could be nice, but breaks too many proofs: | 
| 4089 | 84 | claset_ref() := claset() addbefore split_all_tac; | 
| 3568 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 85 | *) | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 86 | |
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 87 | (*** lemmas for splitting paired `!!' | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 88 | Does not work with simplifier because it also affects premises in | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 89 | congrence rules, where is can lead to premises of the form | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 90 | !!a b. ... = ?P(a,b) | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 91 | which cannot be solved by reflexivity. | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 92 | |
| 3842 | 93 | val [prem] = goal Prod.thy "(!!x. PROP P x) ==> (!!a b. PROP P(a,b))"; | 
| 3568 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 94 | br prem 1; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 95 | val lemma1 = result(); | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 96 | |
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 97 | local | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 98 | val psig = sign_of Prod.thy; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 99 | val pT = Sign.read_typ (psig, K None) "?'a*?'b=>prop"; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 100 |     val PeqP = reflexive(read_cterm psig ("P", pT));
 | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 101 |     val psplit = zero_var_indexes(read_instantiate [("p","x")]
 | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 102 | surjective_pairing RS eq_reflection) | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 103 | in | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 104 | val adhoc = combination PeqP psplit | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 105 | end; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 106 | |
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 107 | |
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 108 | val [prem] = goal Prod.thy "(!!a b. PROP P(a,b)) ==> PROP P x"; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 109 | bw adhoc; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 110 | br prem 1; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 111 | val lemma = result(); | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 112 | |
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 113 | val [prem] = goal Prod.thy "(!!a b. PROP P(a,b)) ==> (!!x. PROP P x)"; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 114 | br lemma 1; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 115 | br prem 1; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 116 | val lemma2 = result(); | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 117 | |
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 118 | bind_thm("split_paired_all", equal_intr lemma1 lemma2);
 | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 119 | Addsimps [split_paired_all]; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 120 | ***) | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 121 | |
| 1301 | 122 | goal Prod.thy "(!x. P x) = (!a b. P(a,b))"; | 
| 4089 | 123 | by (fast_tac (claset() addbefore split_all_tac) 1); | 
| 1301 | 124 | qed "split_paired_All"; | 
| 3568 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 125 | Addsimps [split_paired_All]; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 126 | (* 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: 
3429diff
changeset | 127 | |
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 128 | goal Prod.thy "(? x. P x) = (? a b. P(a,b))"; | 
| 4089 | 129 | by (fast_tac (claset() addbefore split_all_tac) 1); | 
| 3568 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 130 | qed "split_paired_Ex"; | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 131 | (* Addsimps [split_paired_Ex]; breaks a number of IOA proofs *) | 
| 1301 | 132 | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 133 | 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: 
1465diff
changeset | 134 | by (EVERY1[stac fst_conv, stac snd_conv]); | 
| 923 | 135 | by (rtac refl 1); | 
| 136 | qed "split"; | |
| 137 | ||
| 4134 | 138 | val split_select = prove_goalw Prod.thy [split_def] | 
| 139 | "(®(x,y). P x y) = (®xy. P (fst xy) (snd xy))" (K [rtac refl 1]); | |
| 140 | ||
| 3568 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 141 | Addsimps [fst_conv, snd_conv, split]; | 
| 923 | 142 | |
| 143 | goal Prod.thy "(s=t) = (fst(s)=fst(t) & snd(s)=snd(t))"; | |
| 144 | by (res_inst_tac[("p","s")] PairE 1);
 | |
| 145 | by (res_inst_tac[("p","t")] PairE 1);
 | |
| 1264 | 146 | by (Asm_simp_tac 1); | 
| 923 | 147 | qed "Pair_fst_snd_eq"; | 
| 148 | ||
| 149 | (*Prevents simplification of c: much faster*) | |
| 150 | qed_goal "split_weak_cong" Prod.thy | |
| 151 | "p=q ==> split c p = split c q" | |
| 152 | (fn [prem] => [rtac (prem RS arg_cong) 1]); | |
| 153 | ||
| 154 | (* Do not add as rewrite rule: invalidates some proofs in IMP *) | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 155 | goal Prod.thy "p = (fst(p),snd(p))"; | 
| 923 | 156 | by (res_inst_tac [("p","p")] PairE 1);
 | 
| 1264 | 157 | by (Asm_simp_tac 1); | 
| 923 | 158 | qed "surjective_pairing"; | 
| 159 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 160 | goal Prod.thy "p = split (%x y.(x,y)) p"; | 
| 923 | 161 | by (res_inst_tac [("p","p")] PairE 1);
 | 
| 1264 | 162 | by (Asm_simp_tac 1); | 
| 923 | 163 | qed "surjective_pairing2"; | 
| 164 | ||
| 4134 | 165 | val surj_pair = prove_goal Prod.thy "? x y. z = (x, y)" (K [ | 
| 166 | rtac exI 1, rtac exI 1, rtac surjective_pairing 1]); | |
| 167 | Addsimps [surj_pair]; | |
| 168 | ||
| 1655 | 169 | qed_goal "split_eta" Prod.thy "(%(x,y). f(x,y)) = f" | 
| 170 | (fn _ => [rtac ext 1, split_all_tac 1, rtac split 1]); | |
| 171 | ||
| 4134 | 172 | val split_beta = prove_goal Prod.thy "(%(x,y). P x y) z = P (fst z) (snd z)" | 
| 173 | (fn _ => [stac surjective_pairing 1, stac split 1, rtac refl 1]); | |
| 174 | ||
| 923 | 175 | (*For use with split_tac and the simplifier*) | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 176 | goal Prod.thy "R(split c p) = (! x y. p = (x,y) --> R(c x y))"; | 
| 923 | 177 | by (stac surjective_pairing 1); | 
| 178 | by (stac split 1); | |
| 2935 | 179 | by (Blast_tac 1); | 
| 923 | 180 | qed "expand_split"; | 
| 181 | ||
| 3568 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 182 | (* could be done after split_tac has been speeded up significantly: | 
| 4089 | 183 | simpset_ref() := (simpset() addsplits [expand_split]); | 
| 3568 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 184 | 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: 
3429diff
changeset | 185 | 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: 
3429diff
changeset | 186 | *) | 
| 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 187 | |
| 923 | 188 | (** split used as a logical connective or set former **) | 
| 189 | ||
| 2935 | 190 | (*These rules are for use with blast_tac. | 
| 923 | 191 | Could instead call simp_tac/asm_full_simp_tac using split as rewrite.*) | 
| 192 | ||
| 1454 
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
 nipkow parents: 
1301diff
changeset | 193 | goal Prod.thy "!!p. [| !!a b. p=(a,b) ==> c a b |] ==> split c p"; | 
| 1552 | 194 | by (split_all_tac 1); | 
| 1454 
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
 nipkow parents: 
1301diff
changeset | 195 | by (Asm_simp_tac 1); | 
| 
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
 nipkow parents: 
1301diff
changeset | 196 | qed "splitI2"; | 
| 
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
 nipkow parents: 
1301diff
changeset | 197 | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 198 | goal Prod.thy "!!a b c. c a b ==> split c (a,b)"; | 
| 1264 | 199 | by (Asm_simp_tac 1); | 
| 923 | 200 | qed "splitI"; | 
| 201 | ||
| 202 | val prems = goalw Prod.thy [split_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 203 | "[| split c p; !!x y. [| p = (x,y); c x y |] ==> Q |] ==> Q"; | 
| 923 | 204 | by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); | 
| 205 | qed "splitE"; | |
| 206 | ||
| 4134 | 207 | val splitE2 = prove_goal Prod.thy | 
| 208 | "[|Q (split P z); !!x y. [|z = (x, y); Q (P x y)|] ==> R|] ==> R" (fn prems => [ | |
| 209 | REPEAT (resolve_tac (prems@[surjective_pairing]) 1), | |
| 210 | rtac (split_beta RS subst) 1, | |
| 211 | rtac (hd prems) 1]); | |
| 212 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 213 | goal Prod.thy "!!R a b. split R (a,b) ==> R a b"; | 
| 923 | 214 | by (etac (split RS iffD1) 1); | 
| 215 | qed "splitD"; | |
| 216 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 217 | goal Prod.thy "!!a b c. z: c a b ==> z: split c (a,b)"; | 
| 1264 | 218 | by (Asm_simp_tac 1); | 
| 923 | 219 | qed "mem_splitI"; | 
| 220 | ||
| 1454 
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
 nipkow parents: 
1301diff
changeset | 221 | goal Prod.thy "!!p. [| !!a b. p=(a,b) ==> z: c a b |] ==> z: split c p"; | 
| 1552 | 222 | by (split_all_tac 1); | 
| 1454 
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
 nipkow parents: 
1301diff
changeset | 223 | by (Asm_simp_tac 1); | 
| 
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
 nipkow parents: 
1301diff
changeset | 224 | qed "mem_splitI2"; | 
| 
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
 nipkow parents: 
1301diff
changeset | 225 | |
| 923 | 226 | val prems = goalw Prod.thy [split_def] | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 227 | "[| z: split c p; !!x y. [| p = (x,y); z: c x y |] ==> Q |] ==> Q"; | 
| 923 | 228 | by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); | 
| 229 | qed "mem_splitE"; | |
| 230 | ||
| 2856 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 231 | AddSIs [splitI, splitI2, mem_splitI, mem_splitI2]; | 
| 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 232 | AddSEs [splitE, mem_splitE]; | 
| 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 233 | |
| 923 | 234 | (*** prod_fun -- action of the product functor upon functions ***) | 
| 235 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 236 | goalw Prod.thy [prod_fun_def] "prod_fun f g (a,b) = (f(a),g(b))"; | 
| 923 | 237 | by (rtac split 1); | 
| 238 | qed "prod_fun"; | |
| 239 | ||
| 240 | goal Prod.thy | |
| 241 | "prod_fun (f1 o f2) (g1 o g2) = ((prod_fun f1 g1) o (prod_fun f2 g2))"; | |
| 242 | by (rtac ext 1); | |
| 243 | by (res_inst_tac [("p","x")] PairE 1);
 | |
| 4089 | 244 | by (asm_simp_tac (simpset() addsimps [prod_fun,o_def]) 1); | 
| 923 | 245 | qed "prod_fun_compose"; | 
| 246 | ||
| 3842 | 247 | goal Prod.thy "prod_fun (%x. x) (%y. y) = (%z. z)"; | 
| 923 | 248 | by (rtac ext 1); | 
| 249 | by (res_inst_tac [("p","z")] PairE 1);
 | |
| 4089 | 250 | by (asm_simp_tac (simpset() addsimps [prod_fun]) 1); | 
| 923 | 251 | qed "prod_fun_ident"; | 
| 252 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 253 | val prems = goal Prod.thy "(a,b):r ==> (f(a),g(b)) : (prod_fun f g)``r"; | 
| 923 | 254 | by (rtac image_eqI 1); | 
| 255 | by (rtac (prod_fun RS sym) 1); | |
| 256 | by (resolve_tac prems 1); | |
| 257 | qed "prod_fun_imageI"; | |
| 258 | ||
| 259 | val major::prems = goal Prod.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 260 | "[| c: (prod_fun f g)``r; !!x y. [| c=(f(x),g(y)); (x,y):r |] ==> P \ | 
| 923 | 261 | \ |] ==> P"; | 
| 262 | by (rtac (major RS imageE) 1); | |
| 263 | by (res_inst_tac [("p","x")] PairE 1);
 | |
| 264 | by (resolve_tac prems 1); | |
| 2935 | 265 | by (Blast_tac 2); | 
| 4089 | 266 | by (blast_tac (claset() addIs [prod_fun]) 1); | 
| 923 | 267 | qed "prod_fun_imageE"; | 
| 268 | ||
| 269 | (*** Disjoint union of a family of sets - Sigma ***) | |
| 270 | ||
| 271 | qed_goalw "SigmaI" Prod.thy [Sigma_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 272 | "[| a:A; b:B(a) |] ==> (a,b) : Sigma A B" | 
| 923 | 273 | (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]); | 
| 274 | ||
| 2856 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 275 | AddSIs [SigmaI]; | 
| 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 276 | |
| 923 | 277 | (*The general elimination rule*) | 
| 278 | qed_goalw "SigmaE" Prod.thy [Sigma_def] | |
| 279 | "[| c: Sigma A B; \ | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 280 | \ !!x y.[| x:A; y:B(x); c=(x,y) |] ==> P \ | 
| 923 | 281 | \ |] ==> P" | 
| 282 | (fn major::prems=> | |
| 283 | [ (cut_facts_tac [major] 1), | |
| 284 | (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); | |
| 285 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 286 | (** Elimination of (a,b):A*B -- introduces no eigenvariables **) | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 287 | qed_goal "SigmaD1" Prod.thy "(a,b) : Sigma A B ==> a : A" | 
| 923 | 288 | (fn [major]=> | 
| 289 | [ (rtac (major RS SigmaE) 1), | |
| 290 | (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); | |
| 291 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 292 | qed_goal "SigmaD2" Prod.thy "(a,b) : Sigma A B ==> b : B(a)" | 
| 923 | 293 | (fn [major]=> | 
| 294 | [ (rtac (major RS SigmaE) 1), | |
| 295 | (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); | |
| 296 | ||
| 297 | qed_goal "SigmaE2" Prod.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 298 | "[| (a,b) : Sigma A B; \ | 
| 923 | 299 | \ [| a:A; b:B(a) |] ==> P \ | 
| 300 | \ |] ==> P" | |
| 301 | (fn [major,minor]=> | |
| 302 | [ (rtac minor 1), | |
| 303 | (rtac (major RS SigmaD1) 1), | |
| 304 | (rtac (major RS SigmaD2) 1) ]); | |
| 305 | ||
| 2856 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 306 | AddSEs [SigmaE2, SigmaE]; | 
| 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 307 | |
| 1515 | 308 | val prems = goal Prod.thy | 
| 1642 | 309 | "[| A<=C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D"; | 
| 1515 | 310 | by (cut_facts_tac prems 1); | 
| 4089 | 311 | by (blast_tac (claset() addIs (prems RL [subsetD])) 1); | 
| 1515 | 312 | qed "Sigma_mono"; | 
| 313 | ||
| 1618 | 314 | qed_goal "Sigma_empty1" Prod.thy "Sigma {} B = {}"
 | 
| 2935 | 315 | (fn _ => [ (Blast_tac 1) ]); | 
| 1618 | 316 | |
| 1642 | 317 | qed_goal "Sigma_empty2" Prod.thy "A Times {} = {}"
 | 
| 2935 | 318 | (fn _ => [ (Blast_tac 1) ]); | 
| 1618 | 319 | |
| 320 | Addsimps [Sigma_empty1,Sigma_empty2]; | |
| 321 | ||
| 322 | goal Prod.thy "((a,b): Sigma A B) = (a:A & b:B(a))"; | |
| 2935 | 323 | by (Blast_tac 1); | 
| 1618 | 324 | qed "mem_Sigma_iff"; | 
| 3568 
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
 nipkow parents: 
3429diff
changeset | 325 | AddIffs [mem_Sigma_iff]; | 
| 1618 | 326 | |
| 4134 | 327 | val Collect_Prod = prove_goal Prod.thy | 
| 328 | 	"{(a,b). P a & Q b} = Collect P Times Collect Q" (K [Blast_tac 1]);
 | |
| 329 | Addsimps [Collect_Prod]; | |
| 1515 | 330 | |
| 2856 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 331 | (*Suggested by Pierre Chartier*) | 
| 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 332 | goal Prod.thy | 
| 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 333 | "(UN (a,b):(A Times B). E a Times F b) = (UNION A E) Times (UNION B F)"; | 
| 2935 | 334 | by (Blast_tac 1); | 
| 2856 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 335 | qed "UNION_Times_distrib"; | 
| 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 336 | |
| 923 | 337 | (*** Domain of a relation ***) | 
| 338 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 339 | val prems = goalw Prod.thy [image_def] "(a,b) : r ==> a : fst``r"; | 
| 923 | 340 | by (rtac CollectI 1); | 
| 341 | by (rtac bexI 1); | |
| 342 | by (rtac (fst_conv RS sym) 1); | |
| 343 | by (resolve_tac prems 1); | |
| 344 | qed "fst_imageI"; | |
| 345 | ||
| 346 | val major::prems = goal Prod.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 347 | "[| a : fst``r; !!y.[| (a,y) : r |] ==> P |] ==> P"; | 
| 923 | 348 | by (rtac (major RS imageE) 1); | 
| 349 | by (resolve_tac prems 1); | |
| 350 | by (etac ssubst 1); | |
| 351 | by (rtac (surjective_pairing RS subst) 1); | |
| 352 | by (assume_tac 1); | |
| 353 | qed "fst_imageE"; | |
| 354 | ||
| 355 | (*** Range of a relation ***) | |
| 356 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 357 | val prems = goalw Prod.thy [image_def] "(a,b) : r ==> b : snd``r"; | 
| 923 | 358 | by (rtac CollectI 1); | 
| 359 | by (rtac bexI 1); | |
| 360 | by (rtac (snd_conv RS sym) 1); | |
| 361 | by (resolve_tac prems 1); | |
| 362 | qed "snd_imageI"; | |
| 363 | ||
| 364 | val major::prems = goal Prod.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 365 | "[| a : snd``r; !!y.[| (y,a) : r |] ==> P |] ==> P"; | 
| 923 | 366 | by (rtac (major RS imageE) 1); | 
| 367 | by (resolve_tac prems 1); | |
| 368 | by (etac ssubst 1); | |
| 369 | by (rtac (surjective_pairing RS subst) 1); | |
| 370 | by (assume_tac 1); | |
| 371 | qed "snd_imageE"; | |
| 372 | ||
| 373 | (** Exhaustion rule for unit -- a degenerate form of induction **) | |
| 374 | ||
| 375 | goalw Prod.thy [Unity_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 376 | "u = ()"; | 
| 2886 | 377 | by (stac (rewrite_rule [unit_def] Rep_unit RS singletonD RS sym) 1); | 
| 2880 | 378 | by (rtac (Rep_unit_inverse RS sym) 1); | 
| 923 | 379 | qed "unit_eq"; | 
| 1754 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1746diff
changeset | 380 | |
| 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1746diff
changeset | 381 | AddIs [fst_imageI, snd_imageI, prod_fun_imageI]; | 
| 2856 
cdb908486a96
Reorganization of how classical rules are installed
 paulson parents: 
2637diff
changeset | 382 | AddSEs [fst_imageE, snd_imageE, prod_fun_imageE]; | 
| 923 | 383 | |
| 1746 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 384 | structure Prod_Syntax = | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 385 | struct | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 386 | |
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 387 | val unitT = Type("unit",[]);
 | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 388 | |
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 389 | fun mk_prod (T1,T2) = Type("*", [T1,T2]);
 | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 390 | |
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 391 | (*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: 
1727diff
changeset | 392 | fun factors (Type("*", [T1,T2])) = factors T1 @ factors T2
 | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 393 | | factors T = [T]; | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 394 | |
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 395 | (*Make a correctly typed ordered pair*) | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 396 | fun mk_Pair (t1,t2) = | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 397 | let val T1 = fastype_of t1 | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 398 | and T2 = fastype_of t2 | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 399 |   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: 
1727diff
changeset | 400 | |
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 401 | fun split_const(Ta,Tb,Tc) = | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 402 |     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: 
1727diff
changeset | 403 | |
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 404 | (*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: 
1727diff
changeset | 405 | 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: 
1727diff
changeset | 406 | of type S.*) | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 407 | fun ap_split (Type("*", [T1,T2])) T3 u = 
 | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 408 | split_const(T1,T2,T3) $ | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 409 |       Abs("v", T1, 
 | 
| 2031 | 410 | ap_split T2 T3 | 
| 411 | ((ap_split T1 (factors T2 ---> T3) (incr_boundvars 1 u)) $ | |
| 412 | Bound 0)) | |
| 1746 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 413 | | ap_split T T3 u = u; | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 414 | |
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 415 | (*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: 
1727diff
changeset | 416 | fun mk_tuple (Type("*", [T1,T2])) tms = 
 | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 417 | mk_Pair (mk_tuple T1 tms, | 
| 2031 | 418 | mk_tuple T2 (drop (length (factors T1), tms))) | 
| 1746 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 419 | | mk_tuple T (t::_) = t; | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 420 | |
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 421 | (*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: 
1727diff
changeset | 422 | val remove_split = rewrite_rule [split RS eq_reflection] o | 
| 2031 | 423 | rule_by_tactic (ALLGOALS split_all_tac); | 
| 1746 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 424 | |
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 425 | (*Uncurries any Var of function type in the rule*) | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 426 | 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: 
1727diff
changeset | 427 | let val T' = factors T1 ---> T2 | 
| 2031 | 428 | val newt = ap_split T1 T2 (Var(v,T')) | 
| 429 | 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: 
1727diff
changeset | 430 | in | 
| 2031 | 431 | remove_split (instantiate ([], [(cterm t, cterm newt)]) rl) | 
| 1746 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 432 | end | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 433 | | split_rule_var (t,rl) = rl; | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 434 | |
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 435 | (*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: 
1727diff
changeset | 436 | 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: 
1727diff
changeset | 437 | |> standard; | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 438 | |
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1727diff
changeset | 439 | end; |