--- a/src/HOL/Sum.ML Thu Jul 13 23:09:03 2000 +0200
+++ b/src/HOL/Sum.ML Thu Jul 13 23:09:25 2000 +0200
@@ -88,18 +88,18 @@
(** Introduction rules for the injections **)
-Goalw [sum_def] "a : A ==> Inl(a) : A Plus B";
+Goalw [sum_def] "a : A ==> Inl(a) : A <+> B";
by (Blast_tac 1);
qed "InlI";
-Goalw [sum_def] "b : B ==> Inr(b) : A Plus B";
+Goalw [sum_def] "b : B ==> Inr(b) : A <+> B";
by (Blast_tac 1);
qed "InrI";
(** Elimination rules **)
val major::prems = Goalw [sum_def]
- "[| u: A Plus B; \
+ "[| u: A <+> B; \
\ !!x. [| x:A; u=Inl(x) |] ==> P; \
\ !!y. [| y:B; u=Inr(y) |] ==> P \
\ |] ==> P";
--- a/src/HOL/Sum.thy Thu Jul 13 23:09:03 2000 +0200
+++ b/src/HOL/Sum.thy Thu Jul 13 23:09:25 2000 +0200
@@ -31,7 +31,7 @@
Inr :: "'b => 'a + 'b"
(*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
local
@@ -40,7 +40,7 @@
Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))"
Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))"
- sum_def "A Plus B == (Inl``A) Un (Inr``B)"
+ sum_def "A <+> 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)}"