sum.ML
changeset 0 7949f97df77a
child 2 befa4e9f7c90
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/sum.ML	Thu Sep 16 12:21:07 1993 +0200
@@ -0,0 +1,130 @@
+(*  Title: 	HOL/sum
+    ID:         $Id$
+    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1991  University of Cambridge
+
+For sum.ML.  The disjoint sum of two types
+*)
+
+open Sum;
+
+(** Inl_Rep and Inr_Rep: Representations of the constructors **)
+
+(*This counts as a non-emptiness result for admitting 'a+'b as a type*)
+goalw Sum.thy [Sum_def] "Inl_Rep(a) : Sum";
+by (EVERY1 [rtac CollectI, rtac disjI1, rtac exI, rtac refl]);
+val Inl_RepI = result();
+
+goalw Sum.thy [Sum_def] "Inr_Rep(b) : Sum";
+by (EVERY1 [rtac CollectI, rtac disjI2, rtac exI, rtac refl]);
+val Inr_RepI = result();
+
+goal Sum.thy "inj_onto(Abs_Sum,Sum)";
+by (rtac inj_onto_inverseI 1);
+by (etac Abs_Sum_inverse 1);
+val inj_onto_Abs_Sum = result();
+
+(** Distinctness of Inl and Inr **)
+
+goalw Sum.thy [Inl_Rep_def, Inr_Rep_def] "~ (Inl_Rep(a) = Inr_Rep(b))";
+by (EVERY1 [rtac notI,
+	    etac (fun_cong RS fun_cong RS fun_cong RS iffE), 
+	    rtac (notE RS ccontr),  etac (mp RS conjunct2), 
+	    REPEAT o (ares_tac [refl,conjI]) ]);
+val Inl_Rep_not_Inr_Rep = result();
+
+goalw Sum.thy [Inl_def,Inr_def] "~ Inl(a)=Inr(b)";
+by (rtac (inj_onto_Abs_Sum RS inj_onto_contraD) 1);
+by (rtac Inl_Rep_not_Inr_Rep 1);
+by (rtac Inl_RepI 1);
+by (rtac Inr_RepI 1);
+val Inl_not_Inr = result();
+
+val Inl_neq_Inr = standard (Inl_not_Inr RS notE);
+val Inr_neq_Inl = sym RS Inl_neq_Inr;
+
+(** Injectiveness of Inl and Inr **)
+
+val [major] = goalw Sum.thy [Inl_Rep_def] "Inl_Rep(a) = Inl_Rep(c) ==> a=c";
+by (rtac (major RS fun_cong RS fun_cong RS fun_cong RS iffE) 1);
+by (fast_tac HOL_cs 1);
+val Inl_Rep_inject = result();
+
+val [major] = goalw Sum.thy [Inr_Rep_def] "Inr_Rep(b) = Inr_Rep(d) ==> b=d";
+by (rtac (major RS fun_cong RS fun_cong RS fun_cong RS iffE) 1);
+by (fast_tac HOL_cs 1);
+val Inr_Rep_inject = result();
+
+goalw Sum.thy [Inl_def] "inj(Inl)";
+by (rtac injI 1);
+by (etac (inj_onto_Abs_Sum RS inj_ontoD RS Inl_Rep_inject) 1);
+by (rtac Inl_RepI 1);
+by (rtac Inl_RepI 1);
+val inj_Inl = result();
+val Inl_inject = inj_Inl RS injD;
+
+goalw Sum.thy [Inr_def] "inj(Inr)";
+by (rtac injI 1);
+by (etac (inj_onto_Abs_Sum RS inj_ontoD RS Inr_Rep_inject) 1);
+by (rtac Inr_RepI 1);
+by (rtac Inr_RepI 1);
+val inj_Inr = result();
+val Inr_inject = inj_Inr RS injD;
+
+goal Sum.thy "(Inl(x)=Inl(y)) = (x=y)";
+br iffI 1;
+be (rewrite_rule [inj_def] Inl_inject) 1;
+be ssubst 1;
+br refl 1;
+val Inl_inj = result();
+
+goal Sum.thy "(Inr(x)=Inr(y)) = (x=y)";
+br iffI 1;
+be (rewrite_rule [inj_def] Inr_inject) 1;
+be ssubst 1;
+br refl 1;
+val Inr_inj = result();
+
+(** case -- the selection operator for sums **)
+
+goalw Sum.thy [case_def] "case(Inl(x), f, g) = f(x)";
+by (fast_tac (set_cs addIs [select_equality] 
+		     addSEs [make_elim Inl_inject, Inl_neq_Inr]) 1);
+val case_Inl = result();
+
+goalw Sum.thy [case_def] "case(Inr(x), f, g) = g(x)";
+by (fast_tac (set_cs addIs [select_equality] 
+		     addSEs [make_elim Inr_inject, Inr_neq_Inl]) 1);
+val case_Inr = result();
+
+(** Exhaustion rule for sums -- a degenerate form of induction **)
+
+val prems = goalw Sum.thy [Inl_def,Inr_def]
+    "[| !!x::'a. s = Inl(x) ==> P;  !!y::'b. s = Inr(y) ==> P \
+\    |] ==> P";
+by (rtac (rewrite_rule [Sum_def] Rep_Sum RS CollectE) 1);
+by (REPEAT (eresolve_tac [disjE,exE] 1
+     ORELSE EVERY1 [resolve_tac prems, 
+		    etac subst,
+		    rtac (Rep_Sum_inverse RS sym)]));
+val sumE = result();
+
+goal Sum.thy "case(s, %x::'a. f(Inl(x)), %y::'b. f(Inr(y))) = f(s)";
+by (EVERY1 [res_inst_tac [("s","s")] sumE, 
+	    etac ssubst, rtac case_Inl,
+	    etac ssubst, rtac case_Inr]);
+val surjective_sum = result();
+
+goal Sum.thy "R(case(s,f,g)) = \
+\             ((! x. s = Inl(x) --> R(f(x))) & (! y. s = Inr(y) --> R(g(y))))";
+by (rtac sumE 1);
+by (etac ssubst 1);
+by (stac case_Inl 1);
+by (fast_tac (set_cs addSEs [make_elim Inl_inject, Inl_neq_Inr]) 1);
+by (etac ssubst 1);
+by (stac case_Inr 1);
+by (fast_tac (set_cs addSEs [make_elim Inr_inject, Inr_neq_Inl]) 1);
+val expand_case = result();
+
+(*FIXME: case's congruence rule only should simplifies the first argument*)
+val sum_ss = pair_ss addsimps [case_Inl, case_Inr];