--- a/src/HOL/Sum.ML Tue Nov 19 14:20:02 1996 +0100
+++ b/src/HOL/Sum.ML Wed Nov 20 10:32:58 1996 +0100
@@ -91,29 +91,29 @@
(** Introduction rules for the injections **)
-goalw Sum.thy [sum_def] "!!a A B. a : A ==> Inl(a) : A plus B";
+goalw Sum.thy [sum_def] "!!a A B. a : A ==> Inl(a) : A Plus B";
by (Fast_tac 1);
qed "InlI";
-goalw Sum.thy [sum_def] "!!b A B. b : B ==> Inr(b) : A plus B";
+goalw Sum.thy [sum_def] "!!b A B. b : B ==> Inr(b) : A Plus B";
by (Fast_tac 1);
qed "InrI";
(** Elimination rules **)
val major::prems = goalw Sum.thy [sum_def]
- "[| u: A plus B; \
+ "[| 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));
-qed "plusE";
+qed "PlusE";
AddSIs [InlI, InrI];
-AddSEs [plusE];
+AddSEs [PlusE];
(** sum_case -- the selection operator for sums **)
--- a/src/HOL/Sum.thy Tue Nov 19 14:20:02 1996 +0100
+++ b/src/HOL/Sum.thy Wed Nov 20 10:32:58 1996 +0100
@@ -30,7 +30,7 @@
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)
+ Plus :: "['a set, 'b set] => ('a + 'b) set" (infixr 65)
Part :: ['a set, 'b => 'a] => 'a set
translations
@@ -42,7 +42,7 @@
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)"
+ 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)}"