Prod.ML
changeset 0 7949f97df77a
child 2 befa4e9f7c90
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Prod.ML	Thu Sep 16 12:21:07 1993 +0200
@@ -0,0 +1,204 @@
+(*  Title: 	HOL/prod
+    ID:         $Id$
+    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1991  University of Cambridge
+
+For prod.thy.  Ordered Pairs, the Cartesian product type, the unit type
+*)
+
+open Prod;
+
+(*This counts as a non-emptiness result for admitting 'a * 'b as a type*)
+goalw Prod.thy [Prod_def] "Pair_Rep(a,b) : Prod";
+by (EVERY1 [rtac CollectI, rtac exI, rtac exI, rtac refl]);
+val ProdI = result();
+
+val [major] = goalw Prod.thy [Pair_Rep_def]
+    "Pair_Rep(a, b) = Pair_Rep(a',b') ==> a=a' & b=b'";
+by (EVERY1 [rtac (major RS fun_cong RS fun_cong RS subst), 
+	    rtac conjI, rtac refl, rtac refl]);
+val Pair_Rep_inject = result();
+
+goal Prod.thy "inj_onto(Abs_Prod,Prod)";
+by (rtac inj_onto_inverseI 1);
+by (etac Abs_Prod_inverse 1);
+val inj_onto_Abs_Prod = result();
+
+val prems = goalw Prod.thy [Pair_def]
+    "[| <a, b> = <a',b'>;  [| a=a';  b=b' |] ==> R |] ==> R";
+by (rtac (inj_onto_Abs_Prod RS inj_ontoD RS Pair_Rep_inject RS conjE) 1);
+by (REPEAT (ares_tac (prems@[ProdI]) 1));
+val Pair_inject = result();
+
+goal Prod.thy "(<a,b> = <a',b'>) = (a=a' & b=b')";
+by (fast_tac (set_cs addIs [Pair_inject]) 1);
+val Pair_eq = result();
+
+goalw Prod.thy [fst_def] "fst(<a,b>) = a";
+by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1);
+val fst_conv = result();
+
+goalw Prod.thy [snd_def] "snd(<a,b>) = b";
+by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1);
+val snd_conv = result();
+
+goalw Prod.thy [Pair_def] "? x y. p = <x,y>";
+by (rtac (rewrite_rule [Prod_def] Rep_Prod RS CollectE) 1);
+by (EVERY1[etac exE, etac exE, rtac exI, rtac exI,
+	   rtac (Rep_Prod_inverse RS sym RS trans),  etac arg_cong]);
+val PairE_lemma = result();
+
+val [prem] = goal Prod.thy "[| !!x y. p = <x,y> ==> Q |] ==> Q";
+by (rtac (PairE_lemma RS exE) 1);
+by (REPEAT (eresolve_tac [prem,exE] 1));
+val PairE = result();
+
+goalw Prod.thy [split_def] "split(<a,b>, c) = c(a,b)";
+by (sstac [fst_conv, snd_conv] 1);
+by (rtac refl 1);
+val split = result();
+
+(*FIXME: split's congruence rule should only simplifies the pair*)
+val pair_ss = set_ss addsimps [fst_conv, snd_conv, split];
+
+goal Prod.thy "p = <fst(p),snd(p)>";
+by (res_inst_tac [("p","p")] PairE 1);
+by(asm_simp_tac pair_ss 1);
+val surjective_pairing = result();
+
+goal Prod.thy "p = split(p, %x y.<x,y>)";
+by (res_inst_tac [("p","p")] PairE 1);
+by(asm_simp_tac pair_ss 1);
+val surjective_pairing2 = result();
+
+(** split used as a logical connective, with result type bool **)
+
+val prems = goal Prod.thy "c(a,b) ==> split(<a,b>, c)";
+by (stac split 1);
+by (resolve_tac prems 1);
+val splitI = result();
+
+val prems = goalw Prod.thy [split_def]
+    "[| split(p,c);  !!x y. [| p = <x,y>;  c(x,y) |] ==> Q |] ==> Q";
+by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
+val splitE = result();
+
+goal Prod.thy "R(split(p,c)) = (! x y. p = <x,y> --> R(c(x,y)))";
+by (stac surjective_pairing 1);
+by (stac split 1);
+by (fast_tac (HOL_cs addSEs [Pair_inject]) 1);
+val expand_split = result();
+
+(*** prod_fun -- action of the product functor upon functions ***)
+
+goalw Prod.thy [prod_fun_def] "prod_fun(f,g,<a,b>) = <f(a),g(b)>";
+by (rtac split 1);
+val prod_fun = result();
+
+goal Prod.thy 
+    "prod_fun(f1 o f2, g1 o g2) = (prod_fun(f1,g1) o prod_fun(f2,g2))";
+by (rtac ext 1);
+by (res_inst_tac [("p","x")] PairE 1);
+by(asm_simp_tac (pair_ss addsimps [prod_fun,o_def]) 1);
+val prod_fun_compose = result();
+
+goal Prod.thy "prod_fun(%x.x, %y.y) = (%z.z)";
+by (rtac ext 1);
+by (res_inst_tac [("p","z")] PairE 1);
+by(asm_simp_tac (pair_ss addsimps [prod_fun]) 1);
+val prod_fun_ident = result();
+
+val prems = goal Prod.thy "<a,b>:r ==> <f(a),g(b)> : prod_fun(f,g)``r";
+by (rtac image_eqI 1);
+by (rtac (prod_fun RS sym) 1);
+by (resolve_tac prems 1);
+val prod_fun_imageI = result();
+
+val major::prems = goal Prod.thy
+    "[| c: prod_fun(f,g)``r;  !!x y. [| c=<f(x),g(y)>;  <x,y>:r |] ==> P  \
+\    |] ==> P";
+by (rtac (major RS imageE) 1);
+by (res_inst_tac [("p","x")] PairE 1);
+by (resolve_tac prems 1);
+by (fast_tac HOL_cs 2);
+by (fast_tac (HOL_cs addIs [prod_fun]) 1);
+val prod_fun_imageE = result();
+
+(*** Disjoint union of a family of sets - Sigma ***)
+
+val SigmaI = prove_goalw Prod.thy [Sigma_def]
+    "[| a:A;  b:B(a) |] ==> <a,b> : Sigma(A,B)"
+ (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]);
+
+(*The general elimination rule*)
+val SigmaE = prove_goalw Prod.thy [Sigma_def]
+    "[| c: Sigma(A,B);  \
+\       !!x y.[| x:A;  y:B(x);  c=<x,y> |] ==> P \
+\    |] ==> P"
+ (fn major::prems=>
+  [ (cut_facts_tac [major] 1),
+    (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]);
+
+(** Elimination of <a,b>:A*B -- introduces no eigenvariables **)
+val SigmaD1 = prove_goal Prod.thy "<a,b> : Sigma(A,B) ==> a : A"
+ (fn [major]=>
+  [ (rtac (major RS SigmaE) 1),
+    (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
+
+val SigmaD2 = prove_goal Prod.thy "<a,b> : Sigma(A,B) ==> b : B(a)"
+ (fn [major]=>
+  [ (rtac (major RS SigmaE) 1),
+    (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
+
+val SigmaE2 = prove_goal Prod.thy
+    "[| <a,b> : Sigma(A,B);    \
+\       [| a:A;  b:B(a) |] ==> P   \
+\    |] ==> P"
+ (fn [major,minor]=>
+  [ (rtac minor 1),
+    (rtac (major RS SigmaD1) 1),
+    (rtac (major RS SigmaD2) 1) ]);
+
+(*** Domain of a relation ***)
+
+val prems = goalw Prod.thy [image_def] "<a,b> : r ==> a : fst``r";
+by (rtac CollectI 1);
+by (rtac bexI 1);
+by (rtac (fst_conv RS sym) 1);
+by (resolve_tac prems 1);
+val fst_imageI = result();
+
+val major::prems = goal Prod.thy
+    "[| a : fst``r;  !!y.[| <a,y> : r |] ==> P |] ==> P"; 
+by (rtac (major RS imageE) 1);
+by (resolve_tac prems 1);
+by (etac ssubst 1);
+by (rtac (surjective_pairing RS subst) 1);
+by (assume_tac 1);
+val fst_imageE = result();
+
+(*** Range of a relation ***)
+
+val prems = goalw Prod.thy [image_def] "<a,b> : r ==> b : snd``r";
+by (rtac CollectI 1);
+by (rtac bexI 1);
+by (rtac (snd_conv RS sym) 1);
+by (resolve_tac prems 1);
+val snd_imageI = result();
+
+val major::prems = goal Prod.thy
+    "[| a : snd``r;  !!y.[| <y,a> : r |] ==> P |] ==> P"; 
+by (rtac (major RS imageE) 1);
+by (resolve_tac prems 1);
+by (etac ssubst 1);
+by (rtac (surjective_pairing RS subst) 1);
+by (assume_tac 1);
+val snd_imageE = result();
+
+(** Exhaustion rule for unit -- a degenerate form of induction **)
+
+goalw Prod.thy [Unity_def]
+    "u = Unity";
+by (stac (rewrite_rule [Unit_def] Rep_Unit RS CollectD RS sym) 1);
+by (rtac (Rep_Unit_inverse RS sym) 1);
+val unit_eq = result();