dropped group_simps, ring_simps, field_eq_simps
authorhaftmann
Mon, 26 Apr 2010 11:34:19 +0200
changeset 36350 bc7982c54e37
parent 36349 39be26d1bc28
child 36351 85ee44388f7b
child 36409 d323e7773aa8
dropped group_simps, ring_simps, field_eq_simps
src/HOL/Decision_Procs/Polynomial_List.thy
src/HOL/GCD.thy
src/HOL/Library/Binomial.thy
src/HOL/Library/Formal_Power_Series.thy
src/HOL/Library/Numeral_Type.thy
src/HOL/Library/Polynomial.thy
src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML
src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy
src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Derivative.thy
src/HOL/Multivariate_Analysis/Determinants.thy
src/HOL/Multivariate_Analysis/Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Integration.thy
src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
src/HOL/NSA/HyperDef.thy
src/HOL/Number_Theory/Binomial.thy
src/HOL/Number_Theory/Cong.thy
src/HOL/Number_Theory/Fib.thy
src/HOL/Number_Theory/Residues.thy
src/HOL/SMT/Tools/z3_proof_rules.ML
src/HOL/SMT/Z3.thy
src/HOL/SetInterval.thy
src/HOL/ex/Lagrange.thy
--- a/src/HOL/Decision_Procs/Polynomial_List.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -283,11 +283,11 @@
 apply (drule_tac x = "%m. if m = Suc n then a else i m" in spec, safe)
 apply (drule poly_mult_eq_zero_disj [THEN iffD1], safe)
 apply (drule_tac x = "Suc (length q)" in spec)
-apply (auto simp add: ring_simps)
+apply (auto simp add: field_simps)
 apply (drule_tac x = xa in spec)
-apply (clarsimp simp add: ring_simps)
+apply (clarsimp simp add: field_simps)
 apply (drule_tac x = m in spec)
-apply (auto simp add:ring_simps)
+apply (auto simp add:field_simps)
 done
 lemmas poly_roots_index_lemma1 = conjI [THEN poly_roots_index_lemma0, standard]
 
@@ -327,7 +327,7 @@
 apply (drule_tac x = "a#i" in spec)
 apply (auto simp only: poly_mult List.list.size)
 apply (drule_tac x = xa in spec)
-apply (clarsimp simp add: ring_simps)
+apply (clarsimp simp add: field_simps)
 done
 
 lemmas poly_roots_index_lemma2 = conjI [THEN poly_roots_index_lemma, standard]
@@ -413,7 +413,7 @@
 by (auto intro!: ext)
 
 lemma poly_add_minus_zero_iff: "(poly (p +++ -- q) = poly []) = (poly p = poly q)"
-by (auto simp add: ring_simps poly_add poly_minus_def fun_eq poly_cmult)
+by (auto simp add: field_simps poly_add poly_minus_def fun_eq poly_cmult)
 
 lemma poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
 by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult right_distrib)
--- a/src/HOL/GCD.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/GCD.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -1034,11 +1034,11 @@
     thus "fst (bezw m n) * int m + snd (bezw m n) * int n = int (gcd m n)"
       apply (simp add: bezw_non_0 gcd_non_0_nat)
       apply (erule subst)
-      apply (simp add: ring_simps)
+      apply (simp add: field_simps)
       apply (subst mod_div_equality [of m n, symmetric])
       (* applying simp here undoes the last substitution!
          what is procedure cancel_div_mod? *)
-      apply (simp only: ring_simps zadd_int [symmetric]
+      apply (simp only: field_simps zadd_int [symmetric]
         zmult_int [symmetric])
       done
 qed
@@ -1389,7 +1389,7 @@
   show "lcm (lcm n m) p = lcm n (lcm m p)"
     by (rule lcm_unique_nat [THEN iffD1]) (metis dvd.order_trans lcm_unique_nat)
   show "lcm m n = lcm n m"
-    by (simp add: lcm_nat_def gcd_commute_nat ring_simps)
+    by (simp add: lcm_nat_def gcd_commute_nat field_simps)
 qed
 
 interpretation lcm_int!: abel_semigroup "lcm :: int \<Rightarrow> int \<Rightarrow> int"
--- a/src/HOL/Library/Binomial.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Library/Binomial.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -236,10 +236,10 @@
     have th1: "(\<Prod>n\<in>{1\<Colon>nat..n}. a + of_nat n) =
       (\<Prod>n\<in>{0\<Colon>nat..n - 1}. a + 1 + of_nat n)"
       apply (rule setprod_reindex_cong[where f = "Suc"])
-      using n0 by (auto simp add: expand_fun_eq ring_simps)
+      using n0 by (auto simp add: expand_fun_eq field_simps)
     have ?thesis apply (simp add: pochhammer_def)
     unfolding setprod_insert[OF th0, unfolded eq]
-    using th1 by (simp add: ring_simps)}
+    using th1 by (simp add: field_simps)}
 ultimately show ?thesis by blast
 qed
 
@@ -378,10 +378,10 @@
       by simp
     from n h th0 
     have "fact k * fact (n - k) * (n choose k) = k * (fact h * fact (m - h) * (m choose h)) +  (m - h) * (fact k * fact (m - k) * (m choose k))"
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     also have "\<dots> = (k + (m - h)) * fact m"
       using H[rule_format, OF mn hm'] H[rule_format, OF mn km]
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     finally have ?ths using h n km by simp}
   moreover have "n=0 \<or> k = 0 \<or> k = n \<or> (EX m h. n=Suc m \<and> k = Suc h \<and> h < m)" using kn by presburger
   ultimately show ?ths by blast
@@ -391,13 +391,13 @@
   assumes kn: "k \<le> n" 
   shows "(of_nat (n choose k) :: 'a::field_char_0) = of_nat (fact n) / (of_nat (fact k) * of_nat (fact (n - k)))"
   using binomial_fact_lemma[OF kn]
-  by (simp add: field_eq_simps of_nat_mult [symmetric])
+  by (simp add: field_simps of_nat_mult [symmetric])
 
 lemma binomial_gbinomial: "of_nat (n choose k) = of_nat n gchoose k"
 proof-
   {assume kn: "k > n" 
     from kn binomial_eq_0[OF kn] have ?thesis 
-      by (simp add: gbinomial_pochhammer field_eq_simps
+      by (simp add: gbinomial_pochhammer field_simps
         pochhammer_of_nat_eq_0_iff)}
   moreover
   {assume "k=0" then have ?thesis by simp}
@@ -414,13 +414,13 @@
       apply clarsimp
       apply (presburger)
       apply presburger
-      by (simp add: expand_fun_eq ring_simps of_nat_add[symmetric] del: of_nat_add)
+      by (simp add: expand_fun_eq field_simps of_nat_add[symmetric] del: of_nat_add)
     have th0: "finite {1..n - Suc h}" "finite {n - h .. n}" 
 "{1..n - Suc h} \<inter> {n - h .. n} = {}" and eq3: "{1..n - Suc h} \<union> {n - h .. n} = {1..n}" using h kn by auto
     from eq[symmetric]
     have ?thesis using kn
       apply (simp add: binomial_fact[OF kn, where ?'a = 'a] 
-        gbinomial_pochhammer field_eq_simps pochhammer_Suc_setprod)
+        gbinomial_pochhammer field_simps pochhammer_Suc_setprod)
       apply (simp add: pochhammer_Suc_setprod fact_altdef_nat h of_nat_setprod setprod_timesf[symmetric] eq' del: One_nat_def power_Suc)
       unfolding setprod_Un_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \<Rightarrow> 'a"] eq[unfolded h]
       unfolding mult_assoc[symmetric] 
@@ -449,9 +449,9 @@
   have "?r = ((- 1) ^n * pochhammer (- a) n / of_nat (fact n)) * (of_nat n - (- a + of_nat n))"
     unfolding gbinomial_pochhammer
     pochhammer_Suc fact_Suc of_nat_mult right_diff_distrib power_Suc
-    by (simp add:  field_eq_simps del: of_nat_Suc)
+    by (simp add:  field_simps del: of_nat_Suc)
   also have "\<dots> = ?l" unfolding gbinomial_pochhammer
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
   finally show ?thesis ..
 qed
 
@@ -482,17 +482,17 @@
 
     have "of_nat (fact (Suc k)) * (a gchoose k + (a gchoose (Suc k))) = ((a gchoose Suc h) * of_nat (fact (Suc h)) * of_nat (Suc k)) + (\<Prod>i\<in>{0\<Colon>nat..Suc h}. a - of_nat i)" 
       unfolding h
-      apply (simp add: ring_simps del: fact_Suc)
+      apply (simp add: field_simps del: fact_Suc)
       unfolding gbinomial_mult_fact'
       apply (subst fact_Suc)
       unfolding of_nat_mult 
       apply (subst mult_commute)
       unfolding mult_assoc
       unfolding gbinomial_mult_fact
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     also have "\<dots> = (\<Prod>i\<in>{0..h}. a - of_nat i) * (a + 1)"
       unfolding gbinomial_mult_fact' setprod_nat_ivl_Suc
-      by (simp add: ring_simps h)
+      by (simp add: field_simps h)
     also have "\<dots> = (\<Prod>i\<in>{0..k}. (a + 1) - of_nat i)"
       using eq0
       unfolding h  setprod_nat_ivl_1_Suc
--- a/src/HOL/Library/Formal_Power_Series.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Library/Formal_Power_Series.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -588,7 +588,7 @@
   from k have "real k > - log y x" by simp
   then have "ln y * real k > - ln x" unfolding log_def
     using ln_gt_zero_iff[OF yp] y1
-    by (simp add: minus_divide_left field_simps field_eq_simps del:minus_divide_left[symmetric])
+    by (simp add: minus_divide_left field_simps del:minus_divide_left[symmetric])
   then have "ln y * real k + ln x > 0" by simp
   then have "exp (real k * ln y + ln x) > exp 0"
     by (simp add: mult_ac)
@@ -596,7 +596,7 @@
     unfolding exp_zero exp_add exp_real_of_nat_mult
     exp_ln[OF xp] exp_ln[OF yp] by simp
   then have "x > (1/y)^k" using yp 
-    by (simp add: field_simps field_eq_simps nonzero_power_divide)
+    by (simp add: field_simps nonzero_power_divide)
   then show ?thesis using kp by blast
 qed
 lemma X_nth[simp]: "X$n = (if n = 1 then 1 else 0)" by (simp add: X_def)
@@ -693,7 +693,7 @@
     from th0[symmetric, unfolded nonzero_divide_eq_eq[OF f0]]
     have th1: "setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n} =
       - (f$0) * (inverse f)$n"
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     have "(f * inverse f) $ n = (\<Sum>i = 0..n. f $i * natfun_inverse f (n - i))"
       unfolding fps_mult_nth ifn ..
     also have "\<dots> = f$0 * natfun_inverse f n
@@ -766,7 +766,7 @@
 lemma fps_deriv_nth[simp]: "fps_deriv f $ n = of_nat (n +1) * f $ (n+1)" by (simp add: fps_deriv_def)
 
 lemma fps_deriv_linear[simp]: "fps_deriv (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) = fps_const a * fps_deriv f + fps_const b * fps_deriv g"
-  unfolding fps_eq_iff fps_add_nth  fps_const_mult_left fps_deriv_nth by (simp add: ring_simps)
+  unfolding fps_eq_iff fps_add_nth  fps_const_mult_left fps_deriv_nth by (simp add: field_simps)
 
 lemma fps_deriv_mult[simp]:
   fixes f :: "('a :: comm_ring_1) fps"
@@ -817,7 +817,7 @@
       unfolding s0 s1
       unfolding setsum_addf[symmetric] setsum_right_distrib
       apply (rule setsum_cong2)
-      by (auto simp add: of_nat_diff ring_simps)
+      by (auto simp add: of_nat_diff field_simps)
     finally have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" .}
   then show ?thesis unfolding fps_eq_iff by auto
 qed
@@ -878,7 +878,7 @@
 proof-
   have "fps_deriv f = fps_deriv g \<longleftrightarrow> fps_deriv (f - g) = 0" by simp
   also have "\<dots> \<longleftrightarrow> f - g = fps_const ((f-g)$0)" unfolding fps_deriv_eq_0_iff ..
-  finally show ?thesis by (simp add: ring_simps)
+  finally show ?thesis by (simp add: field_simps)
 qed
 
 lemma fps_deriv_eq_iff_ex: "(fps_deriv f = fps_deriv g) \<longleftrightarrow> (\<exists>(c::'a::{idom,semiring_char_0}). f = fps_const c + g)"
@@ -929,7 +929,7 @@
 qed
 
 lemma fps_deriv_maclauren_0: "(fps_nth_deriv k (f:: ('a::comm_semiring_1) fps)) $ 0 = of_nat (fact k) * f$(k)"
-  by (induct k arbitrary: f) (auto simp add: ring_simps of_nat_mult)
+  by (induct k arbitrary: f) (auto simp add: field_simps of_nat_mult)
 
 subsection {* Powers*}
 
@@ -943,7 +943,7 @@
   case (Suc n)
   note h = Suc.hyps[OF `a$0 = 1`]
   show ?case unfolding power_Suc fps_mult_nth
-    using h `a$0 = 1`  fps_power_zeroth_eq_one[OF `a$0=1`] by (simp add: ring_simps)
+    using h `a$0 = 1`  fps_power_zeroth_eq_one[OF `a$0=1`] by (simp add: field_simps)
 qed
 
 lemma startsby_one_power:"a $ 0 = (1::'a::comm_ring_1) \<Longrightarrow> a^n $ 0 = 1"
@@ -1005,7 +1005,7 @@
   case 0 thus ?case by (simp add: power_0)
 next
   case (Suc n)
-  have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)" by (simp add: ring_simps power_Suc)
+  have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)" by (simp add: field_simps power_Suc)
   also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}" by (simp add: fps_mult_nth)
   also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
     apply (rule setsum_mono_zero_right)
@@ -1045,8 +1045,8 @@
 qed
 
 lemma fps_deriv_power: "fps_deriv (a ^ n) = fps_const (of_nat n :: 'a:: comm_ring_1) * fps_deriv a * a ^ (n - 1)"
-  apply (induct n, auto simp add: power_Suc ring_simps fps_const_add[symmetric] simp del: fps_const_add)
-  by (case_tac n, auto simp add: power_Suc ring_simps)
+  apply (induct n, auto simp add: power_Suc field_simps fps_const_add[symmetric] simp del: fps_const_add)
+  by (case_tac n, auto simp add: power_Suc field_simps)
 
 lemma fps_inverse_deriv:
   fixes a:: "('a :: field) fps"
@@ -1060,11 +1060,11 @@
   with inverse_mult_eq_1[OF a0]
   have "inverse a ^ 2 * fps_deriv a + fps_deriv (inverse a) = 0"
     unfolding power2_eq_square
-    apply (simp add: ring_simps)
+    apply (simp add: field_simps)
     by (simp add: mult_assoc[symmetric])
   hence "inverse a ^ 2 * fps_deriv a + fps_deriv (inverse a) - fps_deriv a * inverse a ^ 2 = 0 - fps_deriv a * inverse a ^ 2"
     by simp
-  then show "fps_deriv (inverse a) = - fps_deriv a * inverse a ^ 2" by (simp add: ring_simps)
+  then show "fps_deriv (inverse a) = - fps_deriv a * inverse a ^ 2" by (simp add: field_simps)
 qed
 
 lemma fps_inverse_mult:
@@ -1084,7 +1084,7 @@
     from inverse_mult_eq_1[OF ab0]
     have "inverse (a*b) * (a*b) * inverse a * inverse b = 1 * inverse a * inverse b" by simp
     then have "inverse (a*b) * (inverse a * a) * (inverse b * b) = inverse a * inverse b"
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     then have ?thesis using inverse_mult_eq_1[OF a0] inverse_mult_eq_1[OF b0] by simp}
 ultimately show ?thesis by blast
 qed
@@ -1105,7 +1105,7 @@
   assumes a0: "b$0 \<noteq> 0"
   shows "fps_deriv (a / b) = (fps_deriv a * b - a * fps_deriv b) / b ^ 2"
   using fps_inverse_deriv[OF a0]
-  by (simp add: fps_divide_def ring_simps power2_eq_square fps_inverse_mult inverse_mult_eq_1'[OF a0])
+  by (simp add: fps_divide_def field_simps power2_eq_square fps_inverse_mult inverse_mult_eq_1'[OF a0])
 
 
 lemma fps_inverse_gp': "inverse (Abs_fps(\<lambda>n. (1::'a::field)))
@@ -1121,7 +1121,7 @@
 proof-
   have eq: "(1 + X) * ?r = 1"
     unfolding minus_one_power_iff
-    by (auto simp add: ring_simps fps_eq_iff)
+    by (auto simp add: field_simps fps_eq_iff)
   show ?thesis by (auto simp add: eq intro: fps_inverse_unique)
 qed
 
@@ -1185,7 +1185,7 @@
     next
       case (Suc k)
       note th = Suc.hyps[symmetric]
-      have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n = (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n" by (simp add: ring_simps)
+      have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n = (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n" by (simp add: field_simps)
       also  have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n"
         using th
         unfolding fps_sub_nth by simp
@@ -1209,10 +1209,10 @@
 definition "XD = op * X o fps_deriv"
 
 lemma XD_add[simp]:"XD (a + b) = XD a + XD (b :: ('a::comm_ring_1) fps)"
-  by (simp add: XD_def ring_simps)
+  by (simp add: XD_def field_simps)
 
 lemma XD_mult_const[simp]:"XD (fps_const (c::'a::comm_ring_1) * a) = fps_const c * XD a"
-  by (simp add: XD_def ring_simps)
+  by (simp add: XD_def field_simps)
 
 lemma XD_linear[simp]: "XD (fps_const c * a + fps_const d * b) = fps_const c * XD a + fps_const d * XD (b :: ('a::comm_ring_1) fps)"
   by simp
@@ -1226,7 +1226,7 @@
 
 lemma fps_mult_XD_shift:
   "(XD ^^ k) (a:: ('a::{comm_ring_1}) fps) = Abs_fps (\<lambda>n. (of_nat n ^ k) * a$n)"
-  by (induct k arbitrary: a) (simp_all add: power_Suc XD_def fps_eq_iff ring_simps del: One_nat_def)
+  by (induct k arbitrary: a) (simp_all add: power_Suc XD_def fps_eq_iff field_simps del: One_nat_def)
 
 subsubsection{* Rule 3 is trivial and is given by @{text fps_times_def}*}
 subsubsection{* Rule 5 --- summation and "division" by (1 - X)*}
@@ -1688,7 +1688,7 @@
         then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
           by (simp add: natpermute_max_card[OF nz, simplified])
         also have "\<dots> = a$n - setsum ?f ?Pnknn"
-          unfolding n1 using r00 a0 by (simp add: field_eq_simps fps_radical_def del: of_nat_Suc)
+          unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc)
         finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
         have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
           unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
@@ -1764,7 +1764,7 @@
   shows "a = b / c"
 proof-
   from eq have "a * c * inverse c = b * inverse c" by simp
-  hence "a * (inverse c * c) = b/c" by (simp only: field_eq_simps divide_inverse)
+  hence "a * (inverse c * c) = b/c" by (simp only: field_simps divide_inverse)
   then show "a = b/c" unfolding  field_inverse[OF c0] by simp
 qed
 
@@ -1837,7 +1837,7 @@
           show "a$(xs !i) = ?r$(xs!i)" .
         qed
         have th00: "\<And>(x::'a). of_nat (Suc k) * (x * inverse (of_nat (Suc k))) = x"
-          by (simp add: field_eq_simps del: of_nat_Suc)
+          by (simp add: field_simps del: of_nat_Suc)
         from H have "b$n = a^Suc k $ n" by (simp add: fps_eq_iff)
         also have "a ^ Suc k$n = setsum ?g ?Pnkn + setsum ?g ?Pnknn"
           unfolding fps_power_nth_Suc
@@ -1854,7 +1854,7 @@
         then have "a$n = ?r $n"
           apply (simp del: of_nat_Suc)
           unfolding fps_radical_def n1
-          by (simp add: field_eq_simps n1 th00 del: of_nat_Suc)}
+          by (simp add: field_simps n1 th00 del: of_nat_Suc)}
         ultimately show "a$n = ?r $ n" by (cases n, auto)
       qed}
     then have "a = ?r" by (simp add: fps_eq_iff)}
@@ -2018,11 +2018,11 @@
 proof-
   {fix n
     have "(fps_deriv (a oo b))$n = setsum (\<lambda>i. a $ i * (fps_deriv (b^i))$n) {0.. Suc n}"
-      by (simp add: fps_compose_def ring_simps setsum_right_distrib del: of_nat_Suc)
+      by (simp add: fps_compose_def field_simps setsum_right_distrib del: of_nat_Suc)
     also have "\<dots> = setsum (\<lambda>i. a$i * ((fps_const (of_nat i)) * (fps_deriv b * (b^(i - 1))))$n) {0.. Suc n}"
-      by (simp add: ring_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc)
+      by (simp add: field_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc)
   also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (((b^(i - 1)) * fps_deriv b))$n) {0.. Suc n}"
-    unfolding fps_mult_left_const_nth  by (simp add: ring_simps)
+    unfolding fps_mult_left_const_nth  by (simp add: field_simps)
   also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}"
     unfolding fps_mult_nth ..
   also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}"
@@ -2170,7 +2170,7 @@
   by (auto simp add: fps_eq_iff fps_compose_nth power_0_left setsum_0')
 
 lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)"
-  by (simp add: fps_eq_iff fps_compose_nth ring_simps setsum_addf)
+  by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_addf)
 
 lemma fps_compose_setsum_distrib: "(setsum f S) oo a = setsum (\<lambda>i. f i oo a) S"
 proof-
@@ -2212,7 +2212,7 @@
     apply (simp add: fps_mult_nth setsum_right_distrib)
     apply (subst setsum_commute)
     apply (rule setsum_cong2)
-    by (auto simp add: ring_simps)
+    by (auto simp add: field_simps)
   also have "\<dots> = ?l"
     apply (simp add: fps_mult_nth fps_compose_nth setsum_product)
     apply (rule setsum_cong2)
@@ -2312,7 +2312,7 @@
 qed
 
 lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)"
-  by (simp add: fps_eq_iff fps_compose_nth ring_simps setsum_negf[symmetric])
+  by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_negf[symmetric])
 
 lemma fps_compose_sub_distrib:
   shows "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
@@ -2469,7 +2469,7 @@
 proof-
   let ?r = "fps_inv"
   have ra0: "?r a $ 0 = 0" by (simp add: fps_inv_def)
-  from a1 have ra1: "?r a $ 1 \<noteq> 0" by (simp add: fps_inv_def field_eq_simps)
+  from a1 have ra1: "?r a $ 1 \<noteq> 0" by (simp add: fps_inv_def field_simps)
   have X0: "X$0 = 0" by simp
   from fps_inv[OF ra0 ra1] have "?r (?r a) oo ?r a = X" .
   then have "?r (?r a) oo ?r a oo a = X oo a" by simp
@@ -2486,7 +2486,7 @@
 proof-
   let ?r = "fps_ginv"
   from c0 have rca0: "?r c a $0 = 0" by (simp add: fps_ginv_def)
-  from a1 c1 have rca1: "?r c a $ 1 \<noteq> 0" by (simp add: fps_ginv_def field_eq_simps)
+  from a1 c1 have rca1: "?r c a $ 1 \<noteq> 0" by (simp add: fps_ginv_def field_simps)
   from fps_ginv[OF rca0 rca1] 
   have "?r b (?r c a) oo ?r c a = b" .
   then have "?r b (?r c a) oo ?r c a oo a = b oo a" by simp
@@ -2534,8 +2534,8 @@
 proof-
   {fix n
     have "?l$n = ?r $ n"
-  apply (auto simp add: E_def field_eq_simps power_Suc[symmetric]simp del: fact_Suc of_nat_Suc power_Suc)
-  by (simp add: of_nat_mult ring_simps)}
+  apply (auto simp add: E_def field_simps power_Suc[symmetric]simp del: fact_Suc of_nat_Suc power_Suc)
+  by (simp add: of_nat_mult field_simps)}
 then show ?thesis by (simp add: fps_eq_iff)
 qed
 
@@ -2545,15 +2545,15 @@
 proof-
   {assume d: ?lhs
   from d have th: "\<And>n. a $ Suc n = c * a$n / of_nat (Suc n)"
-    by (simp add: fps_deriv_def fps_eq_iff field_eq_simps del: of_nat_Suc)
+    by (simp add: fps_deriv_def fps_eq_iff field_simps del: of_nat_Suc)
   {fix n have "a$n = a$0 * c ^ n/ (of_nat (fact n))"
       apply (induct n)
       apply simp
       unfolding th
       using fact_gt_zero_nat
-      apply (simp add: field_eq_simps del: of_nat_Suc fact_Suc)
+      apply (simp add: field_simps del: of_nat_Suc fact_Suc)
       apply (drule sym)
-      by (simp add: ring_simps of_nat_mult power_Suc)}
+      by (simp add: field_simps of_nat_mult power_Suc)}
   note th' = this
   have ?rhs
     by (auto simp add: fps_eq_iff fps_const_mult_left E_def intro : th')}
@@ -2570,7 +2570,7 @@
 lemma E_add_mult: "E (a + b) = E (a::'a::field_char_0) * E b" (is "?l = ?r")
 proof-
   have "fps_deriv (?r) = fps_const (a+b) * ?r"
-    by (simp add: fps_const_add[symmetric] ring_simps del: fps_const_add)
+    by (simp add: fps_const_add[symmetric] field_simps del: fps_const_add)
   then have "?r = ?l" apply (simp only: E_unique_ODE)
     by (simp add: fps_mult_nth E_def)
   then show ?thesis ..
@@ -2618,13 +2618,13 @@
   (is "inverse ?l = ?r")
 proof-
   have th: "?l * ?r = 1"
-    by (auto simp add: ring_simps fps_eq_iff minus_one_power_iff)
+    by (auto simp add: field_simps fps_eq_iff minus_one_power_iff)
   have th': "?l $ 0 \<noteq> 0" by (simp add: )
   from fps_inverse_unique[OF th' th] show ?thesis .
 qed
 
 lemma E_power_mult: "(E (c::'a::field_char_0))^n = E (of_nat n * c)"
-  by (induct n, auto simp add: ring_simps E_add_mult power_Suc)
+  by (induct n, auto simp add: field_simps E_add_mult power_Suc)
 
 lemma radical_E:
   assumes r: "r (Suc k) 1 = 1" 
@@ -2649,18 +2649,18 @@
 text{* The generalized binomial theorem as a  consequence of @{thm E_add_mult} *}
 
 lemma gbinomial_theorem: 
-  "((a::'a::{field_char_0, division_by_zero})+b) ^ n = (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
+  "((a::'a::{field_char_0, division_ring_inverse_zero})+b) ^ n = (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
 proof-
   from E_add_mult[of a b] 
   have "(E (a + b)) $ n = (E a * E b)$n" by simp
   then have "(a + b) ^ n = (\<Sum>i\<Colon>nat = 0\<Colon>nat..n. a ^ i * b ^ (n - i)  * (of_nat (fact n) / of_nat (fact i * fact (n - i))))"
-    by (simp add: field_eq_simps fps_mult_nth of_nat_mult[symmetric] setsum_right_distrib)
+    by (simp add: field_simps fps_mult_nth of_nat_mult[symmetric] setsum_right_distrib)
   then show ?thesis 
     apply simp
     apply (rule setsum_cong2)
     apply simp
     apply (frule binomial_fact[where ?'a = 'a, symmetric])
-    by (simp add: field_eq_simps of_nat_mult)
+    by (simp add: field_simps of_nat_mult)
 qed
 
 text{* And the nat-form -- also available from Binomial.thy *}
@@ -2683,7 +2683,7 @@
   by (simp add: L_def fps_eq_iff del: of_nat_Suc)
 
 lemma L_nth: "L c $ n = (if n=0 then 0 else 1/c * ((- 1) ^ (n - 1) / of_nat n))"
-  by (simp add: L_def field_eq_simps)
+  by (simp add: L_def field_simps)
 
 lemma L_0[simp]: "L c $ 0 = 0" by (simp add: L_def)
 lemma L_E_inv:
@@ -2694,9 +2694,9 @@
   have b0: "?b $ 0 = 0" by simp
   have b1: "?b $ 1 \<noteq> 0" by (simp add: a)
   have "fps_deriv (E a - 1) oo fps_inv (E a - 1) = (fps_const a * (E a - 1) + fps_const a) oo fps_inv (E a - 1)"
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
   also have "\<dots> = fps_const a * (X + 1)" apply (simp add: fps_compose_add_distrib fps_const_mult_apply_left[symmetric] fps_inv_right[OF b0 b1])
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
   finally have eq: "fps_deriv (E a - 1) oo fps_inv (E a - 1) = fps_const a * (X + 1)" .
   from fps_inv_deriv[OF b0 b1, unfolded eq]
   have "fps_deriv (fps_inv ?b) = fps_const (inverse a) / (X + 1)"
@@ -2713,7 +2713,7 @@
   shows "L c + L d = fps_const (c+d) * L (c*d)"
   (is "?r = ?l")
 proof-
-  from c0 d0 have eq: "1/c + 1/d = (c+d)/(c*d)" by (simp add: field_eq_simps)
+  from c0 d0 have eq: "1/c + 1/d = (c+d)/(c*d)" by (simp add: field_simps)
   have "fps_deriv ?r = fps_const (1/c + 1/d) * inverse (1 + X)"
     by (simp add: fps_deriv_L fps_const_add[symmetric] algebra_simps del: fps_const_add)
   also have "\<dots> = fps_deriv ?l"
@@ -2743,7 +2743,7 @@
   have "?l = ?r \<longleftrightarrow> inverse ?x1 * ?l = inverse ?x1 * ?r" by simp
   also have "\<dots> \<longleftrightarrow> ?da = (fps_const c * a) / ?x1"
     apply (simp only: fps_divide_def  mult_assoc[symmetric] inverse_mult_eq_1[OF x10])
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
   finally have eq: "?l = ?r \<longleftrightarrow> ?lhs" by simp
   moreover
   {assume h: "?l = ?r" 
@@ -2752,8 +2752,8 @@
       
       from lrn 
       have "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n" 
-        apply (simp add: ring_simps del: of_nat_Suc)
-        by (cases n, simp_all add: field_eq_simps del: of_nat_Suc)
+        apply (simp add: field_simps del: of_nat_Suc)
+        by (cases n, simp_all add: field_simps del: of_nat_Suc)
     }
     note th0 = this
     {fix n have "a$n = (c gchoose n) * a$0"
@@ -2762,24 +2762,24 @@
       next
         case (Suc m)
         thus ?case unfolding th0
-          apply (simp add: field_eq_simps del: of_nat_Suc)
+          apply (simp add: field_simps del: of_nat_Suc)
           unfolding mult_assoc[symmetric] gbinomial_mult_1
-          by (simp add: ring_simps)
+          by (simp add: field_simps)
       qed}
     note th1 = this
     have ?rhs
       apply (simp add: fps_eq_iff)
       apply (subst th1)
-      by (simp add: ring_simps)}
+      by (simp add: field_simps)}
   moreover
   {assume h: ?rhs
   have th00:"\<And>x y. x * (a$0 * y) = a$0 * (x*y)" by (simp add: mult_commute)
     have "?l = ?r" 
       apply (subst h)
       apply (subst (2) h)
-      apply (clarsimp simp add: fps_eq_iff ring_simps)
+      apply (clarsimp simp add: fps_eq_iff field_simps)
       unfolding mult_assoc[symmetric] th00 gbinomial_mult_1
-      by (simp add: ring_simps gbinomial_mult_1)}
+      by (simp add: field_simps gbinomial_mult_1)}
   ultimately show ?thesis by blast
 qed
 
@@ -2798,9 +2798,9 @@
   have "fps_deriv ?P = ?db c * ?b d + ?b c * ?db d - ?db (c + d)"  by simp
   also have "\<dots> = inverse (1 + X) * (fps_const c * ?b c * ?b d + fps_const d * ?b c * ?b d - fps_const (c+d) * ?b (c + d))"
     unfolding fps_binomial_deriv
-    by (simp add: fps_divide_def ring_simps)
+    by (simp add: fps_divide_def field_simps)
   also have "\<dots> = (fps_const (c + d)/ (1 + X)) * ?P"
-    by (simp add: ring_simps fps_divide_def fps_const_add[symmetric] del: fps_const_add)
+    by (simp add: field_simps fps_divide_def fps_const_add[symmetric] del: fps_const_add)
   finally have th0: "fps_deriv ?P = fps_const (c+d) * ?P / (1 + X)"
     by (simp add: fps_divide_def)
   have "?P = fps_const (?P$0) * ?b (c + d)"
@@ -2880,7 +2880,7 @@
           using kn pochhammer_minus'[where k=k and n=n and b=b]
           apply (simp add:  pochhammer_same)
           using bn0
-          by (simp add: field_eq_simps power_add[symmetric])}
+          by (simp add: field_simps power_add[symmetric])}
       moreover
       {assume nk: "k \<noteq> n"
         have m1nk: "?m1 n = setprod (%i. - 1) {0..m}" 
@@ -2905,7 +2905,7 @@
           unfolding m1nk 
           
           unfolding m h pochhammer_Suc_setprod
-          apply (simp add: field_eq_simps del: fact_Suc id_def)
+          apply (simp add: field_simps del: fact_Suc id_def)
           unfolding fact_altdef_nat id_def
           unfolding of_nat_setprod
           unfolding setprod_timesf[symmetric]
@@ -2942,10 +2942,10 @@
           apply auto
           done
         then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k = setprod (%i. b - of_nat i) {0.. n - k - 1}" 
-          using nz' by (simp add: field_eq_simps)
+          using nz' by (simp add: field_simps)
         have "(?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k) = ((?m1 k * ?p (of_nat n) k) / ?f n) * ((?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k)"
           using bnz0
-          by (simp add: field_eq_simps)
+          by (simp add: field_simps)
         also have "\<dots> = b gchoose (n - k)" 
           unfolding th1 th2
           using kn' by (simp add: gbinomial_def)
@@ -2959,15 +2959,15 @@
   note th00 = this
   have "?r = ((a + b) gchoose n) * (of_nat (fact n)/ (?m1 n * pochhammer (- b) n))"
     unfolding gbinomial_pochhammer 
-    using bn0 by (auto simp add: field_eq_simps)
+    using bn0 by (auto simp add: field_simps)
   also have "\<dots> = ?l"
     unfolding gbinomial_Vandermonde[symmetric]
     apply (simp add: th00)
     unfolding gbinomial_pochhammer
-    using bn0 apply (simp add: setsum_left_distrib setsum_right_distrib field_eq_simps)
+    using bn0 apply (simp add: setsum_left_distrib setsum_right_distrib field_simps)
     apply (rule setsum_cong2)
     apply (drule th00(2))
-    by (simp add: field_eq_simps power_add[symmetric])
+    by (simp add: field_simps power_add[symmetric])
   finally show ?thesis by simp
 qed 
 
@@ -2992,7 +2992,7 @@
   have nz: "pochhammer c n \<noteq> 0" using c
     by (simp add: pochhammer_eq_0_iff)
   from Vandermonde_pochhammer_lemma[where a = "?a" and b="?b" and n=n, OF h, unfolded th0 th1]
-  show ?thesis using nz by (simp add: field_eq_simps setsum_right_distrib)
+  show ?thesis using nz by (simp add: field_simps setsum_right_distrib)
 qed
 
 subsubsection{* Formal trigonometric functions  *}
@@ -3014,11 +3014,11 @@
         using en by (simp add: fps_sin_def)
       also have "\<dots> = (- 1)^(n div 2) * c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
         unfolding fact_Suc of_nat_mult
-        by (simp add: field_eq_simps del: of_nat_add of_nat_Suc)
+        by (simp add: field_simps del: of_nat_add of_nat_Suc)
       also have "\<dots> = (- 1)^(n div 2) *c^Suc n / of_nat (fact n)"
-        by (simp add: field_eq_simps del: of_nat_add of_nat_Suc)
+        by (simp add: field_simps del: of_nat_add of_nat_Suc)
       finally have "?lhs $n = ?rhs$n" using en
-        by (simp add: fps_cos_def ring_simps power_Suc )}
+        by (simp add: fps_cos_def field_simps power_Suc )}
     then show "?lhs $ n = ?rhs $ n"
       by (cases "even n", simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
 qed
@@ -3038,13 +3038,13 @@
         using en by (simp add: fps_cos_def)
       also have "\<dots> = (- 1)^((n + 1) div 2)*c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
         unfolding fact_Suc of_nat_mult
-        by (simp add: field_eq_simps del: of_nat_add of_nat_Suc)
+        by (simp add: field_simps del: of_nat_add of_nat_Suc)
       also have "\<dots> = (- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact n)"
-        by (simp add: field_eq_simps del: of_nat_add of_nat_Suc)
+        by (simp add: field_simps del: of_nat_add of_nat_Suc)
       also have "\<dots> = (- ((- 1)^((n - 1) div 2))) * c^Suc n / of_nat (fact n)"
         unfolding th0 unfolding th1[OF en] by simp
       finally have "?lhs $n = ?rhs$n" using en
-        by (simp add: fps_sin_def ring_simps power_Suc)}
+        by (simp add: fps_sin_def field_simps power_Suc)}
     then show "?lhs $ n = ?rhs $ n"
       by (cases "even n", simp_all add: fps_deriv_def fps_sin_def
         fps_cos_def)
@@ -3055,7 +3055,7 @@
 proof-
   have "fps_deriv ?lhs = 0"
     apply (simp add:  fps_deriv_power fps_sin_deriv fps_cos_deriv power_Suc)
-    by (simp add: ring_simps fps_const_neg[symmetric] del: fps_const_neg)
+    by (simp add: field_simps fps_const_neg[symmetric] del: fps_const_neg)
   then have "?lhs = fps_const (?lhs $ 0)"
     unfolding fps_deriv_eq_0_iff .
   also have "\<dots> = 1"
@@ -3177,7 +3177,7 @@
   have th0: "fps_cos c $ 0 \<noteq> 0" by (simp add: fps_cos_def)
   show ?thesis
     using fps_sin_cos_sum_of_squares[of c]
-    apply (simp add: fps_tan_def fps_divide_deriv[OF th0] fps_sin_deriv fps_cos_deriv add: fps_const_neg[symmetric] ring_simps power2_eq_square del: fps_const_neg)
+    apply (simp add: fps_tan_def fps_divide_deriv[OF th0] fps_sin_deriv fps_cos_deriv add: fps_const_neg[symmetric] field_simps power2_eq_square del: fps_const_neg)
     unfolding right_distrib[symmetric]
     by simp
 qed
@@ -3252,7 +3252,7 @@
 subsection {* Hypergeometric series *}
 
 
-definition "F as bs (c::'a::{field_char_0, division_by_zero}) = Abs_fps (%n. (foldl (%r a. r* pochhammer a n) 1 as * c^n)/ (foldl (%r b. r * pochhammer b n) 1 bs * of_nat (fact n)))"
+definition "F as bs (c::'a::{field_char_0, division_ring_inverse_zero}) = Abs_fps (%n. (foldl (%r a. r* pochhammer a n) 1 as * c^n)/ (foldl (%r b. r * pochhammer b n) 1 bs * of_nat (fact n)))"
 
 lemma F_nth[simp]: "F as bs c $ n =  (foldl (%r a. r* pochhammer a n) 1 as * c^n)/ (foldl (%r b. r * pochhammer b n) 1 bs * of_nat (fact n))"
   by (simp add: F_def)
@@ -3321,9 +3321,9 @@
   by (simp add: fps_eq_iff fps_integral_def)
 
 lemma F_minus_nat: 
-  "F [- of_nat n] [- of_nat (n + m)] (c::'a::{field_char_0, division_by_zero}) $ k = (if k <= n then pochhammer (- of_nat n) k * c ^ k /
+  "F [- of_nat n] [- of_nat (n + m)] (c::'a::{field_char_0, division_ring_inverse_zero}) $ k = (if k <= n then pochhammer (- of_nat n) k * c ^ k /
     (pochhammer (- of_nat (n + m)) k * of_nat (fact k)) else 0)"
-  "F [- of_nat m] [- of_nat (m + n)] (c::'a::{field_char_0, division_by_zero}) $ k = (if k <= m then pochhammer (- of_nat m) k * c ^ k /
+  "F [- of_nat m] [- of_nat (m + n)] (c::'a::{field_char_0, division_ring_inverse_zero}) $ k = (if k <= m then pochhammer (- of_nat m) k * c ^ k /
     (pochhammer (- of_nat (m + n)) k * of_nat (fact k)) else 0)"
   by (auto simp add: pochhammer_eq_0_iff)
 
--- a/src/HOL/Library/Numeral_Type.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Library/Numeral_Type.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -213,7 +213,7 @@
 
 lemma comm_ring_1: "OFCLASS('a, comm_ring_1_class)"
 apply (intro_classes, unfold definitions)
-apply (simp_all add: Rep_simps zmod_simps ring_simps)
+apply (simp_all add: Rep_simps zmod_simps field_simps)
 done
 
 end
--- a/src/HOL/Library/Polynomial.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Library/Polynomial.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -1093,10 +1093,10 @@
 apply (cases "r = 0")
 apply (cases "r' = 0")
 apply (simp add: pdivmod_rel_def)
-apply (simp add: pdivmod_rel_def ring_simps degree_mult_eq)
+apply (simp add: pdivmod_rel_def field_simps degree_mult_eq)
 apply (cases "r' = 0")
 apply (simp add: pdivmod_rel_def degree_mult_eq)
-apply (simp add: pdivmod_rel_def ring_simps)
+apply (simp add: pdivmod_rel_def field_simps)
 apply (simp add: degree_mult_eq degree_add_less)
 done
 
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Mon Apr 26 11:34:19 2010 +0200
@@ -1282,9 +1282,9 @@
   fun simple_cterm_ord t u = Term_Ord.fast_term_ord (term_of t, term_of u) = LESS
  val concl = Thm.dest_arg o cprop_of
  val shuffle1 =
-   fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps) })
+   fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: field_simps) })
  val shuffle2 =
-    fconv_rule (rewr_conv @{lemma "(x + a == y) ==  (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps)})
+    fconv_rule (rewr_conv @{lemma "(x + a == y) ==  (x == y - (a::real))" by (atomize (full)) (simp add: field_simps)})
  fun substitutable_monomial fvs tm = case term_of tm of
     Free(_,@{typ real}) => if not (member (op aconvc) fvs tm) then (Rat.one,tm)
                            else raise Failure "substitutable_monomial"
--- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -1383,7 +1383,7 @@
     apply(rule_tac x=x in bexI) apply assumption+ apply(rule continuous_on_intros)+
     unfolding frontier_cball subset_eq Ball_def image_iff apply(rule,rule,erule bexE)
     unfolding vector_dist_norm apply(simp add: * norm_minus_commute) . note x = this
-  hence "scaleR 2 a = scaleR 1 x + scaleR 1 x" by(auto simp add:group_simps)
+  hence "scaleR 2 a = scaleR 1 x + scaleR 1 x" by(auto simp add:algebra_simps)
   hence "a = x" unfolding scaleR_left_distrib[THEN sym] by auto 
   thus False using x using assms by auto qed
 
@@ -1396,7 +1396,7 @@
  "interval_bij (a,b) (u,v) = (\<lambda>x. (\<chi> i. (v$i - u$i) / (b$i - a$i) * x$i) +
             (\<chi> i. u$i - (v$i - u$i) / (b$i - a$i) * a$i))"
   apply rule unfolding Cart_eq interval_bij_def vector_component_simps
-  by(auto simp add:group_simps field_simps add_divide_distrib[THEN sym]) 
+  by(auto simp add: field_simps add_divide_distrib[THEN sym]) 
 
 lemma continuous_interval_bij:
   "continuous (at x) (interval_bij (a,b::real^'n) (u,v))" 
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -646,7 +646,7 @@
       using assms(1)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
       using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
       apply - apply(rule add_mono) by auto
-    hence "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> u * (f x + g x) + v * (f y + g y)" by (simp add: ring_simps)  }
+    hence "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> u * (f x + g x) + v * (f y + g y)" by (simp add: field_simps)  }
   thus ?thesis unfolding convex_on_def by auto 
 qed
 
@@ -654,7 +654,7 @@
   assumes "0 \<le> (c::real)" "convex_on s f"
   shows "convex_on s (\<lambda>x. c * f x)"
 proof-
-  have *:"\<And>u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: ring_simps)
+  have *:"\<And>u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: field_simps)
   show ?thesis using assms(2) and mult_mono1[OF _ assms(1)] unfolding convex_on_def and * by auto
 qed
 
@@ -1060,7 +1060,7 @@
 proof-
   have fin:"finite {a,b,c}" "finite {b,c}" "finite {c}" by auto
   have *:"\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z"
-         "\<And>x y z ::real^_. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by (auto simp add: ring_simps)
+         "\<And>x y z ::real^_. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by (auto simp add: field_simps)
   show ?thesis unfolding convex_hull_finite[OF fin(1)] and Collect_def and convex_hull_finite_step[OF fin(2)] and *
     unfolding convex_hull_finite_step[OF fin(3)] apply(rule ext) apply simp apply auto
     apply(rule_tac x=va in exI) apply (rule_tac x="u c" in exI) apply simp
@@ -2310,7 +2310,7 @@
   } moreover
   { fix a b assume "\<not> u * a + v * b \<le> a"
     hence "v * b > (1 - u) * a" unfolding not_le using as(4) by(auto simp add: field_simps)
-    hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: ring_simps)
+    hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: field_simps)
     hence "u * a + v * b \<le> b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) }
   ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> s" apply- apply(rule assms[unfolded is_interval_def, rule_format, OF as(1,2)])
     using as(3-) dimindex_ge_1 apply- by(auto simp add: vector_component) qed
--- a/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -1,11 +1,12 @@
-(*  Title:      HOL/Library/Convex_Euclidean_Space.thy
-    Author:                     John Harrison
-    Translation from HOL light: Robert Himmelmann, TU Muenchen *)
+(*  Title:                       HOL/Multivariate_Analysis/Derivative.thy
+    Author:                      John Harrison
+    Translation from HOL Light:  Robert Himmelmann, TU Muenchen
+*)
 
 header {* Multivariate calculus in Euclidean space. *}
 
 theory Derivative
-  imports Brouwer_Fixpoint RealVector
+imports Brouwer_Fixpoint RealVector
 begin
 
 
@@ -40,7 +41,7 @@
   show ?l unfolding deriv_def LIM_def apply safe apply(drule as,safe)
     apply(rule_tac x=d in exI,safe) apply(erule_tac x="xa + x" in allE)
     unfolding vector_dist_norm diff_0_right norm_scaleR
-    unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:group_simps *) qed
+    unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:algebra_simps *) qed
 
 lemma FDERIV_conv_has_derivative:"FDERIV f (x::'a::{real_normed_vector,perfect_space}) :> f' = (f has_derivative f') (at x)" (is "?l = ?r") proof 
   assume ?l note as = this[unfolded fderiv_def]
@@ -50,14 +51,14 @@
     thus "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow>
       dist ((1 / norm (xa - netlimit (at x))) *\<^sub>R (f xa - (f (netlimit (at x)) + f' (xa - netlimit (at x))))) (0) < e"
       apply(rule_tac x=d in exI) apply(erule conjE,rule,assumption) apply rule apply(erule_tac x="xa - x" in allE)
-      unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:group_simps) qed next
+      unfolding vector_dist_norm netlimit_at[of x] by (auto simp add: diff_diff_eq) qed next
   assume ?r note as = this[unfolded has_derivative_def]
   show ?l unfolding fderiv_def LIM_def apply-apply(rule,rule as[THEN conjunct1]) proof(rule,rule)
     fix e::real assume "e>0"
     guess d using as[THEN conjunct2,unfolded Lim_at,rule_format,OF`e>0`] ..
     thus "\<exists>s>0. \<forall>xa. xa \<noteq> 0 \<and> dist xa 0 < s \<longrightarrow> dist (norm (f (x + xa) - f x - f' xa) / norm xa) 0 < e" apply-
       apply(rule_tac x=d in exI) apply(erule conjE,rule,assumption) apply rule apply(erule_tac x="xa + x" in allE)
-      unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:group_simps) qed qed
+      unfolding vector_dist_norm netlimit_at[of x] by (auto simp add: diff_diff_eq add.commute) qed qed
 
 subsection {* These are the only cases we'll care about, probably. *}
 
@@ -76,7 +77,7 @@
         (\<forall>e>0. \<exists>d>0. \<forall>x'\<in>s. 0 < norm(x' - x) \<and> norm(x' - x) < d
         \<longrightarrow> norm(f x' - f x - f'(x' - x)) / norm(x' - x) < e)"
   unfolding has_derivative_within Lim_within vector_dist_norm
-  unfolding diff_0_right norm_mul by(simp add: group_simps)
+  unfolding diff_0_right norm_mul by (simp add: diff_diff_eq)
 
 lemma has_derivative_at':
  "(f has_derivative f') (at x) \<longleftrightarrow> bounded_linear f' \<and>
@@ -186,14 +187,14 @@
   note as = assms[unfolded has_derivative_def]
   show ?thesis unfolding has_derivative_def apply(rule,rule bounded_linear_add)
     using Lim_add[OF as(1)[THEN conjunct2] as(2)[THEN conjunct2]] and as
-    by(auto simp add:group_simps scaleR_right_diff_distrib scaleR_right_distrib) qed
+    by (auto simp add:algebra_simps scaleR_right_diff_distrib scaleR_right_distrib) qed
 
 lemma has_derivative_add_const:"(f has_derivative f') net \<Longrightarrow> ((\<lambda>x. f x + c) has_derivative f') net"
   apply(drule has_derivative_add) apply(rule has_derivative_const) by auto
 
 lemma has_derivative_sub:
  "(f has_derivative f') net \<Longrightarrow> (g has_derivative g') net \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) has_derivative (\<lambda>h. f'(h) - g'(h))) net"
-  apply(drule has_derivative_add) apply(drule has_derivative_neg,assumption) by(simp add:group_simps)
+  apply(drule has_derivative_add) apply(drule has_derivative_neg,assumption) by(simp add:algebra_simps)
 
 lemma has_derivative_setsum: assumes "finite s" "\<forall>a\<in>s. ((f a) has_derivative (f' a)) net"
   shows "((\<lambda>x. setsum (\<lambda>a. f a x) s) has_derivative (\<lambda>h. setsum (\<lambda>a. f' a h) s)) net"
@@ -391,8 +392,8 @@
       case False hence "norm (f y - (f x + f' (y - x))) < e * norm (y - x)" using as(4)[rule_format, OF `y\<in>s`]
 	unfolding vector_dist_norm diff_0_right norm_mul using as(3)
 	using pos_divide_less_eq[OF False[unfolded dist_nz], unfolded vector_dist_norm]
-	by(auto simp add:linear_0 linear_sub group_simps)
-      thus ?thesis by(auto simp add:group_simps) qed qed next
+	by (auto simp add: linear_0 linear_sub)
+      thus ?thesis by(auto simp add:algebra_simps) qed qed next
   assume ?rhs thus ?lhs unfolding has_derivative_within Lim_within apply-apply(erule conjE,rule,assumption)
     apply(rule,erule_tac x="e/2" in allE,rule,erule impE) defer apply(erule exE,rule_tac x=d in exI)
     apply(erule conjE,rule,assumption,rule,rule) unfolding vector_dist_norm diff_0_right norm_scaleR
@@ -400,7 +401,7 @@
     fix e d y assume "bounded_linear f'" "0 < e" "0 < d" "y \<in> s" "0 < norm (y - x) \<and> norm (y - x) < d"
         "norm (f y - f x - f' (y - x)) \<le> e / 2 * norm (y - x)"
     thus "\<bar>1 / norm (y - x)\<bar> * norm (f y - (f x + f' (y - x))) < e"
-      apply(rule_tac le_less_trans[of _ "e/2"]) by(auto intro!:mult_imp_div_pos_le simp add:group_simps) qed auto qed
+      apply(rule_tac le_less_trans[of _ "e/2"]) by(auto intro!:mult_imp_div_pos_le simp add:algebra_simps) qed auto qed
 
 lemma has_derivative_at_alt:
   "(f has_derivative f') (at x) \<longleftrightarrow> bounded_linear f' \<and>
@@ -435,8 +436,8 @@
     hence 1:"norm (f y - f x - f' (y - x)) \<le> min (norm (y - x)) (e / 2 / B2 * norm (y - x))" using d1 d2 d by auto
 
     have "norm (f y - f x) \<le> norm (f y - f x - f' (y - x)) + norm (f' (y - x))"
-      using norm_triangle_sub[of "f y - f x" "f' (y - x)"] by(auto simp add:group_simps)
-    also have "\<dots> \<le> norm (f y - f x - f' (y - x)) + B1 * norm (y - x)" apply(rule add_left_mono) using B1 by(auto simp add:group_simps)
+      using norm_triangle_sub[of "f y - f x" "f' (y - x)"] by(auto simp add:algebra_simps)
+    also have "\<dots> \<le> norm (f y - f x - f' (y - x)) + B1 * norm (y - x)" apply(rule add_left_mono) using B1 by(auto simp add:algebra_simps)
     also have "\<dots> \<le> min (norm (y - x)) (e / 2 / B2 * norm (y - x)) + B1 * norm (y - x)" apply(rule add_right_mono) using d1 d2 d as by auto
     also have "\<dots> \<le> norm (y - x) + B1 * norm (y - x)" by auto
     also have "\<dots> = norm (y - x) * (1 + B1)" by(auto simp add:field_simps)
@@ -453,8 +454,8 @@
     interpret g': bounded_linear g' using assms(2) by auto
     interpret f': bounded_linear f' using assms(1) by auto
     have "norm (- g' (f' (y - x)) + g' (f y - f x)) = norm (g' (f y - f x - f' (y - x)))"
-      by(auto simp add:group_simps f'.diff g'.diff g'.add)
-    also have "\<dots> \<le> B2 * norm (f y - f x - f' (y - x))" using B2 by(auto simp add:group_simps)
+      by(auto simp add:algebra_simps f'.diff g'.diff g'.add)
+    also have "\<dots> \<le> B2 * norm (f y - f x - f' (y - x))" using B2 by(auto simp add:algebra_simps)
     also have "\<dots> \<le> B2 * (e / 2 / B2 * norm (y - x))" apply(rule mult_left_mono) using as d1 d2 d B2 by auto 
     also have "\<dots> \<le> e / 2 * norm (y - x)" using B2 by auto
     finally have 5:"norm (- g' (f' (y - x)) + g' (f y - f x)) \<le> e / 2 * norm (y - x)" by auto
@@ -535,7 +536,7 @@
       unfolding scaleR_right_distrib by auto
     also have "\<dots> = norm ((1 / abs c) *\<^sub>R (c *\<^sub>R (- (f' (basis i)) + f'' (basis i))))"  
       unfolding f'.scaleR f''.scaleR unfolding scaleR_right_distrib scaleR_minus_right by auto
-    also have "\<dots> = e" unfolding e_def norm_mul using c[THEN conjunct1] using norm_minus_cancel[of "f' (basis i) - f'' (basis i)"] by(auto simp add:group_simps)
+    also have "\<dots> = e" unfolding e_def norm_mul using c[THEN conjunct1] using norm_minus_cancel[of "f' (basis i) - f'' (basis i)"] by (auto simp add: add.commute ab_diff_minus)
     finally show False using c using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R basis i"] using norm_basis[of i] unfolding vector_dist_norm 
       unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib by auto qed qed
 
@@ -623,7 +624,7 @@
   have ***:"\<And>y y1 y2 d dx::real. (y1\<le>y\<and>y2\<le>y) \<or> (y\<le>y1\<and>y\<le>y2) \<Longrightarrow> d < abs dx \<Longrightarrow> abs(y1 - y - - dx) \<le> d \<Longrightarrow> (abs (y2 - y - dx) \<le> d) \<Longrightarrow> False" by arith
   show False apply(rule ***[OF **, where dx="d * D $ k $ j" and d="\<bar>D $ k $ j\<bar> / 2 * \<bar>d\<bar>"]) 
     using *[of "-d"] and *[of d] and d[THEN conjunct1] and j unfolding mult_minus_left
-    unfolding abs_mult diff_minus_eq_add scaleR.minus_left unfolding group_simps by (auto intro: mult_pos_pos)
+    unfolding abs_mult diff_minus_eq_add scaleR.minus_left unfolding algebra_simps by (auto intro: mult_pos_pos)
 qed
 
 subsection {* In particular if we have a mapping into @{typ "real^1"}. *}
@@ -727,7 +728,7 @@
   shows "norm(f x - f y) \<le> B * norm(x - y)" proof-
   let ?p = "\<lambda>u. x + u *\<^sub>R (y - x)"
   have *:"\<And>u. u\<in>{0..1} \<Longrightarrow> x + u *\<^sub>R (y - x) \<in> s"
-    using assms(1)[unfolded convex_alt,rule_format,OF x y] unfolding scaleR_left_diff_distrib scaleR_right_diff_distrib by(auto simp add:group_simps)
+    using assms(1)[unfolded convex_alt,rule_format,OF x y] unfolding scaleR_left_diff_distrib scaleR_right_diff_distrib by(auto simp add:algebra_simps)
   hence 1:"continuous_on {0..1} (f \<circ> ?p)" apply- apply(rule continuous_on_intros continuous_on_vmul)+
     unfolding continuous_on_eq_continuous_within apply(rule,rule differentiable_imp_continuous_within)
     unfolding differentiable_def apply(rule_tac x="f' xa" in exI)
@@ -862,7 +863,7 @@
   assumes "compact t" "convex t"  "t \<noteq> {}" "continuous_on t f"
   "\<forall>x\<in>s. \<forall>y\<in>t. x + (y - f y) \<in> t" "x\<in>s"
   shows "\<exists>y\<in>t. f y = x" proof-
-  have *:"\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y" by(auto simp add:group_simps)
+  have *:"\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y" by(auto simp add:algebra_simps)
   show ?thesis  unfolding * apply(rule brouwer[OF assms(1-3), of "\<lambda>y. x + (y - f y)"])
     apply(rule continuous_on_intros assms)+ using assms(4-6) by auto qed
 
@@ -907,8 +908,8 @@
     finally have *:"norm (x + g' (z - f x) - x) < e0" by auto
     have **:"f x + f' (x + g' (z - f x) - x) = z" using assms(6)[unfolded o_def id_def,THEN cong] by auto
     have "norm (f x - (y + (z - f (x + g' (z - f x))))) \<le> norm (f (x + g' (z - f x)) - z) + norm (f x - y)"
-      using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] by(auto simp add:group_simps)
-    also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" using e0[THEN conjunct2,rule_format,OF *] unfolding group_simps ** by auto 
+      using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] by(auto simp add:algebra_simps)
+    also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" using e0[THEN conjunct2,rule_format,OF *] unfolding algebra_simps ** by auto 
     also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + e/2" using as(1)[unfolded mem_cball vector_dist_norm] by auto
     also have "\<dots> \<le> 1 / (B * 2) * B * norm (z - f x) + e/2" using * and B by(auto simp add:field_simps)
     also have "\<dots> \<le> 1 / 2 * norm (z - f x) + e/2" by auto
@@ -983,7 +984,7 @@
 (* we know for some other reason that the inverse function exists, it's OK. *}
 
 lemma bounded_linear_sub: "bounded_linear f \<Longrightarrow> bounded_linear g ==> bounded_linear (\<lambda>x. f x - g x)"
-  using bounded_linear_add[of f "\<lambda>x. - g x"] bounded_linear_minus[of g] by(auto simp add:group_simps)
+  using bounded_linear_add[of f "\<lambda>x. - g x"] bounded_linear_minus[of g] by(auto simp add:algebra_simps)
 
 lemma has_derivative_locally_injective: fixes f::"real^'n \<Rightarrow> real^'m"
   assumes "a \<in> s" "open s" "bounded_linear g'" "g' o f'(a) = id"
@@ -1004,7 +1005,7 @@
     show "\<forall>x\<in>ball a d. \<forall>x'\<in>ball a d. f x' = f x \<longrightarrow> x' = x" proof(intro strip)
       fix x y assume as:"x\<in>ball a d" "y\<in>ball a d" "f x = f y"
       def ph \<equiv> "\<lambda>w. w - g'(f w - f x)" have ph':"ph = g' \<circ> (\<lambda>w. f' a w - (f w - f x))"
-	unfolding ph_def o_def unfolding diff using f'g' by(auto simp add:group_simps)
+	unfolding ph_def o_def unfolding diff using f'g' by(auto simp add:algebra_simps)
       have "norm (ph x - ph y) \<le> (1/2) * norm (x - y)"
 	apply(rule differentiable_bound[OF convex_ball _ _ as(1-2), where f'="\<lambda>x v. v - g'(f' x v)"])
 	apply(rule_tac[!] ballI) proof- fix u assume u:"u \<in> ball a d" hence "u\<in>s" using d d2 by auto
@@ -1020,7 +1021,7 @@
 	  unfolding linear_conv_bounded_linear by(rule assms(3) **)+ 
 	also have "\<dots> \<le> onorm g' * k" apply(rule mult_left_mono) 
 	  using d1[THEN conjunct2,rule_format,of u] using onorm_neg[OF **(1)[unfolded linear_linear]]
-	  using d and u and onorm_pos_le[OF assms(3)[unfolded linear_linear]] by(auto simp add:group_simps) 
+	  using d and u and onorm_pos_le[OF assms(3)[unfolded linear_linear]] by(auto simp add:algebra_simps) 
 	also have "\<dots> \<le> 1/2" unfolding k_def by auto
 	finally show "onorm (\<lambda>v. v - g' (f' u v)) \<le> 1 / 2" by assumption qed
       moreover have "norm (ph y - ph x) = norm (y - x)" apply(rule arg_cong[where f=norm])
@@ -1039,7 +1040,7 @@
     fix x assume "x\<in>s" show "((\<lambda>a. f m a - f n a) has_derivative (\<lambda>h. f' m x h - f' n x h)) (at x within s)"
       by(rule has_derivative_intros assms(2)[rule_format] `x\<in>s`)+
     { fix h have "norm (f' m x h - f' n x h) \<le> norm (f' m x h - g' x h) + norm (f' n x h - g' x h)"
-	using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] unfolding norm_minus_commute by(auto simp add:group_simps) 
+	using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] unfolding norm_minus_commute by(auto simp add:algebra_simps) 
       also have "\<dots> \<le> e * norm h+ e * norm h"  using assms(3)[rule_format,OF `N\<le>m` `x\<in>s`, of h] assms(3)[rule_format,OF `N\<le>n` `x\<in>s`, of h]
 	by(auto simp add:field_simps)
       finally have "norm (f' m x h - f' n x h) \<le> 2 * e * norm h" by auto }
@@ -1083,7 +1084,7 @@
       have "eventually (\<lambda>xa. norm (f n x - f n y - (f xa x - f xa y)) \<le> e * norm (x - y)) sequentially" 
 	unfolding eventually_sequentially apply(rule_tac x=N in exI) proof(rule,rule)
 	fix m assume "N\<le>m" thus "norm (f n x - f n y - (f m x - f m y)) \<le> e * norm (x - y)"
-	  using N[rule_format, of n m x y] and as by(auto simp add:group_simps) qed
+	  using N[rule_format, of n m x y] and as by(auto simp add:algebra_simps) qed
       thus "norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)" apply-
 	apply(rule Lim_norm_ubound[OF trivial_limit_sequentially, where f="\<lambda>m. (f n x - f n y) - (f m x - f m y)"])
 	apply(rule Lim_sub Lim_const g[rule_format] as)+ by assumption qed qed
@@ -1120,10 +1121,10 @@
 	have "norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)" using d1 and as by auto ultimately
 	have "norm (g y - g x - f' ?N x (y - x)) \<le> 2 * e / 3 * norm (y - x)" 
 	  using norm_triangle_le[of "g y - g x - (f ?N y - f ?N x)" "f ?N y - f ?N x - f' ?N x (y - x)" "2 * e / 3 * norm (y - x)"] 
-	  by (auto simp add:group_simps) moreover
+	  by (auto simp add:algebra_simps) moreover
 	have " norm (f' ?N x (y - x) - g' x (y - x)) \<le> e / 3 * norm (y - x)" using N1 `x\<in>s` by auto
 	ultimately show "norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)"
-	  using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] by(auto simp add:group_simps)
+	  using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] by(auto simp add:algebra_simps)
 	qed qed qed qed
 
 subsection {* Can choose to line up antiderivatives if we want. *}
@@ -1274,7 +1275,7 @@
   unfolding has_vector_derivative_def using has_derivative_id by auto
 
 lemma has_vector_derivative_cmul:  "(f has_vector_derivative f') net \<Longrightarrow> ((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net"
-  unfolding has_vector_derivative_def apply(drule has_derivative_cmul) by(auto simp add:group_simps)
+  unfolding has_vector_derivative_def apply(drule has_derivative_cmul) by(auto simp add:algebra_simps)
 
 lemma has_vector_derivative_cmul_eq: assumes "c \<noteq> 0"
   shows "(((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net \<longleftrightarrow> (f has_vector_derivative f') net)"
--- a/src/HOL/Multivariate_Analysis/Determinants.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Determinants.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -55,7 +55,7 @@
 done
 
   (* FIXME: In Finite_Set there is a useless further assumption *)
-lemma setprod_inversef: "finite A ==> setprod (inverse \<circ> f) A = (inverse (setprod f A) :: 'a:: {division_by_zero, field})"
+lemma setprod_inversef: "finite A ==> setprod (inverse \<circ> f) A = (inverse (setprod f A) :: 'a:: {division_ring_inverse_zero, field})"
   apply (erule finite_induct)
   apply (simp)
   apply simp
@@ -352,13 +352,13 @@
     apply (rule setprod_insert)
     apply simp
     by blast
-  also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" by (simp add: ring_simps)
+  also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" by (simp add: field_simps)
   also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?h i $ p i) ?Uk)" by (metis th1 th2)
   also have "\<dots> = setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + setprod (\<lambda>i. ?h i $ p i) (insert k ?Uk)"
     unfolding  setprod_insert[OF th3] by simp
   finally have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?g i $ p i) ?U + setprod (\<lambda>i. ?h i $ p i) ?U" unfolding kU[symmetric] .
   then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U + of_int (sign p) * setprod (\<lambda>i. ?h i $ p i) ?U"
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
 qed
 
 lemma det_row_mul:
@@ -389,14 +389,14 @@
     apply (rule setprod_insert)
     apply simp
     by blast
-  also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" by (simp add: ring_simps)
+  also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" by (simp add: field_simps)
   also have "\<dots> = c* (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk)"
     unfolding th1 by (simp add: mult_ac)
   also have "\<dots> = c* (setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk))"
     unfolding  setprod_insert[OF th3] by simp
   finally have "setprod (\<lambda>i. ?f i $ p i) ?U = c* (setprod (\<lambda>i. ?g i $ p i) ?U)" unfolding kU[symmetric] .
   then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = c * (of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U)"
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
 qed
 
 lemma det_row_0:
@@ -604,7 +604,7 @@
   have "setprod (\<lambda>i. c i * a i $ p i) ?U = setprod c ?U * setprod (\<lambda>i. a i $ p i) ?U"
     unfolding setprod_timesf ..
   then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) =
-        setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: ring_simps)
+        setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: field_simps)
 qed
 
 lemma det_mul:
@@ -681,7 +681,7 @@
         using permutes_in_image[OF q] by vector
       show "?s q * setprod (\<lambda>i. (((\<chi> i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = ?s p * (setprod (\<lambda>i. A$i$p i) ?U) * (?s (q o inv p) * setprod (\<lambda>i. B$i$(q o inv p) i) ?U)"
         using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
-        by (simp add: sign_nz th00 ring_simps sign_idempotent sign_compose)
+        by (simp add: sign_nz th00 field_simps sign_idempotent sign_compose)
     qed
   }
   then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B"
@@ -772,7 +772,7 @@
   have fUk: "finite ?Uk" by simp
   have kUk: "k \<notin> ?Uk" by simp
   have th00: "\<And>k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s"
-    by (vector ring_simps)
+    by (vector field_simps)
   have th001: "\<And>f k . (\<lambda>x. if x = k then f k else f x) = f" by (auto intro: ext)
   have "(\<chi> i. row i A) = A" by (vector row_def)
   then have thd1: "det (\<chi> i. row i A) = det A"  by simp
@@ -793,7 +793,7 @@
     unfolding thd0
     unfolding det_row_mul
     unfolding th001[of k "\<lambda>i. row i A"]
-    unfolding thd1  by (simp add: ring_simps)
+    unfolding thd1  by (simp add: field_simps)
 qed
 
 lemma cramer_lemma:
@@ -901,7 +901,7 @@
   have th: "\<And>x::'a. x = 1 \<or> x = - 1 \<longleftrightarrow> x*x = 1" (is "\<And>x::'a. ?ths x")
   proof-
     fix x:: 'a
-    have th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: ring_simps)
+    have th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: field_simps)
     have th1: "\<And>(x::'a) y. x = - y \<longleftrightarrow> x + y = 0"
       apply (subst eq_iff_diff_eq_0) by simp
     have "x*x = 1 \<longleftrightarrow> x*x - 1 = 0" by simp
@@ -929,7 +929,7 @@
       unfolding dot_norm_neg dist_norm[symmetric]
       unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)}
   note fc = this
-  show ?thesis unfolding linear_def vector_eq smult_conv_scaleR by (simp add: inner_simps fc ring_simps)
+  show ?thesis unfolding linear_def vector_eq smult_conv_scaleR by (simp add: inner_simps fc field_simps)
 qed
 
 lemma isometry_linear:
@@ -980,7 +980,7 @@
       using H(5-9)
       apply (simp add: norm_eq norm_eq_1)
       apply (simp add: inner_simps smult_conv_scaleR) unfolding *
-      by (simp add: ring_simps) }
+      by (simp add: field_simps) }
   note th0 = this
   let ?g = "\<lambda>x. if x = 0 then 0 else norm x *s f (inverse (norm x) *s x)"
   {fix x:: "real ^'n" assume nx: "norm x = 1"
@@ -1079,7 +1079,7 @@
   unfolding permutes_sing
   apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq)
   apply (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31))
-  by (simp add: ring_simps)
+  by (simp add: field_simps)
 qed
 
 end
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -257,14 +257,14 @@
   | "vector_power x (Suc n) = x * vector_power x n"
 
 instance cart :: (semiring,finite) semiring
-  apply (intro_classes) by (vector ring_simps)+
+  apply (intro_classes) by (vector field_simps)+
 
 instance cart :: (semiring_0,finite) semiring_0
-  apply (intro_classes) by (vector ring_simps)+
+  apply (intro_classes) by (vector field_simps)+
 instance cart :: (semiring_1,finite) semiring_1
   apply (intro_classes) by vector
 instance cart :: (comm_semiring,finite) comm_semiring
-  apply (intro_classes) by (vector ring_simps)+
+  apply (intro_classes) by (vector field_simps)+
 
 instance cart :: (comm_semiring_0,finite) comm_semiring_0 by (intro_classes)
 instance cart :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..
@@ -278,7 +278,7 @@
 
 instance cart :: (real_algebra,finite) real_algebra
   apply intro_classes
-  apply (simp_all add: vector_scaleR_def ring_simps)
+  apply (simp_all add: vector_scaleR_def field_simps)
   apply vector
   apply vector
   done
@@ -318,19 +318,19 @@
 lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x"
   by (vector mult_assoc)
 lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x"
-  by (vector ring_simps)
+  by (vector field_simps)
 lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y"
-  by (vector ring_simps)
+  by (vector field_simps)
 lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector
 lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector
 lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y"
-  by (vector ring_simps)
+  by (vector field_simps)
 lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector
 lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector
 lemma vector_sneg_minus1: "-x = (- (1::'a::ring_1)) *s x" by vector
 lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector
 lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x"
-  by (vector ring_simps)
+  by (vector field_simps)
 
 lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"
   by (simp add: Cart_eq)
@@ -752,7 +752,7 @@
 lemma square_bound_lemma: "(x::real) < (1 + x) * (1 + x)"
 proof-
   have "(x + 1/2)^2 + 3/4 > 0" using zero_le_power2[of "x+1/2"] by arith
-  thus ?thesis by (simp add: ring_simps power2_eq_square)
+  thus ?thesis by (simp add: field_simps power2_eq_square)
 qed
 
 lemma square_continuous: "0 < (e::real) ==> \<exists>d. 0 < d \<and> (\<forall>y. abs(y - x) < d \<longrightarrow> abs(y * y - x * x) < e)"
@@ -828,7 +828,7 @@
 lemma norm_triangle_sub:
   fixes x y :: "'a::real_normed_vector"
   shows "norm x \<le> norm y  + norm (x - y)"
-  using norm_triangle_ineq[of "y" "x - y"] by (simp add: ring_simps)
+  using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps)
 
 lemma component_le_norm: "\<bar>x$i\<bar> <= norm x"
   apply (simp add: norm_vector_def)
@@ -867,7 +867,7 @@
 
 lemma real_abs_le_square_iff: "\<bar>x\<bar> \<le> \<bar>y\<bar> \<longleftrightarrow> (x::real)^2 \<le> y^2"
 proof-
-  have "x^2 \<le> y^2 \<longleftrightarrow> (x -y) * (y + x) \<le> 0" by (simp add: ring_simps power2_eq_square)
+  have "x^2 \<le> y^2 \<longleftrightarrow> (x -y) * (y + x) \<le> 0" by (simp add: field_simps power2_eq_square)
   also have "\<dots> \<longleftrightarrow> \<bar>x\<bar> \<le> \<bar>y\<bar>" apply (simp add: zero_compare_simps real_abs_def not_less) by arith
 finally show ?thesis ..
 qed
@@ -898,7 +898,7 @@
   unfolding power2_norm_eq_inner inner_simps inner_commute by auto 
 
 lemma dot_norm_neg: "x \<bullet> y = ((norm x ^ 2 + norm y ^ 2) - norm(x - y) ^ 2) / 2"
-  unfolding power2_norm_eq_inner inner_simps inner_commute by(auto simp add:group_simps)
+  unfolding power2_norm_eq_inner inner_simps inner_commute by(auto simp add:algebra_simps)
 
 text{* Equality of vectors in terms of @{term "op \<bullet>"} products.    *}
 
@@ -909,7 +909,7 @@
   assume ?rhs
   then have "x \<bullet> x - x \<bullet> y = 0 \<and> x \<bullet> y - y \<bullet> y = 0" by simp
   hence "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0" by (simp add: inner_simps inner_commute)
-  then have "(x - y) \<bullet> (x - y) = 0" by (simp add: ring_simps inner_simps inner_commute)
+  then have "(x - y) \<bullet> (x - y) = 0" by (simp add: field_simps inner_simps inner_commute)
   then show "x = y" by (simp)
 qed
 
@@ -930,7 +930,7 @@
   by (rule order_trans [OF norm_triangle_ineq add_mono])
 
 lemma ge_iff_diff_ge_0: "(a::'a::linordered_ring) \<ge> b == a - b \<ge> 0"
-  by (simp add: ring_simps)
+  by (simp add: field_simps)
 
 lemma pth_1:
   fixes x :: "'a::real_normed_vector"
@@ -1430,15 +1430,15 @@
   shows "linear f" using assms unfolding linear_def by auto
 
 lemma linear_compose_cmul: "linear f ==> linear (\<lambda>x. (c::'a::comm_semiring) *s f x)"
-  by (vector linear_def Cart_eq ring_simps)
+  by (vector linear_def Cart_eq field_simps)
 
 lemma linear_compose_neg: "linear (f :: 'a ^'n \<Rightarrow> 'a::comm_ring ^'m) ==> linear (\<lambda>x. -(f(x)))" by (vector linear_def Cart_eq)
 
 lemma linear_compose_add: "linear (f :: 'a ^'n \<Rightarrow> 'a::semiring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f(x) + g(x))"
-  by (vector linear_def Cart_eq ring_simps)
+  by (vector linear_def Cart_eq field_simps)
 
 lemma linear_compose_sub: "linear (f :: 'a ^'n \<Rightarrow> 'a::ring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f x - g x)"
-  by (vector linear_def Cart_eq ring_simps)
+  by (vector linear_def Cart_eq field_simps)
 
 lemma linear_compose: "linear f \<Longrightarrow> linear g ==> linear (g o f)"
   by (simp add: linear_def)
@@ -1460,7 +1460,7 @@
   shows "linear (\<lambda>x. f x $ k *s v)"
   using lf
   apply (auto simp add: linear_def )
-  by (vector ring_simps)+
+  by (vector field_simps)+
 
 lemma linear_0: "linear f ==> f 0 = (0::'a::semiring_1 ^'n)"
   unfolding linear_def
@@ -1536,7 +1536,7 @@
       unfolding norm_mul
       apply (simp only: mult_commute)
       apply (rule mult_mono)
-      by (auto simp add: ring_simps norm_ge_zero) }
+      by (auto simp add: field_simps norm_ge_zero) }
     then have th: "\<forall>i\<in> ?S. norm ((x$i) *s f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x" by metis
     from real_setsum_norm_le[OF fS, of "\<lambda>i. (x$i) *s (f (basis i))", OF th]
     have "norm (f x) \<le> ?B * norm x" unfolding th0 setsum_left_distrib by metis}
@@ -1562,7 +1562,7 @@
     {fix x::"real ^ 'n"
       have "norm (f x) \<le> ?K *  norm x"
       using B[rule_format, of x] norm_ge_zero[of x] norm_ge_zero[of "f x"] Bp
-      apply (auto simp add: ring_simps split add: abs_split)
+      apply (auto simp add: field_simps split add: abs_split)
       apply (erule order_trans, simp)
       done
   }
@@ -1631,12 +1631,12 @@
 lemma bilinear_lzero:
   fixes h :: "'a::ring^'n \<Rightarrow> _" assumes bh: "bilinear h" shows "h 0 x = 0"
   using bilinear_ladd[OF bh, of 0 0 x]
-    by (simp add: eq_add_iff ring_simps)
+    by (simp add: eq_add_iff field_simps)
 
 lemma bilinear_rzero:
   fixes h :: "'a::ring^_ \<Rightarrow> _" assumes bh: "bilinear h" shows "h x 0 = 0"
   using bilinear_radd[OF bh, of x 0 0 ]
-    by (simp add: eq_add_iff ring_simps)
+    by (simp add: eq_add_iff field_simps)
 
 lemma bilinear_lsub: "bilinear h ==> h (x - (y:: 'a::ring_1 ^ _)) z = h x z - h y z"
   by (simp  add: diff_def bilinear_ladd bilinear_lneg)
@@ -1677,7 +1677,7 @@
       apply (rule real_setsum_norm_le)
       using fN fM
       apply simp
-      apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] norm_mul ring_simps)
+      apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] norm_mul field_simps)
       apply (rule mult_mono)
       apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm)
       apply (rule mult_mono)
@@ -1767,7 +1767,7 @@
         by (simp add: linear_cmul[OF lf])
       finally have "f x \<bullet> y = x \<bullet> ?w"
         apply (simp only: )
-        apply (simp add: inner_vector_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] ring_simps)
+        apply (simp add: inner_vector_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] field_simps)
         done}
   }
   then show ?thesis unfolding adjoint_def
@@ -1832,7 +1832,7 @@
 
 lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
 lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)"
-  by (vector matrix_matrix_mult_def setsum_addf[symmetric] ring_simps)
+  by (vector matrix_matrix_mult_def setsum_addf[symmetric] field_simps)
 
 lemma matrix_mul_lid:
   fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
@@ -1951,7 +1951,7 @@
 where "matrix f = (\<chi> i j. (f(basis j))$i)"
 
 lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::'a::comm_semiring_1 ^ _))"
-  by (simp add: linear_def matrix_vector_mult_def Cart_eq ring_simps setsum_right_distrib setsum_addf)
+  by (simp add: linear_def matrix_vector_mult_def Cart_eq field_simps setsum_right_distrib setsum_addf)
 
 lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::'a::comm_ring_1 ^ 'n)"
 apply (simp add: matrix_def matrix_vector_mult_def Cart_eq mult_commute)
@@ -2005,7 +2005,7 @@
 proof-
   have uv': "u = 0 \<longrightarrow> v \<noteq> 0" using u v uv by arith
   have "a = a * (u + v)" unfolding uv  by simp
-  hence th: "u * a + v * a = a" by (simp add: ring_simps)
+  hence th: "u * a + v * a = a" by (simp add: field_simps)
   from xa u have "u \<noteq> 0 \<Longrightarrow> u*x < u*a" by (simp add: mult_compare_simps)
   from ya v have "v \<noteq> 0 \<Longrightarrow> v * y < v * a" by (simp add: mult_compare_simps)
   from xa ya u v have "u * x + v * y < u * a + v * a"
@@ -2028,7 +2028,7 @@
   shows "u * x + v * y \<le> a"
 proof-
   from xa ya u v have "u * x + v * y \<le> u * a + v * a" by (simp add: add_mono mult_left_mono)
-  also have "\<dots> \<le> (u + v) * a" by (simp add: ring_simps)
+  also have "\<dots> \<le> (u + v) * a" by (simp add: field_simps)
   finally show ?thesis unfolding uv by simp
 qed
 
@@ -2049,7 +2049,7 @@
   shows "x <= y + z"
 proof-
   have "y^2 + z^2 \<le> y^2 + 2*y*z + z^2" using z y  by (simp add: zero_compare_simps)
-  with xy have th: "x ^2 \<le> (y+z)^2" by (simp add: power2_eq_square ring_simps)
+  with xy have th: "x ^2 \<le> (y+z)^2" by (simp add: power2_eq_square field_simps)
   from y z have yz: "y + z \<ge> 0" by arith
   from power2_le_imp_le[OF th yz] show ?thesis .
 qed
@@ -2534,9 +2534,9 @@
   from h have p: "1 \<le> (1 + x) ^ n" using Suc.prems by simp
   from h have "1 + real n * x + x \<le> (1 + x) ^ n + x" by simp
   also have "\<dots> \<le> (1 + x) ^ Suc n" apply (subst diff_le_0_iff_le[symmetric])
-    apply (simp add: ring_simps)
+    apply (simp add: field_simps)
     using mult_left_mono[OF p Suc.prems] by simp
-  finally show ?case  by (simp add: real_of_nat_Suc ring_simps)
+  finally show ?case  by (simp add: real_of_nat_Suc field_simps)
 qed
 
 lemma real_arch_pow: assumes x: "1 < (x::real)" shows "\<exists>n. y < x^n"
@@ -2602,10 +2602,10 @@
     from geometric_sum[OF x1, of "Suc n", unfolded x1']
     have "(- (1 - x)) * setsum (\<lambda>i. x^i) {0 .. n} = - (1 - x^(Suc n))"
       unfolding atLeastLessThanSuc_atLeastAtMost
-      using x1' apply (auto simp only: field_eq_simps)
-      apply (simp add: ring_simps)
+      using x1' apply (auto simp only: field_simps)
+      apply (simp add: field_simps)
       done
-    then have ?thesis by (simp add: ring_simps) }
+    then have ?thesis by (simp add: field_simps) }
   ultimately show ?thesis by metis
 qed
 
@@ -2624,7 +2624,7 @@
   from setsum_reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]
   have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})" by simp
   then show ?thesis unfolding sum_gp_basic using mn
-    by (simp add: ring_simps power_add[symmetric])
+    by (simp add: field_simps power_add[symmetric])
 qed
 
 lemma sum_gp: "setsum (op ^ (x::'a::{field})) {m .. n} =
@@ -2637,7 +2637,7 @@
     {assume x: "x = 1"  hence ?thesis by simp}
     moreover
     {assume x: "x \<noteq> 1" hence nz: "1 - x \<noteq> 0" by simp
-      from sum_gp_multiplied[OF nm, of x] nz have ?thesis by (simp add: field_eq_simps)}
+      from sum_gp_multiplied[OF nm, of x] nz have ?thesis by (simp add: field_simps)}
     ultimately have ?thesis by metis
   }
   ultimately show ?thesis by metis
@@ -2646,7 +2646,7 @@
 lemma sum_gp_offset: "setsum (op ^ (x::'a::{field})) {m .. m+n} =
   (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
   unfolding sum_gp[of x m "m + n"] power_Suc
-  by (simp add: ring_simps power_add)
+  by (simp add: field_simps power_add)
 
 
 subsection{* A bit of linear algebra. *}
@@ -2920,14 +2920,14 @@
     apply (simp only: )
     apply (rule span_add[unfolded mem_def])
     apply assumption+
-    apply (vector ring_simps)
+    apply (vector field_simps)
     apply (clarsimp simp add: mem_def)
     apply (rule_tac x= "c*k" in exI)
     apply (subgoal_tac "c *s x - (c * k) *s b = c*s (x - k*s b)")
     apply (simp only: )
     apply (rule span_mul[unfolded mem_def])
     apply assumption
-    by (vector ring_simps)
+    by (vector field_simps)
   ultimately show "?P a" using aS span_induct[where S=S and P= "?P"] by metis
 qed
 
@@ -3073,7 +3073,7 @@
           setsum_clauses(2)[OF fS] cong del: if_weak_cong)
       also have "\<dots> = (\<Sum>v\<in>S. u v *s v) + c *s x"
         apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]])
-        by (vector ring_simps)
+        by (vector field_simps)
       also have "\<dots> = c*s x + y"
         by (simp add: add_commute u)
       finally have "setsum (\<lambda>v. ?u v *s v) ?S = c*s x + y" .
@@ -3110,7 +3110,7 @@
     from fS SP aP have th0: "finite ?S" "?S \<subseteq> P" "?v \<in> ?S" "?u ?v \<noteq> 0" by auto
     have s0: "setsum (\<lambda>v. ?u v *s v) ?S = 0"
       using fS aS
-      apply (simp add: vector_smult_lneg vector_smult_lid setsum_clauses ring_simps )
+      apply (simp add: vector_smult_lneg vector_smult_lid setsum_clauses field_simps )
       apply (subst (2) ua[symmetric])
       apply (rule setsum_cong2)
       by auto
@@ -3643,7 +3643,7 @@
   from C(1) have fC: "finite ?C" by simp
   from fB aB C(1,2) have cC: "card ?C \<le> card (insert a B)" by (simp add: card_insert_if)
   {fix x k
-    have th0: "\<And>(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: ring_simps)
+    have th0: "\<And>(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: field_simps)
     have "x - k *s (a - (\<Sum>x\<in>C. (x \<bullet> a / (x \<bullet> x)) *s x)) \<in> span C \<longleftrightarrow> x - k *s a \<in> span C"
       apply (simp only: vector_ssub_ldistrib th0)
       apply (rule span_add_eq)
@@ -3863,7 +3863,7 @@
       using z .
     {fix k assume k: "z - k *s a \<in> span b"
       have eq: "z - ?h z *s a - (z - k*s a) = (k - ?h z) *s a"
-        by (simp add: ring_simps vector_sadd_rdistrib[symmetric])
+        by (simp add: field_simps vector_sadd_rdistrib[symmetric])
       from span_sub[OF th0 k]
       have khz: "(k - ?h z) *s a \<in> span b" by (simp add: eq)
       {assume "k \<noteq> ?h z" hence k0: "k - ?h z \<noteq> 0" by simp
@@ -3877,7 +3877,7 @@
   let ?g = "\<lambda>z. ?h z *s f a + g (z - ?h z *s a)"
   {fix x y assume x: "x \<in> span (insert a b)" and y: "y \<in> span (insert a b)"
     have tha: "\<And>(x::'a^'n) y a k l. (x + y) - (k + l) *s a = (x - k *s a) + (y - l *s a)"
-      by (vector ring_simps)
+      by (vector field_simps)
     have addh: "?h (x + y) = ?h x + ?h y"
       apply (rule conjunct2[OF h, rule_format, symmetric])
       apply (rule span_add[OF x y])
@@ -3890,14 +3890,14 @@
   moreover
   {fix x:: "'a^'n" and c:: 'a  assume x: "x \<in> span (insert a b)"
     have tha: "\<And>(x::'a^'n) c k a. c *s x - (c * k) *s a = c *s (x - k *s a)"
-      by (vector ring_simps)
+      by (vector field_simps)
     have hc: "?h (c *s x) = c * ?h x"
       apply (rule conjunct2[OF h, rule_format, symmetric])
       apply (metis span_mul x)
       by (metis tha span_mul x conjunct1[OF h])
     have "?g (c *s x) = c*s ?g x"
       unfolding hc tha g(2)[rule_format, OF conjunct1[OF h, OF x]]
-      by (vector ring_simps)}
+      by (vector field_simps)}
   moreover
   {fix x assume x: "x \<in> (insert a b)"
     {assume xa: "x = a"
@@ -4276,7 +4276,7 @@
             fix j
             have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
            else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" using i(1)
-              by (simp add: ring_simps)
+              by (simp add: field_simps)
             have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
            else (x$xa) * ((column xa A$j))) ?U = setsum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
               apply (rule setsum_cong[OF refl])
@@ -4619,7 +4619,7 @@
   from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"]
   have ths: "infnorm x \<le> infnorm (x - y) + infnorm y"
     "infnorm y \<le> infnorm (x - y) + infnorm x"
-    by (simp_all add: ring_simps infnorm_neg diff_def[symmetric])
+    by (simp_all add: field_simps infnorm_neg diff_def[symmetric])
   from th[OF ths]  show ?thesis .
 qed
 
@@ -4718,9 +4718,9 @@
       using x y
       unfolding inner_simps smult_conv_scaleR
       unfolding power2_norm_eq_inner[symmetric] power2_eq_square diff_eq_0_iff_eq apply (simp add: inner_commute)
-      apply (simp add: ring_simps) by metis
+      apply (simp add: field_simps) by metis
     also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" using x y
-      by (simp add: ring_simps inner_commute)
+      by (simp add: field_simps inner_commute)
     also have "\<dots> \<longleftrightarrow> ?lhs" using x y
       apply simp
       by metis
@@ -4766,7 +4766,7 @@
     also have "\<dots> \<longleftrightarrow> norm x *s y = norm y *s x"
       unfolding norm_cauchy_schwarz_eq[symmetric]
       unfolding power2_norm_eq_inner inner_simps
-      by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute ring_simps)
+      by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps)
     finally have ?thesis .}
   ultimately show ?thesis by blast
 qed
@@ -4852,10 +4852,10 @@
 unfolding vector_smult_assoc
 unfolding norm_mul
 apply (subgoal_tac "norm x * c = \<bar>c\<bar> * norm x \<or> norm x * c = - \<bar>c\<bar> * norm x")
-apply (case_tac "c <= 0", simp add: ring_simps)
-apply (simp add: ring_simps)
-apply (case_tac "c <= 0", simp add: ring_simps)
-apply (simp add: ring_simps)
+apply (case_tac "c <= 0", simp add: field_simps)
+apply (simp add: field_simps)
+apply (case_tac "c <= 0", simp add: field_simps)
+apply (simp add: field_simps)
 apply simp
 apply simp
 done
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -1131,7 +1131,7 @@
     guess d2 by(rule has_integralD[OF goal1(2) e]) note d2=this
     guess p by(rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)],of a b]) note p=this
     let ?c = "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" have "norm (k1 - k2) \<le> norm (?c - k2) + norm (?c - k1)"
-      using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] by(auto simp add:group_simps norm_minus_commute)
+      using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] by(auto simp add:algebra_simps norm_minus_commute)
     also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2"
       apply(rule add_strict_mono) apply(rule_tac[!] d2(2) d1(2)) using p unfolding fine_def by auto
     finally show False by auto
@@ -1244,7 +1244,7 @@
           unfolding scaleR_right_distrib setsum_addf[of "\<lambda>(x,k). content k *\<^sub>R f x" "\<lambda>(x,k). content k *\<^sub>R g x" p,THEN sym]
           by(rule setsum_cong2,auto)
         have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) = norm (((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l))"
-          unfolding * by(auto simp add:group_simps) also let ?res = "\<dots>"
+          unfolding * by(auto simp add:algebra_simps) also let ?res = "\<dots>"
         from as have *:"d1 fine p" "d2 fine p" unfolding fine_inter by auto
         have "?res < e/2 + e/2" apply(rule le_less_trans[OF norm_triangle_ineq])
           apply(rule add_strict_mono) using d1(2)[OF as(1) *(1)] and d2(2)[OF as(1) *(2)] by auto
@@ -1268,7 +1268,7 @@
 
 lemma has_integral_sub:
   shows "(f has_integral k) s \<Longrightarrow> (g has_integral l) s \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) has_integral (k - l)) s"
-  using has_integral_add[OF _ has_integral_neg,of f k s g l] unfolding group_simps by auto
+  using has_integral_add[OF _ has_integral_neg,of f k s g l] unfolding algebra_simps by auto
 
 lemma integral_0: "integral s (\<lambda>x::real^'n. 0::real^'m) = 0"
   by(rule integral_unique has_integral_0)+
@@ -1703,7 +1703,7 @@
       proof- guess p using fine_division_exists[OF d(1), of a' b] . note p=this
         show ?thesis using norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]]
           using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
-          using p using assms by(auto simp add:group_simps)
+          using p using assms by(auto simp add:algebra_simps)
       qed qed  
     show "?P {x. x $ k \<ge> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
     proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p2"
@@ -1711,7 +1711,7 @@
       proof- guess p using fine_division_exists[OF d(1), of a b'] . note p=this
         show ?thesis using norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]]
           using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
-          using p using assms by(auto simp add:group_simps) qed qed qed qed
+          using p using assms by(auto simp add:algebra_simps) qed qed qed qed
 
 subsection {* Generalized notion of additivity. *}
 
@@ -1848,7 +1848,7 @@
 
 lemma monoidal_monoid[intro]:
   shows "monoidal ((op +)::('a::comm_monoid_add) \<Rightarrow> 'a \<Rightarrow> 'a)"
-  unfolding monoidal_def neutral_monoid by(auto simp add: group_simps) 
+  unfolding monoidal_def neutral_monoid by(auto simp add: algebra_simps) 
 
 lemma operative_integral: fixes f::"real^'n \<Rightarrow> 'a::banach"
   shows "operative (lifted(op +)) (\<lambda>i. if f integrable_on i then Some(integral i f) else None)"
@@ -2381,8 +2381,8 @@
       have lem2:"\<And>s1 s2 i1 i2. norm(s2 - s1) \<le> e/2 \<Longrightarrow> norm(s1 - i1) < e / 4 \<Longrightarrow> norm(s2 - i2) < e / 4 \<Longrightarrow>norm(i1 - i2) < e"
       proof- case goal1 have "norm (i1 - i2) \<le> norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)"
           using norm_triangle_ineq[of "i1 - s1" "s1 - i2"]
-          using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by(auto simp add:group_simps)
-        also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_simps)
+          using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by(auto simp add:algebra_simps)
+        also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:algebra_simps)
         finally show ?case .
       qed
       show ?case unfolding vector_dist_norm apply(rule lem2) defer
@@ -2399,7 +2399,7 @@
         also have "\<dots> = 2 / real M" unfolding real_divide_def by auto
         finally show "norm (g n x - g m x) \<le> 2 / real M"
           using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"]
-          by(auto simp add:group_simps simp add:norm_minus_commute)
+          by(auto simp add:algebra_simps simp add:norm_minus_commute)
       qed qed qed
   from this[unfolded convergent_eq_cauchy[THEN sym]] guess s .. note s=this
 
@@ -2413,8 +2413,8 @@
     have lem:"\<And>sf sg i. norm(sf - sg) \<le> e / 3 \<Longrightarrow> norm(i - s) < e / 3 \<Longrightarrow> norm(sg - i) < e / 3 \<Longrightarrow> norm(sf - s) < e"
     proof- case goal1 have "norm (sf - s) \<le> norm (sf - sg) + norm (sg - i) + norm (i - s)"
         using norm_triangle_ineq[of "sf - sg" "sg - s"]
-        using norm_triangle_ineq[of "sg -  i" " i - s"] by(auto simp add:group_simps)
-      also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_simps)
+        using norm_triangle_ineq[of "sg -  i" " i - s"] by(auto simp add:algebra_simps)
+      also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:algebra_simps)
       finally show ?case .
     qed
     show ?case apply(rule_tac x=g' in exI) apply(rule,rule g')
@@ -2956,7 +2956,7 @@
       have ball:"\<forall>xa\<in>k. xa \<in> ball x (d (dest_vec1 x))" using as(2)[unfolded fine_def,rule_format,OF `(x,k)\<in>p`,unfolded split_conv subset_eq] .
       have "norm ((v$1 - u$1) *\<^sub>R f' x - (f v - f u)) \<le> norm (f u - f x - (u$1 - x$1) *\<^sub>R f' x) + norm (f v - f x - (v$1 - x$1) *\<^sub>R f' x)"
         apply(rule order_trans[OF _ norm_triangle_ineq4]) apply(rule eq_refl) apply(rule arg_cong[where f=norm])
-        unfolding scaleR.diff_left by(auto simp add:group_simps)
+        unfolding scaleR.diff_left by(auto simp add:algebra_simps)
       also have "... \<le> e * norm (dest_vec1 u - dest_vec1 x) + e * norm (dest_vec1 v - dest_vec1 x)"
         apply(rule add_mono) apply(rule d(2)[of "x$1" "u$1",unfolded o_def vec1_dest_vec1]) prefer 4
         apply(rule d(2)[of "x$1" "v$1",unfolded o_def vec1_dest_vec1])
@@ -3098,7 +3098,7 @@
   proof(rule,rule,rule d,safe) case goal1 show ?case proof(cases "y < x")
       case False have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 y}" apply(rule integrable_subinterval,rule integrable_continuous)
         apply(rule continuous_on_o_dest_vec1 assms)+  unfolding not_less using assms(2) goal1 by auto
-      hence *:"?I a y - ?I a x = ?I x y" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine)
+      hence *:"?I a y - ?I a x = ?I x y" unfolding algebra_simps apply(subst eq_commute) apply(rule integral_combine)
         using False unfolding not_less using assms(2) goal1 by auto
       have **:"norm (y - x) = content {vec1 x..vec1 y}" apply(subst content_1) using False unfolding not_less by auto
       show ?thesis unfolding ** apply(rule has_integral_bound[where f="(\<lambda>u. f u - f x) o dest_vec1"]) unfolding * unfolding o_def
@@ -3112,7 +3112,7 @@
       qed(insert e,auto)
     next case True have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 x}" apply(rule integrable_subinterval,rule integrable_continuous)
         apply(rule continuous_on_o_dest_vec1 assms)+  unfolding not_less using assms(2) goal1 by auto
-      hence *:"?I a x - ?I a y = ?I y x" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine)
+      hence *:"?I a x - ?I a y = ?I y x" unfolding algebra_simps apply(subst eq_commute) apply(rule integral_combine)
         using True using assms(2) goal1 by auto
       have **:"norm (y - x) = content {vec1 y..vec1 x}" apply(subst content_1) using True unfolding not_less by auto
       have ***:"\<And>fy fx c::'a. fx - fy - (y - x) *\<^sub>R c = -(fy - fx - (x - y) *\<^sub>R c)" unfolding scaleR_left.diff by auto 
@@ -3194,7 +3194,7 @@
           apply(rule_tac X="g ` X" in UnionI) defer apply(rule_tac x="h x" in image_eqI)
           using X(2) assms(3)[rule_format,of x] by auto
       qed note ** = d(2)[OF this] have *:"inj_on (\<lambda>(x, k). (g x, g ` k)) p" using inj(1) unfolding inj_on_def by fastsimp
-       have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _") unfolding group_simps add_left_cancel
+       have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _") unfolding algebra_simps add_left_cancel
         unfolding setsum_reindex[OF *] apply(subst scaleR_right.setsum) defer apply(rule setsum_cong2) unfolding o_def split_paired_all split_conv
         apply(drule p(4)) apply safe unfolding assms(7)[rule_format] using p by auto
       also have "... = r *\<^sub>R ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r") unfolding scaleR.diff_right scaleR.scaleR_left[THEN sym]
@@ -3332,7 +3332,7 @@
 
 lemma norm_triangle_le_sub: "norm x + norm y \<le> e \<Longrightarrow> norm (x - y) \<le> e"
   apply(subst(asm)(2) norm_minus_cancel[THEN sym])
-  apply(drule norm_triangle_le) by(auto simp add:group_simps)
+  apply(drule norm_triangle_le) by(auto simp add:algebra_simps)
 
 lemma fundamental_theorem_of_calculus_interior:
   assumes"a \<le> b" "continuous_on {a..b} f" "\<forall>x\<in>{a<..<b}. (f has_vector_derivative f'(x)) (at x)"
@@ -3641,11 +3641,11 @@
   proof safe show "0 < ?d" using d(1) assms(3) unfolding Cart_simps by auto
     fix t::"_^1" assume as:"c \<le> t" "t$1 < c$1 + ?d"
     have *:"integral{a..c} f = integral{a..b} f - integral{c..b} f"
-      "integral{a..t} f = integral{a..b} f - integral{t..b} f" unfolding group_simps
+      "integral{a..t} f = integral{a..b} f - integral{t..b} f" unfolding algebra_simps
       apply(rule_tac[!] integral_combine) using assms as unfolding Cart_simps by auto
     have "(- c)$1 - d < (- t)$1 \<and> - t \<le> - c" using as by auto note d(2)[rule_format,OF this]
     thus "norm (integral {a..c} f - integral {a..t} f) < e" unfolding * 
-      unfolding integral_reflect apply-apply(subst norm_minus_commute) by(auto simp add:group_simps) qed qed
+      unfolding integral_reflect apply-apply(subst norm_minus_commute) by(auto simp add:algebra_simps) qed qed
 
 declare dest_vec1_eq[simp del] not_less[simp] not_le[simp]
 
@@ -3715,7 +3715,7 @@
     apply safe apply(rule conv) using assms(4,7) by auto
   have *:"\<And>t xa. (1 - t) *\<^sub>R c + t *\<^sub>R x = (1 - xa) *\<^sub>R c + xa *\<^sub>R x \<Longrightarrow> t = xa"
   proof- case goal1 hence "(t - xa) *\<^sub>R x = (t - xa) *\<^sub>R c" 
-      unfolding scaleR_simps by(auto simp add:group_simps)
+      unfolding scaleR_simps by(auto simp add:algebra_simps)
     thus ?case using `x\<noteq>c` by auto qed
   have as2:"finite {t. ((1 - t) *\<^sub>R c + t *\<^sub>R x) \<in> k}" using assms(2) 
     apply(rule finite_surj[where f="\<lambda>z. SOME t. (1-t) *\<^sub>R c + t *\<^sub>R x = z"])
@@ -4390,7 +4390,7 @@
   have *:"\<And>ir ip i cr cp. norm((cp + cr) - i) < e \<Longrightarrow> norm(cr - ir) < k \<Longrightarrow> 
     ip + ir = i \<Longrightarrow> norm(cp - ip) \<le> e + k" 
   proof- case goal1 thus ?case  using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"]  
-      unfolding goal1(3)[THEN sym] norm_minus_cancel by(auto simp add:group_simps) qed
+      unfolding goal1(3)[THEN sym] norm_minus_cancel by(auto simp add:algebra_simps) qed
   
   have "?x =  norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p. integral k f))"
     unfolding split_def setsum_subtractf ..
@@ -4501,7 +4501,7 @@
             norm(c - d) < e / 4 \<longrightarrow> norm(a - d) < e" 
       proof safe case goal1 thus ?case using norm_triangle_lt[of "a - b" "b - c" "3* e/4"]
           norm_triangle_lt[of "a - b + (b - c)" "c - d" e] unfolding norm_minus_cancel
-          by(auto simp add:group_simps) qed
+          by(auto simp add:algebra_simps) qed
       show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - i) < e" apply(rule *[rule_format,where
           b="\<Sum>(x, k)\<in>p. content k *\<^sub>R f (m x) x" and c="\<Sum>(x, k)\<in>p. integral k (f (m x))"])
       proof safe case goal1
@@ -5152,7 +5152,7 @@
   assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s"
   shows "(\<lambda>x. f(x) - g(x)) absolutely_integrable_on s"
   using absolutely_integrable_add[OF assms(1) absolutely_integrable_neg[OF assms(2)]]
-  unfolding group_simps .
+  unfolding algebra_simps .
 
 lemma absolutely_integrable_linear: fixes f::"real^'m \<Rightarrow> real^'n" and h::"real^'n \<Rightarrow> real^'p"
   assumes "f absolutely_integrable_on s" "bounded_linear h"
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -4442,7 +4442,7 @@
   let ?D = "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}"
   obtain a where a:"\<forall>x\<in>s. norm x \<le> a" using assms[unfolded bounded_iff] by auto
   { fix x y assume "x \<in> s" "y \<in> s"
-    hence "norm (x - y) \<le> 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps)  }
+    hence "norm (x - y) \<le> 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: field_simps)  }
   note * = this
   { fix x y assume "x\<in>s" "y\<in>s"  hence "s \<noteq> {}" by auto
     have "norm(x - y) \<le> diameter s" unfolding diameter_def using `s\<noteq>{}` *[OF `x\<in>s` `y\<in>s`] `x\<in>s` `y\<in>s`  
@@ -5752,7 +5752,7 @@
   shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)"
 proof
   assume h: "m *s x + c = y"
-  hence "m *s x = y - c" by (simp add: ring_simps)
+  hence "m *s x = y - c" by (simp add: field_simps)
   hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp
   then show "x = inverse m *s y + - (inverse m *s c)"
     using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
@@ -5854,11 +5854,11 @@
       also have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)"
         using cf_z[of "m + k"] and c by auto
       also have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d"
-        using Suc by (auto simp add: ring_simps)
+        using Suc by (auto simp add: field_simps)
       also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)"
-        unfolding power_add by (auto simp add: ring_simps)
+        unfolding power_add by (auto simp add: field_simps)
       also have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)"
-        using c by (auto simp add: ring_simps)
+        using c by (auto simp add: field_simps)
       finally show ?case by auto
     qed
   } note cf_z2 = this
@@ -6015,7 +6015,7 @@
         apply(erule_tac x="Na+Nb+n" in allE) apply simp
         using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)"
           "-b"  "- f (r (Na + Nb + n)) y"]
-        unfolding ** unfolding group_simps(12) by (auto simp add: dist_commute)
+        unfolding ** by (auto simp add: algebra_simps dist_commute)
       moreover
       have "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) \<ge> dist a b - dist (f n x) (f n y)"
         using distf[of n "r (Na+Nb+n)", OF _ `x\<in>s` `y\<in>s`]
--- a/src/HOL/NSA/HyperDef.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/NSA/HyperDef.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -140,12 +140,12 @@
 
 lemma of_hypreal_inverse [simp]:
   "\<And>x. of_hypreal (inverse x) =
-   inverse (of_hypreal x :: 'a::{real_div_algebra,division_by_zero} star)"
+   inverse (of_hypreal x :: 'a::{real_div_algebra,division_ring_inverse_zero} star)"
 by transfer (rule of_real_inverse)
 
 lemma of_hypreal_divide [simp]:
   "\<And>x y. of_hypreal (x / y) =
-   (of_hypreal x / of_hypreal y :: 'a::{real_field,division_by_zero} star)"
+   (of_hypreal x / of_hypreal y :: 'a::{real_field,division_ring_inverse_zero} star)"
 by transfer (rule of_real_divide)
 
 lemma of_hypreal_eq_iff [simp]:
@@ -454,7 +454,7 @@
 by transfer (rule field_power_not_zero)
 
 lemma hyperpow_inverse:
-  "\<And>r n. r \<noteq> (0::'a::{division_by_zero,field} star)
+  "\<And>r n. r \<noteq> (0::'a::{division_ring_inverse_zero,field} star)
    \<Longrightarrow> inverse (r pow n) = (inverse r) pow n"
 by transfer (rule power_inverse)
   
--- a/src/HOL/Number_Theory/Binomial.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Number_Theory/Binomial.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -208,7 +208,7 @@
   have "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) =
       fact (k + 1) * fact (n - k) * (n choose (k + 1)) + 
       fact (k + 1) * fact (n - k) * (n choose k)" 
-    by (subst choose_reduce_nat, auto simp add: ring_simps)
+    by (subst choose_reduce_nat, auto simp add: field_simps)
   also note one
   also note two
   also with less have "(n - k) * fact n + (k + 1) * fact n= fact (n + 1)" 
@@ -279,7 +279,7 @@
   also have "... = (SUM k=0..n. of_nat (n choose k) * a^k * b^(n+1-k)) +
                   (SUM k=1..n+1. of_nat (n choose (k - 1)) * a^k * b^(n+1-k))"
     by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le 
-             power_Suc ring_simps One_nat_def del:setsum_cl_ivl_Suc)
+             power_Suc field_simps One_nat_def del:setsum_cl_ivl_Suc)
   also have "... = a^(n+1) + b^(n+1) +
                   (SUM k=1..n. of_nat (n choose (k - 1)) * a^k * b^(n+1-k)) +
                   (SUM k=1..n. of_nat (n choose k) * a^k * b^(n+1-k))"
@@ -287,10 +287,10 @@
   also have
       "... = a^(n+1) + b^(n+1) + 
          (SUM k=1..n. of_nat(n+1 choose k) * a^k * b^(n+1-k))"
-    by (auto simp add: ring_simps setsum_addf [symmetric]
+    by (auto simp add: field_simps setsum_addf [symmetric]
       choose_reduce_nat)
   also have "... = (SUM k=0..n+1. of_nat (n+1 choose k) * a^k * b^(n+1-k))"
-    using decomp by (simp add: ring_simps)
+    using decomp by (simp add: field_simps)
   finally show "?P (n + 1)" by simp
 qed
 
--- a/src/HOL/Number_Theory/Cong.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Number_Theory/Cong.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -350,7 +350,7 @@
   apply (subst prime_dvd_mult_eq_int [symmetric], assumption)
   (* any way around this? *)
   apply (subgoal_tac "a * a - 1 = (a - 1) * (a - -1)")
-  apply (auto simp add: ring_simps)
+  apply (auto simp add: field_simps)
 done
 
 lemma cong_mult_rcancel_int:
@@ -416,7 +416,7 @@
 done
 
 lemma cong_iff_lin_int: "([(a::int) = b] (mod m)) = (\<exists>k. b = a + m * k)"
-  apply (auto simp add: cong_altdef_int dvd_def ring_simps)
+  apply (auto simp add: cong_altdef_int dvd_def field_simps)
   apply (rule_tac [!] x = "-k" in exI, auto)
 done
 
@@ -428,14 +428,14 @@
   apply (unfold dvd_def, auto)
   apply (rule_tac x = k in exI)
   apply (rule_tac x = 0 in exI)
-  apply (auto simp add: ring_simps)
+  apply (auto simp add: field_simps)
   apply (subst (asm) cong_sym_eq_nat)
   apply (subst (asm) cong_altdef_nat)
   apply force
   apply (unfold dvd_def, auto)
   apply (rule_tac x = 0 in exI)
   apply (rule_tac x = k in exI)
-  apply (auto simp add: ring_simps)
+  apply (auto simp add: field_simps)
   apply (unfold cong_nat_def)
   apply (subgoal_tac "a mod m = (a + k2 * m) mod m")
   apply (erule ssubst)back
@@ -533,7 +533,7 @@
   apply (auto simp add: cong_iff_lin_nat dvd_def)
   apply (rule_tac x="k1 * k" in exI)
   apply (rule_tac x="k2 * k" in exI)
-  apply (simp add: ring_simps)
+  apply (simp add: field_simps)
 done
 
 lemma cong_dvd_modulus_int: "[(x::int) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> 
@@ -559,7 +559,7 @@
 lemma neg_cong_int: "([(a::int) = b] (mod m)) = ([-a = -b] (mod m))"
   apply (simp add: cong_altdef_int)
   apply (subst dvd_minus_iff [symmetric])
-  apply (simp add: ring_simps)
+  apply (simp add: field_simps)
 done
 
 lemma cong_modulus_neg_int: "([(a::int) = b] (mod m)) = ([a = b] (mod -m))"
@@ -603,7 +603,7 @@
   apply (unfold dvd_def)
   apply auto [1]
   apply (rule_tac x = k in exI)
-  apply (auto simp add: ring_simps) [1]
+  apply (auto simp add: field_simps) [1]
   apply (subst cong_altdef_nat)
   apply (auto simp add: dvd_def)
 done
@@ -611,7 +611,7 @@
 lemma cong_le_nat: "(y::nat) <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
   apply (subst cong_altdef_nat)
   apply assumption
-  apply (unfold dvd_def, auto simp add: ring_simps)
+  apply (unfold dvd_def, auto simp add: field_simps)
   apply (rule_tac x = k in exI)
   apply auto
 done
--- a/src/HOL/Number_Theory/Fib.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Number_Theory/Fib.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -143,9 +143,9 @@
   apply (induct n rule: fib_induct_nat)
   apply auto
   apply (subst fib_reduce_nat)
-  apply (auto simp add: ring_simps)
+  apply (auto simp add: field_simps)
   apply (subst (1 3 5) fib_reduce_nat)
-  apply (auto simp add: ring_simps Suc_eq_plus1)
+  apply (auto simp add: field_simps Suc_eq_plus1)
 (* hmmm. Why doesn't "n + (1 + (1 + k))" simplify to "n + k + 2"? *)
   apply (subgoal_tac "n + (k + 2) = n + (1 + (1 + k))")
   apply (erule ssubst) back back
@@ -184,7 +184,7 @@
 lemma fib_Cassini_aux_int: "fib (int n + 2) * fib (int n) - 
     (fib (int n + 1))^2 = (-1)^(n + 1)"
   apply (induct n)
-  apply (auto simp add: ring_simps power2_eq_square fib_reduce_int
+  apply (auto simp add: field_simps power2_eq_square fib_reduce_int
       power_add)
 done
 
@@ -222,7 +222,7 @@
   apply (subst (2) fib_reduce_nat)
   apply (auto simp add: Suc_eq_plus1) (* again, natdiff_cancel *)
   apply (subst add_commute, auto)
-  apply (subst gcd_commute_nat, auto simp add: ring_simps)
+  apply (subst gcd_commute_nat, auto simp add: field_simps)
 done
 
 lemma coprime_fib_Suc_nat: "coprime (fib n) (fib (Suc n))"
@@ -305,7 +305,7 @@
 theorem fib_mult_eq_setsum_nat:
     "fib ((n::nat) + 1) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
   apply (induct n)
-  apply (auto simp add: atMost_plus_one_nat fib_plus_2_nat ring_simps)
+  apply (auto simp add: atMost_plus_one_nat fib_plus_2_nat field_simps)
 done
 
 theorem fib_mult_eq_setsum'_nat:
--- a/src/HOL/Number_Theory/Residues.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/Number_Theory/Residues.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -69,7 +69,7 @@
   apply (subst mod_add_eq [symmetric])
   apply (subst mult_commute)
   apply (subst zmod_zmult1_eq [symmetric])
-  apply (simp add: ring_simps)
+  apply (simp add: field_simps)
 done
 
 end
--- a/src/HOL/SMT/Tools/z3_proof_rules.ML	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/SMT/Tools/z3_proof_rules.ML	Mon Apr 26 11:34:19 2010 +0200
@@ -1137,7 +1137,8 @@
   handle THM _ => NONE
 in
 val z3_simpset = HOL_ss addsimps @{thms array_rules}
-  addsimps @{thms ring_distribs} addsimps @{thms field_eq_simps}
+  addsimps @{thms ring_distribs} addsimps @{thms field_simps}
+  addsimps [@{thm times_divide_eq_right}, @{thm times_divide_eq_left}]
   addsimps @{thms arith_special} addsimps @{thms less_bin_simps}
   addsimps @{thms le_bin_simps} addsimps @{thms eq_bin_simps}
   addsimps @{thms add_bin_simps} addsimps @{thms succ_bin_simps}
--- a/src/HOL/SMT/Z3.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/SMT/Z3.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -19,7 +19,7 @@
 
 lemmas [z3_rewrite] =
   refl eq_commute conj_commute disj_commute simp_thms nnf_simps
-  ring_distribs field_eq_simps if_True if_False
+  ring_distribs field_simps times_divide_eq_right times_divide_eq_left if_True if_False
 
 lemma [z3_rewrite]: "(P \<noteq> Q) = (Q = (\<not>P))" by fast
 
--- a/src/HOL/SetInterval.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/SetInterval.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -1095,7 +1095,7 @@
   next
     case (Suc n)
     moreover with `y \<noteq> 0` have "(1 + y) ^ n = (y * inverse y) * (1 + y) ^ n" by simp 
-    ultimately show ?case by (simp add: field_eq_simps divide_inverse)
+    ultimately show ?case by (simp add: field_simps divide_inverse)
   qed
   ultimately show ?thesis by simp
 qed
--- a/src/HOL/ex/Lagrange.thy	Mon Apr 26 11:34:17 2010 +0200
+++ b/src/HOL/ex/Lagrange.thy	Mon Apr 26 11:34:19 2010 +0200
@@ -34,7 +34,7 @@
    sq (x1*y2 + x2*y1 + x3*y4 - x4*y3)  +
    sq (x1*y3 - x2*y4 + x3*y1 + x4*y2)  +
    sq (x1*y4 + x2*y3 - x3*y2 + x4*y1)"
-by (simp only: sq_def ring_simps)
+by (simp only: sq_def field_simps)
 
 
 text {* A challenge by John Harrison. Takes about 12s on a 1.6GHz machine. *}
@@ -50,6 +50,6 @@
       sq (p1*u2 + q1*t2 - r1*w2 + s1*v2 - t1*q2 + u1*p2 - v1*s2 + w1*r2) +
       sq (p1*v2 + q1*w2 + r1*t2 - s1*u2 - t1*r2 + u1*s2 + v1*p2 - w1*q2) +
       sq (p1*w2 - q1*v2 + r1*u2 + s1*t2 - t1*s2 - u1*r2 + v1*q2 + w1*p2)"
-by (simp only: sq_def ring_simps)
+by (simp only: sq_def field_simps)
 
 end