HOL/Sum: added disjoint sum of two sets as the plus operator, since + is
overloaded with an incompatible type
--- a/Sum.ML Fri Aug 19 11:10:56 1994 +0200
+++ b/Sum.ML Fri Aug 19 11:15:01 1994 +0200
@@ -88,7 +88,33 @@
by (fast_tac (HOL_cs addSEs [Inr_inject]) 1);
val Inr_eq = result();
-val sum_cs = set_cs addSEs [Inl_neq_Inr, Inr_neq_Inl]
+(*** Rules for the disjoint sum of two SETS ***)
+
+(** Introduction rules for the injections **)
+
+goalw Sum.thy [sum_def] "!!a A B. a : A ==> Inl(a) : A plus B";
+by (REPEAT (ares_tac [UnI1,imageI] 1));
+val InlI = result();
+
+goalw Sum.thy [sum_def] "!!b A B. b : B ==> Inr(b) : A plus B";
+by (REPEAT (ares_tac [UnI2,imageI] 1));
+val InrI = result();
+
+(** Elimination rules **)
+
+val major::prems = goalw Sum.thy [sum_def]
+ "[| u: A plus B; \
+\ !!x. [| x:A; u=Inl(x) |] ==> P; \
+\ !!y. [| y:B; u=Inr(y) |] ==> P \
+\ |] ==> P";
+by (rtac (major RS UnE) 1);
+by (REPEAT (rtac refl 1
+ ORELSE eresolve_tac (prems@[imageE,ssubst]) 1));
+val plusE = result();
+
+
+val sum_cs = set_cs addSIs [InlI, InrI]
+ addSEs [plusE, Inl_neq_Inr, Inr_neq_Inl]
addSDs [Inl_inject, Inr_inject];
@@ -140,6 +166,8 @@
(fn [prem] => [rtac (prem RS arg_cong) 1]);
+
+
(** Rules for the Part primitive **)
goalw Sum.thy [Part_def]
--- a/Sum.thy Fri Aug 19 11:10:56 1994 +0200
+++ b/Sum.thy Fri Aug 19 11:15:01 1994 +0200
@@ -23,6 +23,9 @@
Inl :: "'a => 'a+'b"
Inr :: "'b => 'a+'b"
sum_case :: "['a=>'c,'b=>'c, 'a+'b] =>'c"
+
+ (*disjoint sum for sets; the operator + is overloaded with wrong type!*)
+ "plus" :: "['a set, 'b set] => ('a+'b)set" (infixr 65)
Part :: "['a set, 'a=>'a] => 'a set"
translations
@@ -32,18 +35,20 @@
Inl_Rep_def "Inl_Rep == (%a. %x y p. x=a & p)"
Inr_Rep_def "Inr_Rep == (%b. %x y p. y=b & ~p)"
- Sum_def "Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}"
+ Sum_def "Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}"
(*faking a type definition...*)
Rep_Sum "Rep_Sum(s): Sum"
Rep_Sum_inverse "Abs_Sum(Rep_Sum(s)) = s"
Abs_Sum_inverse "f: Sum ==> Rep_Sum(Abs_Sum(f)) = f"
(*defining the abstract constants*)
- Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))"
- Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))"
+ Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))"
+ Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))"
sum_case_def "sum_case(f,g,p) == @z. (!x. p=Inl(x) --> z=f(x)) \
\ & (!y. p=Inr(y) --> z=g(y))"
+ sum_def "A plus B == (Inl``A) Un (Inr``B)"
+
(*for selecting out the components of a mutually recursive definition*)
Part_def "Part(A,h) == A Int {x. ? z. x = h(z)}"