51 val [prem] = goal Prod.thy "[| !!x y. p = <x,y> ==> Q |] ==> Q"; |
51 val [prem] = goal Prod.thy "[| !!x y. p = <x,y> ==> Q |] ==> Q"; |
52 by (rtac (PairE_lemma RS exE) 1); |
52 by (rtac (PairE_lemma RS exE) 1); |
53 by (REPEAT (eresolve_tac [prem,exE] 1)); |
53 by (REPEAT (eresolve_tac [prem,exE] 1)); |
54 val PairE = result(); |
54 val PairE = result(); |
55 |
55 |
56 goalw Prod.thy [split_def] "split(<a,b>, c) = c(a,b)"; |
56 goalw Prod.thy [split_def] "split(c, <a,b>) = c(a,b)"; |
57 by (sstac [fst_conv, snd_conv] 1); |
57 by (sstac [fst_conv, snd_conv] 1); |
58 by (rtac refl 1); |
58 by (rtac refl 1); |
59 val split = result(); |
59 val split = result(); |
60 |
60 |
61 val pair_ss = set_ss addsimps [fst_conv, snd_conv, split, Pair_eq]; |
61 val pair_ss = set_ss addsimps [fst_conv, snd_conv, split, Pair_eq]; |
66 by (asm_simp_tac pair_ss 1); |
66 by (asm_simp_tac pair_ss 1); |
67 val Pair_fst_snd_eq = result(); |
67 val Pair_fst_snd_eq = result(); |
68 |
68 |
69 (*Prevents simplification of c: much faster*) |
69 (*Prevents simplification of c: much faster*) |
70 val split_weak_cong = prove_goal Prod.thy |
70 val split_weak_cong = prove_goal Prod.thy |
71 "p=q ==> split(p,c) = split(q,c)" |
71 "p=q ==> split(c,p) = split(c,q)" |
72 (fn [prem] => [rtac (prem RS arg_cong) 1]); |
72 (fn [prem] => [rtac (prem RS arg_cong) 1]); |
73 |
73 |
74 goal Prod.thy "p = <fst(p),snd(p)>"; |
74 goal Prod.thy "p = <fst(p),snd(p)>"; |
75 by (res_inst_tac [("p","p")] PairE 1); |
75 by (res_inst_tac [("p","p")] PairE 1); |
76 by (asm_simp_tac pair_ss 1); |
76 by (asm_simp_tac pair_ss 1); |
77 val surjective_pairing = result(); |
77 val surjective_pairing = result(); |
78 |
78 |
79 goal Prod.thy "p = split(p, %x y.<x,y>)"; |
79 goal Prod.thy "p = split(%x y.<x,y>, p)"; |
80 by (res_inst_tac [("p","p")] PairE 1); |
80 by (res_inst_tac [("p","p")] PairE 1); |
81 by (asm_simp_tac pair_ss 1); |
81 by (asm_simp_tac pair_ss 1); |
82 val surjective_pairing2 = result(); |
82 val surjective_pairing2 = result(); |
83 |
83 |
84 (** split used as a logical connective, with result type bool **) |
84 (*For use with split_tac and the simplifier*) |
85 |
85 goal Prod.thy "R(split(c,p)) = (! x y. p = <x,y> --> R(c(x,y)))"; |
86 val prems = goal Prod.thy "c(a,b) ==> split(<a,b>, c)"; |
|
87 by (stac split 1); |
|
88 by (resolve_tac prems 1); |
|
89 val splitI = result(); |
|
90 |
|
91 val prems = goalw Prod.thy [split_def] |
|
92 "[| split(p,c); !!x y. [| p = <x,y>; c(x,y) |] ==> Q |] ==> Q"; |
|
93 by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); |
|
94 val splitE = result(); |
|
95 |
|
96 goal Prod.thy "R(split(p,c)) = (! x y. p = <x,y> --> R(c(x,y)))"; |
|
97 by (stac surjective_pairing 1); |
86 by (stac surjective_pairing 1); |
98 by (stac split 1); |
87 by (stac split 1); |
99 by (fast_tac (HOL_cs addSEs [Pair_inject]) 1); |
88 by (fast_tac (HOL_cs addSEs [Pair_inject]) 1); |
100 val expand_split = result(); |
89 val expand_split = result(); |
|
90 |
|
91 (** split used as a logical connective or set former **) |
|
92 |
|
93 (*These rules are for use with fast_tac. |
|
94 Could instead call simp_tac/asm_full_simp_tac using split as rewrite.*) |
|
95 |
|
96 goal Prod.thy "!!a b c. c(a,b) ==> split(c, <a,b>)"; |
|
97 by (asm_simp_tac pair_ss 1); |
|
98 val splitI = result(); |
|
99 |
|
100 val prems = goalw Prod.thy [split_def] |
|
101 "[| split(c,p); !!x y. [| p = <x,y>; c(x,y) |] ==> Q |] ==> Q"; |
|
102 by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); |
|
103 val splitE = result(); |
|
104 |
|
105 goal Prod.thy "!!a b c. z: c(a,b) ==> z: split(c, <a,b>)"; |
|
106 by (asm_simp_tac pair_ss 1); |
|
107 val mem_splitI = result(); |
|
108 |
|
109 val prems = goalw Prod.thy [split_def] |
|
110 "[| z: split(c,p); !!x y. [| p = <x,y>; z: c(x,y) |] ==> Q |] ==> Q"; |
|
111 by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); |
|
112 val mem_splitE = result(); |
101 |
113 |
102 (*** prod_fun -- action of the product functor upon functions ***) |
114 (*** prod_fun -- action of the product functor upon functions ***) |
103 |
115 |
104 goalw Prod.thy [prod_fun_def] "prod_fun(f,g,<a,b>) = <f(a),g(b)>"; |
116 goalw Prod.thy [prod_fun_def] "prod_fun(f,g,<a,b>) = <f(a),g(b)>"; |
105 by (rtac split 1); |
117 by (rtac split 1); |