Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
authorkrauss
Thu, 12 Apr 2012 23:07:01 +0200
changeset 47445 69e96e5500df
parent 47444 d21c95af2df7
child 47446 ed0795caec95
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
src/HOL/Hahn_Banach/Hahn_Banach.thy
src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy
src/HOL/Hahn_Banach/Subspace.thy
src/HOL/Library/BigO.thy
src/HOL/Library/Set_Algebras.thy
src/HOL/Metis_Examples/Big_O.thy
src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
--- a/src/HOL/Hahn_Banach/Hahn_Banach.thy	Thu Apr 12 22:55:11 2012 +0200
+++ b/src/HOL/Hahn_Banach/Hahn_Banach.thy	Thu Apr 12 23:07:01 2012 +0200
@@ -151,12 +151,12 @@
         qed
       qed
 
-      def H' \<equiv> "H \<oplus> lin x'"
+      def H' \<equiv> "H + lin x'"
         -- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *}
       have HH': "H \<unlhd> H'"
       proof (unfold H'_def)
         from x'E have "vectorspace (lin x')" ..
-        with H show "H \<unlhd> H \<oplus> lin x'" ..
+        with H show "H \<unlhd> H + lin x'" ..
       qed
 
       obtain xi where
--- a/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy	Thu Apr 12 22:55:11 2012 +0200
+++ b/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy	Thu Apr 12 23:07:01 2012 +0200
@@ -90,7 +90,7 @@
 lemma h'_lf:
   assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
       SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
-    and H'_def: "H' \<equiv> H \<oplus> lin x0"
+    and H'_def: "H' \<equiv> H + lin x0"
     and HE: "H \<unlhd> E"
   assumes "linearform H h"
   assumes x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
@@ -106,7 +106,7 @@
     proof (unfold H'_def)
       from `x0 \<in> E`
       have "lin x0 \<unlhd> E" ..
-      with HE show "vectorspace (H \<oplus> lin x0)" using E ..
+      with HE show "vectorspace (H + lin x0)" using E ..
     qed
     {
       fix x1 x2 assume x1: "x1 \<in> H'" and x2: "x2 \<in> H'"
@@ -194,7 +194,7 @@
 lemma h'_norm_pres:
   assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
       SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
-    and H'_def: "H' \<equiv> H \<oplus> lin x0"
+    and H'_def: "H' \<equiv> H + lin x0"
     and x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
   assumes E: "vectorspace E" and HE: "subspace H E"
     and "seminorm E p" and "linearform H h"
--- a/src/HOL/Hahn_Banach/Subspace.thy	Thu Apr 12 22:55:11 2012 +0200
+++ b/src/HOL/Hahn_Banach/Subspace.thy	Thu Apr 12 23:07:01 2012 +0200
@@ -228,38 +228,38 @@
   set of all sums of elements from @{text U} and @{text V}.
 *}
 
-lemma sum_def: "U \<oplus> V = {u + v | u v. u \<in> U \<and> v \<in> V}"
+lemma sum_def: "U + V = {u + v | u v. u \<in> U \<and> v \<in> V}"
   unfolding set_plus_def by auto
 
 lemma sumE [elim]:
-    "x \<in> U \<oplus> V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
+    "x \<in> U + V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
   unfolding sum_def by blast
 
 lemma sumI [intro]:
-    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U \<oplus> V"
+    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U + V"
   unfolding sum_def by blast
 
 lemma sumI' [intro]:
-    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U \<oplus> V"
+    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U + V"
   unfolding sum_def by blast
 
-text {* @{text U} is a subspace of @{text "U \<oplus> V"}. *}
+text {* @{text U} is a subspace of @{text "U + V"}. *}
 
 lemma subspace_sum1 [iff]:
   assumes "vectorspace U" "vectorspace V"
-  shows "U \<unlhd> U \<oplus> V"
+  shows "U \<unlhd> U + V"
 proof -
   interpret vectorspace U by fact
   interpret vectorspace V by fact
   show ?thesis
   proof
     show "U \<noteq> {}" ..
-    show "U \<subseteq> U \<oplus> V"
+    show "U \<subseteq> U + V"
     proof
       fix x assume x: "x \<in> U"
       moreover have "0 \<in> V" ..
-      ultimately have "x + 0 \<in> U \<oplus> V" ..
-      with x show "x \<in> U \<oplus> V" by simp
+      ultimately have "x + 0 \<in> U + V" ..
+      with x show "x \<in> U + V" by simp
     qed
     fix x y assume x: "x \<in> U" and "y \<in> U"
     then show "x + y \<in> U" by simp
@@ -271,30 +271,30 @@
 
 lemma sum_subspace [intro?]:
   assumes "subspace U E" "vectorspace E" "subspace V E"
-  shows "U \<oplus> V \<unlhd> E"
+  shows "U + V \<unlhd> E"
 proof -
   interpret subspace U E by fact
   interpret vectorspace E by fact
   interpret subspace V E by fact
   show ?thesis
   proof
-    have "0 \<in> U \<oplus> V"
+    have "0 \<in> U + V"
     proof
       show "0 \<in> U" using `vectorspace E` ..
       show "0 \<in> V" using `vectorspace E` ..
       show "(0::'a) = 0 + 0" by simp
     qed
-    then show "U \<oplus> V \<noteq> {}" by blast
-    show "U \<oplus> V \<subseteq> E"
+    then show "U + V \<noteq> {}" by blast
+    show "U + V \<subseteq> E"
     proof
-      fix x assume "x \<in> U \<oplus> V"
+      fix x assume "x \<in> U + V"
       then obtain u v where "x = u + v" and
         "u \<in> U" and "v \<in> V" ..
       then show "x \<in> E" by simp
     qed
   next
-    fix x y assume x: "x \<in> U \<oplus> V" and y: "y \<in> U \<oplus> V"
-    show "x + y \<in> U \<oplus> V"
+    fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V"
+    show "x + y \<in> U + V"
     proof -
       from x obtain ux vx where "x = ux + vx" and "ux \<in> U" and "vx \<in> V" ..
       moreover
@@ -306,7 +306,7 @@
         using x y by (simp_all add: add_ac)
       then show ?thesis ..
     qed
-    fix a show "a \<cdot> x \<in> U \<oplus> V"
+    fix a show "a \<cdot> x \<in> U + V"
     proof -
       from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" ..
       then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V"
@@ -319,7 +319,7 @@
 text{* The sum of two subspaces is a vectorspace. *}
 
 lemma sum_vs [intro?]:
-    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U \<oplus> V)"
+    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U + V)"
   by (rule subspace.vectorspace) (rule sum_subspace)
 
 
@@ -481,7 +481,7 @@
 proof -
   interpret vectorspace E by fact
   interpret subspace H E by fact
-  from x y x' have "x \<in> H \<oplus> lin x'" by auto
+  from x y x' have "x \<in> H + lin x'" by auto
   have "\<exists>!p. (\<lambda>(y, a). x = y + a \<cdot> x' \<and> y \<in> H) p" (is "\<exists>!p. ?P p")
   proof (rule ex_ex1I)
     from x y show "\<exists>p. ?P p" by blast
--- a/src/HOL/Library/BigO.thy	Thu Apr 12 22:55:11 2012 +0200
+++ b/src/HOL/Library/BigO.thy	Thu Apr 12 23:07:01 2012 +0200
@@ -92,7 +92,7 @@
   by (auto simp add: bigo_def) 
 
 lemma bigo_plus_self_subset [intro]: 
-  "O(f) \<oplus> O(f) <= O(f)"
+  "O(f) + O(f) <= O(f)"
   apply (auto simp add: bigo_alt_def set_plus_def)
   apply (rule_tac x = "c + ca" in exI)
   apply auto
@@ -104,14 +104,14 @@
   apply force
 done
 
-lemma bigo_plus_idemp [simp]: "O(f) \<oplus> O(f) = O(f)"
+lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)"
   apply (rule equalityI)
   apply (rule bigo_plus_self_subset)
   apply (rule set_zero_plus2) 
   apply (rule bigo_zero)
   done
 
-lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \<oplus> O(g)"
+lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) + O(g)"
   apply (rule subsetI)
   apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def)
   apply (subst bigo_pos_const [symmetric])+
@@ -153,15 +153,15 @@
   apply simp
   done
 
-lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A \<oplus> B <= O(f)"
-  apply (subgoal_tac "A \<oplus> B <= O(f) \<oplus> O(f)")
+lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A + B <= O(f)"
+  apply (subgoal_tac "A + B <= O(f) + O(f)")
   apply (erule order_trans)
   apply simp
   apply (auto del: subsetI simp del: bigo_plus_idemp)
   done
 
 lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL x. 0 <= g x ==> 
-    O(f + g) = O(f) \<oplus> O(g)"
+    O(f + g) = O(f) + O(g)"
   apply (rule equalityI)
   apply (rule bigo_plus_subset)
   apply (simp add: bigo_alt_def set_plus_def func_plus)
@@ -273,12 +273,12 @@
 lemma bigo_abs5: "f =o O(g) ==> (%x. abs(f x)) =o O(g)" 
   by (unfold bigo_def, auto)
 
-lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) \<oplus> O(h)"
+lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) + O(h)"
 proof -
   assume "f : g +o O(h)"
-  also have "... <= O(g) \<oplus> O(h)"
+  also have "... <= O(g) + O(h)"
     by (auto del: subsetI)
-  also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))"
+  also have "... = O(%x. abs(g x)) + O(%x. abs(h x))"
     apply (subst bigo_abs3 [symmetric])+
     apply (rule refl)
     done
@@ -287,13 +287,13 @@
   finally have "f : ...".
   then have "O(f) <= ..."
     by (elim bigo_elt_subset)
-  also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))"
+  also have "... = O(%x. abs(g x)) + O(%x. abs(h x))"
     by (rule bigo_plus_eq, auto)
   finally show ?thesis
     by (simp add: bigo_abs3 [symmetric])
 qed
 
-lemma bigo_mult [intro]: "O(f)\<otimes>O(g) <= O(f * g)"
+lemma bigo_mult [intro]: "O(f)*O(g) <= O(f * g)"
   apply (rule subsetI)
   apply (subst bigo_def)
   apply (auto simp add: bigo_alt_def set_times_def func_times)
@@ -369,7 +369,7 @@
   done
 
 lemma bigo_mult7: "ALL x. f x ~= 0 ==>
-    O(f * g) <= O(f::'a => ('b::linordered_field)) \<otimes> O(g)"
+    O(f * g) <= O(f::'a => ('b::linordered_field)) * O(g)"
   apply (subst bigo_mult6)
   apply assumption
   apply (rule set_times_mono3)
@@ -377,7 +377,7 @@
   done
 
 lemma bigo_mult8: "ALL x. f x ~= 0 ==>
-    O(f * g) = O(f::'a => ('b::linordered_field)) \<otimes> O(g)"
+    O(f * g) = O(f::'a => ('b::linordered_field)) * O(g)"
   apply (rule equalityI)
   apply (erule bigo_mult7)
   apply (rule bigo_mult)
@@ -402,9 +402,9 @@
   show "f +o O(g) <= O(g)"
   proof -
     have "f : O(f)" by auto
-    then have "f +o O(g) <= O(f) \<oplus> O(g)"
+    then have "f +o O(g) <= O(f) + O(g)"
       by (auto del: subsetI)
-    also have "... <= O(g) \<oplus> O(g)"
+    also have "... <= O(g) + O(g)"
     proof -
       from a have "O(f) <= O(g)" by (auto del: subsetI)
       thus ?thesis by (auto del: subsetI)
@@ -656,7 +656,7 @@
 subsection {* Misc useful stuff *}
 
 lemma bigo_useful_intro: "A <= O(f) ==> B <= O(f) ==>
-  A \<oplus> B <= O(f)"
+  A + B <= O(f)"
   apply (subst bigo_plus_idemp [symmetric])
   apply (rule set_plus_mono2)
   apply assumption+
--- a/src/HOL/Library/Set_Algebras.thy	Thu Apr 12 22:55:11 2012 +0200
+++ b/src/HOL/Library/Set_Algebras.thy	Thu Apr 12 23:07:01 2012 +0200
@@ -34,14 +34,6 @@
 
 end
 
-
-text {* Legacy syntax: *}
-
-abbreviation (input) set_plus :: "'a::plus set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "\<oplus>" 65) where
-  "A \<oplus> B \<equiv> A + B"
-abbreviation (input) set_times :: "'a::times set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "\<otimes>" 70) where
-  "A \<otimes> B \<equiv> A * B"
-
 instantiation set :: (zero) zero
 begin
 
@@ -95,14 +87,14 @@
 instance set :: (comm_monoid_mult) comm_monoid_mult
 by default (simp_all add: set_times_def)
 
-lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C \<oplus> D"
+lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C + D"
   by (auto simp add: set_plus_def)
 
 lemma set_plus_intro2 [intro]: "b : C ==> a + b : a +o C"
   by (auto simp add: elt_set_plus_def)
 
-lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) \<oplus>
-    (b +o D) = (a + b) +o (C \<oplus> D)"
+lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) +
+    (b +o D) = (a + b) +o (C + D)"
   apply (auto simp add: elt_set_plus_def set_plus_def add_ac)
    apply (rule_tac x = "ba + bb" in exI)
   apply (auto simp add: add_ac)
@@ -114,8 +106,8 @@
     (a + b) +o C"
   by (auto simp add: elt_set_plus_def add_assoc)
 
-lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) \<oplus> C =
-    a +o (B \<oplus> C)"
+lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) + C =
+    a +o (B + C)"
   apply (auto simp add: elt_set_plus_def set_plus_def)
    apply (blast intro: add_ac)
   apply (rule_tac x = "a + aa" in exI)
@@ -126,8 +118,8 @@
    apply (auto simp add: add_ac)
   done
 
-theorem set_plus_rearrange4: "C \<oplus> ((a::'a::comm_monoid_add) +o D) =
-    a +o (C \<oplus> D)"
+theorem set_plus_rearrange4: "C + ((a::'a::comm_monoid_add) +o D) =
+    a +o (C + D)"
   apply (auto simp add: elt_set_plus_def set_plus_def add_ac)
    apply (rule_tac x = "aa + ba" in exI)
    apply (auto simp add: add_ac)
@@ -140,17 +132,17 @@
   by (auto simp add: elt_set_plus_def)
 
 lemma set_plus_mono2 [intro]: "(C::('a::plus) set) <= D ==> E <= F ==>
-    C \<oplus> E <= D \<oplus> F"
+    C + E <= D + F"
   by (auto simp add: set_plus_def)
 
-lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C \<oplus> D"
+lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C + D"
   by (auto simp add: elt_set_plus_def set_plus_def)
 
 lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) : C ==>
-    a +o D <= D \<oplus> C"
+    a +o D <= D + C"
   by (auto simp add: elt_set_plus_def set_plus_def add_ac)
 
-lemma set_plus_mono5: "a:C ==> B <= D ==> a +o B <= C \<oplus> D"
+lemma set_plus_mono5: "a:C ==> B <= D ==> a +o B <= C + D"
   apply (subgoal_tac "a +o B <= a +o D")
    apply (erule order_trans)
    apply (erule set_plus_mono3)
@@ -163,21 +155,21 @@
   apply auto
   done
 
-lemma set_plus_mono2_b: "C <= D ==> E <= F ==> x : C \<oplus> E ==>
-    x : D \<oplus> F"
+lemma set_plus_mono2_b: "C <= D ==> E <= F ==> x : C + E ==>
+    x : D + F"
   apply (frule set_plus_mono2)
    prefer 2
    apply force
   apply assumption
   done
 
-lemma set_plus_mono3_b: "a : C ==> x : a +o D ==> x : C \<oplus> D"
+lemma set_plus_mono3_b: "a : C ==> x : a +o D ==> x : C + D"
   apply (frule set_plus_mono3)
   apply auto
   done
 
 lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C ==>
-    x : a +o D ==> x : D \<oplus> C"
+    x : a +o D ==> x : D + C"
   apply (frule set_plus_mono4)
   apply auto
   done
@@ -185,7 +177,7 @@
 lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C"
   by (auto simp add: elt_set_plus_def)
 
-lemma set_zero_plus2: "(0::'a::comm_monoid_add) : A ==> B <= A \<oplus> B"
+lemma set_zero_plus2: "(0::'a::comm_monoid_add) : A ==> B <= A + B"
   apply (auto simp add: set_plus_def)
   apply (rule_tac x = 0 in bexI)
    apply (rule_tac x = x in bexI)
@@ -206,14 +198,14 @@
   by (rule iffI, rule set_minus_imp_plus, assumption, rule set_plus_imp_minus,
     assumption)
 
-lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C \<otimes> D"
+lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C * D"
   by (auto simp add: set_times_def)
 
 lemma set_times_intro2 [intro!]: "b : C ==> a * b : a *o C"
   by (auto simp add: elt_set_times_def)
 
-lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) \<otimes>
-    (b *o D) = (a * b) *o (C \<otimes> D)"
+lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) *
+    (b *o D) = (a * b) *o (C * D)"
   apply (auto simp add: elt_set_times_def set_times_def)
    apply (rule_tac x = "ba * bb" in exI)
    apply (auto simp add: mult_ac)
@@ -225,8 +217,8 @@
     (a * b) *o C"
   by (auto simp add: elt_set_times_def mult_assoc)
 
-lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) \<otimes> C =
-    a *o (B \<otimes> C)"
+lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) * C =
+    a *o (B * C)"
   apply (auto simp add: elt_set_times_def set_times_def)
    apply (blast intro: mult_ac)
   apply (rule_tac x = "a * aa" in exI)
@@ -237,8 +229,8 @@
    apply (auto simp add: mult_ac)
   done
 
-theorem set_times_rearrange4: "C \<otimes> ((a::'a::comm_monoid_mult) *o D) =
-    a *o (C \<otimes> D)"
+theorem set_times_rearrange4: "C * ((a::'a::comm_monoid_mult) *o D) =
+    a *o (C * D)"
   apply (auto simp add: elt_set_times_def set_times_def
     mult_ac)
    apply (rule_tac x = "aa * ba" in exI)
@@ -252,17 +244,17 @@
   by (auto simp add: elt_set_times_def)
 
 lemma set_times_mono2 [intro]: "(C::('a::times) set) <= D ==> E <= F ==>
-    C \<otimes> E <= D \<otimes> F"
+    C * E <= D * F"
   by (auto simp add: set_times_def)
 
-lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C \<otimes> D"
+lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C * D"
   by (auto simp add: elt_set_times_def set_times_def)
 
 lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C ==>
-    a *o D <= D \<otimes> C"
+    a *o D <= D * C"
   by (auto simp add: elt_set_times_def set_times_def mult_ac)
 
-lemma set_times_mono5: "a:C ==> B <= D ==> a *o B <= C \<otimes> D"
+lemma set_times_mono5: "a:C ==> B <= D ==> a *o B <= C * D"
   apply (subgoal_tac "a *o B <= a *o D")
    apply (erule order_trans)
    apply (erule set_times_mono3)
@@ -275,21 +267,21 @@
   apply auto
   done
 
-lemma set_times_mono2_b: "C <= D ==> E <= F ==> x : C \<otimes> E ==>
-    x : D \<otimes> F"
+lemma set_times_mono2_b: "C <= D ==> E <= F ==> x : C * E ==>
+    x : D * F"
   apply (frule set_times_mono2)
    prefer 2
    apply force
   apply assumption
   done
 
-lemma set_times_mono3_b: "a : C ==> x : a *o D ==> x : C \<otimes> D"
+lemma set_times_mono3_b: "a : C ==> x : a *o D ==> x : C * D"
   apply (frule set_times_mono3)
   apply auto
   done
 
 lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) : C ==>
-    x : a *o D ==> x : D \<otimes> C"
+    x : a *o D ==> x : D * C"
   apply (frule set_times_mono4)
   apply auto
   done
@@ -301,16 +293,16 @@
     (a * b) +o (a *o C)"
   by (auto simp add: elt_set_plus_def elt_set_times_def ring_distribs)
 
-lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B \<oplus> C) =
-    (a *o B) \<oplus> (a *o C)"
+lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B + C) =
+    (a *o B) + (a *o C)"
   apply (auto simp add: set_plus_def elt_set_times_def ring_distribs)
    apply blast
   apply (rule_tac x = "b + bb" in exI)
   apply (auto simp add: ring_distribs)
   done
 
-lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) \<otimes> D <=
-    a *o D \<oplus> C \<otimes> D"
+lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) * D <=
+    a *o D + C * D"
   apply (auto simp add:
     elt_set_plus_def elt_set_times_def set_times_def
     set_plus_def ring_distribs)
@@ -330,7 +322,7 @@
   by (auto simp add: elt_set_times_def)
 
 lemma set_plus_image:
-  fixes S T :: "'n::semigroup_add set" shows "S \<oplus> T = (\<lambda>(x, y). x + y) ` (S \<times> T)"
+  fixes S T :: "'n::semigroup_add set" shows "S + T = (\<lambda>(x, y). x + y) ` (S \<times> T)"
   unfolding set_plus_def by (fastforce simp: image_iff)
 
 lemma set_setsum_alt:
@@ -339,7 +331,7 @@
     (is "_ = ?setsum I")
 using fin proof induct
   case (insert x F)
-  have "setsum S (insert x F) = S x \<oplus> ?setsum F"
+  have "setsum S (insert x F) = S x + ?setsum F"
     using insert.hyps by auto
   also have "...= {s x + setsum s F |s. \<forall> i\<in>insert x F. s i \<in> S i}"
     unfolding set_plus_def
@@ -355,8 +347,8 @@
 
 lemma setsum_set_cond_linear:
   fixes f :: "('a::comm_monoid_add) set \<Rightarrow> ('b::comm_monoid_add) set"
-  assumes [intro!]: "\<And>A B. P A  \<Longrightarrow> P B  \<Longrightarrow> P (A \<oplus> B)" "P {0}"
-    and f: "\<And>A B. P A  \<Longrightarrow> P B \<Longrightarrow> f (A \<oplus> B) = f A \<oplus> f B" "f {0} = {0}"
+  assumes [intro!]: "\<And>A B. P A  \<Longrightarrow> P B  \<Longrightarrow> P (A + B)" "P {0}"
+    and f: "\<And>A B. P A  \<Longrightarrow> P B \<Longrightarrow> f (A + B) = f A + f B" "f {0} = {0}"
   assumes all: "\<And>i. i \<in> I \<Longrightarrow> P (S i)"
   shows "f (setsum S I) = setsum (f \<circ> S) I"
 proof cases
@@ -372,7 +364,7 @@
 
 lemma setsum_set_linear:
   fixes f :: "('a::comm_monoid_add) set => ('b::comm_monoid_add) set"
-  assumes "\<And>A B. f(A) \<oplus> f(B) = f(A \<oplus> B)" "f {0} = {0}"
+  assumes "\<And>A B. f(A) + f(B) = f(A + B)" "f {0} = {0}"
   shows "f (setsum S I) = setsum (f \<circ> S) I"
   using setsum_set_cond_linear[of "\<lambda>x. True" f I S] assms by auto
 
--- a/src/HOL/Metis_Examples/Big_O.thy	Thu Apr 12 22:55:11 2012 +0200
+++ b/src/HOL/Metis_Examples/Big_O.thy	Thu Apr 12 23:07:01 2012 +0200
@@ -146,17 +146,17 @@
 by (auto simp add: bigo_def)
 
 lemma bigo_plus_self_subset [intro]:
-  "O(f) \<oplus> O(f) <= O(f)"
+  "O(f) + O(f) <= O(f)"
 apply (auto simp add: bigo_alt_def set_plus_def)
 apply (rule_tac x = "c + ca" in exI)
 apply auto
 apply (simp add: ring_distribs func_plus)
 by (metis order_trans abs_triangle_ineq add_mono)
 
-lemma bigo_plus_idemp [simp]: "O(f) \<oplus> O(f) = O(f)"
+lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)"
 by (metis bigo_plus_self_subset bigo_zero set_eq_subset set_zero_plus2)
 
-lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \<oplus> O(g)"
+lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) + O(g)"
 apply (rule subsetI)
 apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def)
 apply (subst bigo_pos_const [symmetric])+
@@ -187,10 +187,10 @@
  apply (metis abs_triangle_ineq mult_le_cancel_left_pos)
 by (metis abs_ge_zero abs_of_pos zero_le_mult_iff)
 
-lemma bigo_plus_subset2 [intro]: "A <= O(f) \<Longrightarrow> B <= O(f) \<Longrightarrow> A \<oplus> B <= O(f)"
+lemma bigo_plus_subset2 [intro]: "A <= O(f) \<Longrightarrow> B <= O(f) \<Longrightarrow> A + B <= O(f)"
 by (metis bigo_plus_idemp set_plus_mono2)
 
-lemma bigo_plus_eq: "\<forall>x. 0 <= f x \<Longrightarrow> \<forall>x. 0 <= g x \<Longrightarrow> O(f + g) = O(f) \<oplus> O(g)"
+lemma bigo_plus_eq: "\<forall>x. 0 <= f x \<Longrightarrow> \<forall>x. 0 <= g x \<Longrightarrow> O(f + g) = O(f) + O(g)"
 apply (rule equalityI)
 apply (rule bigo_plus_subset)
 apply (simp add: bigo_alt_def set_plus_def func_plus)
@@ -284,25 +284,25 @@
 lemma bigo_abs5: "f =o O(g) \<Longrightarrow> (\<lambda>x. abs(f x)) =o O(g)"
 by (unfold bigo_def, auto)
 
-lemma bigo_elt_subset2 [intro]: "f : g +o O(h) \<Longrightarrow> O(f) <= O(g) \<oplus> O(h)"
+lemma bigo_elt_subset2 [intro]: "f : g +o O(h) \<Longrightarrow> O(f) <= O(g) + O(h)"
 proof -
   assume "f : g +o O(h)"
-  also have "... <= O(g) \<oplus> O(h)"
+  also have "... <= O(g) + O(h)"
     by (auto del: subsetI)
-  also have "... = O(\<lambda>x. abs(g x)) \<oplus> O(\<lambda>x. abs(h x))"
+  also have "... = O(\<lambda>x. abs(g x)) + O(\<lambda>x. abs(h x))"
     by (metis bigo_abs3)
   also have "... = O((\<lambda>x. abs(g x)) + (\<lambda>x. abs(h x)))"
     by (rule bigo_plus_eq [symmetric], auto)
   finally have "f : ...".
   then have "O(f) <= ..."
     by (elim bigo_elt_subset)
-  also have "... = O(\<lambda>x. abs(g x)) \<oplus> O(\<lambda>x. abs(h x))"
+  also have "... = O(\<lambda>x. abs(g x)) + O(\<lambda>x. abs(h x))"
     by (rule bigo_plus_eq, auto)
   finally show ?thesis
     by (simp add: bigo_abs3 [symmetric])
 qed
 
-lemma bigo_mult [intro]: "O(f) \<otimes> O(g) <= O(f * g)"
+lemma bigo_mult [intro]: "O(f) * O(g) <= O(f * g)"
 apply (rule subsetI)
 apply (subst bigo_def)
 apply (auto simp del: abs_mult mult_ac
@@ -358,14 +358,14 @@
 declare bigo_mult6 [simp]
 
 lemma bigo_mult7:
-"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) \<le> O(f\<Colon>'a \<Rightarrow> ('b\<Colon>linordered_field)) \<otimes> O(g)"
+"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) \<le> O(f\<Colon>'a \<Rightarrow> ('b\<Colon>linordered_field)) * O(g)"
 by (metis bigo_refl bigo_mult6 set_times_mono3)
 
 declare bigo_mult6 [simp del]
 declare bigo_mult7 [intro!]
 
 lemma bigo_mult8:
-"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = O(f\<Colon>'a \<Rightarrow> ('b\<Colon>linordered_field)) \<otimes> O(g)"
+"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = O(f\<Colon>'a \<Rightarrow> ('b\<Colon>linordered_field)) * O(g)"
 by (metis bigo_mult bigo_mult7 order_antisym_conv)
 
 lemma bigo_minus [intro]: "f : O(g) \<Longrightarrow> - f : O(g)"
@@ -575,7 +575,7 @@
 subsection {* Misc useful stuff *}
 
 lemma bigo_useful_intro: "A <= O(f) \<Longrightarrow> B <= O(f) \<Longrightarrow>
-  A \<oplus> B <= O(f)"
+  A + B <= O(f)"
   apply (subst bigo_plus_idemp [symmetric])
   apply (rule set_plus_mono2)
   apply assumption+
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Thu Apr 12 22:55:11 2012 +0200
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Thu Apr 12 23:07:01 2012 +0200
@@ -5428,13 +5428,13 @@
 
 lemma closure_sum:
   fixes S T :: "('n::euclidean_space) set"
-  shows "closure S \<oplus> closure T \<subseteq> closure (S \<oplus> T)"
+  shows "closure S + closure T \<subseteq> closure (S + T)"
 proof-
-  have "(closure S) \<oplus> (closure T) = (\<lambda>(x,y). x + y) ` (closure S \<times> closure T)"
+  have "(closure S) + (closure T) = (\<lambda>(x,y). x + y) ` (closure S \<times> closure T)"
     by (simp add: set_plus_image)
   also have "... = (\<lambda>(x,y). x + y) ` closure (S \<times> T)"
     using closure_direct_sum by auto
-  also have "... \<subseteq> closure (S \<oplus> T)"
+  also have "... \<subseteq> closure (S + T)"
     using fst_snd_linear closure_linear_image[of "(\<lambda>(x,y). x + y)" "S \<times> T"]
     by (auto simp: set_plus_image)
   finally show ?thesis
@@ -5444,7 +5444,7 @@
 lemma convex_oplus:
 fixes S T :: "('n::euclidean_space) set"
 assumes "convex S" "convex T"
-shows "convex (S \<oplus> T)"
+shows "convex (S + T)"
 proof-
 have "{x + y |x y. x : S & y : T} = {c. EX a:S. EX b:T. c = a + b}" by auto
 thus ?thesis unfolding set_plus_def using convex_sums[of S T] assms by auto
@@ -5452,13 +5452,13 @@
 
 lemma convex_hull_sum:
 fixes S T :: "('n::euclidean_space) set"
-shows "convex hull (S \<oplus> T) = (convex hull S) \<oplus> (convex hull T)"
+shows "convex hull (S + T) = (convex hull S) + (convex hull T)"
 proof-
-have "(convex hull S) \<oplus> (convex hull T) =
+have "(convex hull S) + (convex hull T) =
       (%(x,y). x + y) ` ((convex hull S) <*> (convex hull T))"
    by (simp add: set_plus_image)
 also have "... = (%(x,y). x + y) ` (convex hull (S <*> T))" using convex_hull_direct_sum by auto
-also have "...= convex hull (S \<oplus> T)" using fst_snd_linear linear_conv_bounded_linear
+also have "...= convex hull (S + T)" using fst_snd_linear linear_conv_bounded_linear
    convex_hull_linear_image[of "(%(x,y). x + y)" "S <*> T"] by (auto simp add: set_plus_image)
 finally show ?thesis by auto
 qed
@@ -5466,12 +5466,12 @@
 lemma rel_interior_sum:
 fixes S T :: "('n::euclidean_space) set"
 assumes "convex S" "convex T"
-shows "rel_interior (S \<oplus> T) = (rel_interior S) \<oplus> (rel_interior T)"
+shows "rel_interior (S + T) = (rel_interior S) + (rel_interior T)"
 proof-
-have "(rel_interior S) \<oplus> (rel_interior T) = (%(x,y). x + y) ` (rel_interior S <*> rel_interior T)"
+have "(rel_interior S) + (rel_interior T) = (%(x,y). x + y) ` (rel_interior S <*> rel_interior T)"
    by (simp add: set_plus_image)
 also have "... = (%(x,y). x + y) ` rel_interior (S <*> T)" using rel_interior_direct_sum assms by auto
-also have "...= rel_interior (S \<oplus> T)" using fst_snd_linear convex_direct_sum assms
+also have "...= rel_interior (S + T)" using fst_snd_linear convex_direct_sum assms
    rel_interior_convex_linear_image[of "(%(x,y). x + y)" "S <*> T"] by (auto simp add: set_plus_image)
 finally show ?thesis by auto
 qed
@@ -5507,7 +5507,7 @@
 lemma convex_rel_open_sum:
 fixes S T :: "('n::euclidean_space) set"
 assumes "convex S" "rel_open S" "convex T" "rel_open T"
-shows "convex (S \<oplus> T) & rel_open (S \<oplus> T)"
+shows "convex (S + T) & rel_open (S + T)"
 by (metis assms convex_oplus rel_interior_sum rel_open_def)
 
 lemma convex_hull_finite_union_cones:
@@ -5547,7 +5547,7 @@
 fixes S T :: "('m::euclidean_space) set"
 assumes "convex S" "cone S" "S ~= {}"
 assumes "convex T" "cone T" "T ~= {}"
-shows "convex hull (S Un T) = S \<oplus> T"
+shows "convex hull (S Un T) = S + T"
 proof-
 def I == "{(1::nat),2}"
 def A == "(%i. (if i=(1::nat) then S else T))"
@@ -5556,7 +5556,7 @@
 moreover have "convex hull Union (A ` I) = setsum A I"
     apply (subst convex_hull_finite_union_cones[of I A]) using assms A_def I_def by auto
 moreover have
-  "setsum A I = S \<oplus> T" using A_def I_def
+  "setsum A I = S + T" using A_def I_def
      unfolding set_plus_def apply auto unfolding set_plus_def by auto
 ultimately show ?thesis by auto
 qed