Merged
authoreberlm <eberlm@in.tum.de>
Tue, 20 Sep 2016 11:35:10 +0200
changeset 63922 d184a824aa63
parent 63921 0a5184877cb7 (current diff)
parent 63920 003622e08379 (diff)
child 63923 c9bba9dba73b
Merged
src/HOL/Library/Multiset.thy
src/HOL/Quotient_Examples/FSet.thy
--- a/NEWS	Mon Sep 19 17:37:22 2016 +0200
+++ b/NEWS	Tue Sep 20 11:35:10 2016 +0200
@@ -635,10 +635,20 @@
   msetsum ~> sum_mset
   msetprod ~> prod_mset
 
+* The symbols for intersection and union of multisets have been changed:
+  #\<inter> ~> \<inter>#
+  #\<union> ~> \<union>#
+INCOMPATIBILITY.
+
 * The lemma one_step_implies_mult_aux on multisets has been removed, use
 one_step_implies_mult instead.
 INCOMPATIBILITY.
 
+* The following theorems have been renamed:
+  setsum_left_distrib ~> setsum_distrib_right
+  setsum_right_distrib ~> setsum_distrib_left
+INCOMPATIBILITY.
+
 * Compound constants INFIMUM and SUPREMUM are mere abbreviations now.
 INCOMPATIBILITY.
 
@@ -785,6 +795,9 @@
   nn_integral :: 'a measure => ('a => ennreal) => ennreal
 INCOMPATIBILITY.
 
+* Renamed HOL/Quotient_Examples/FSet.thy to 
+HOL/Quotient_Examples/Quotient_FSet.thy 
+INCOMPATIBILITY.
 
 *** ML ***
 
--- a/src/HOL/Algebra/Divisibility.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Algebra/Divisibility.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -2306,10 +2306,10 @@
     by (fast elim: wfactorsE)
 
   have "\<exists>c cs. c \<in> carrier G \<and> set cs \<subseteq> carrier G \<and> wfactors G cs c \<and>
-    fmset G cs = fmset G as #\<inter> fmset G bs"
+    fmset G cs = fmset G as \<inter># fmset G bs"
   proof (intro mset_wfactorsEx)
     fix X
-    assume "X \<in># fmset G as #\<inter> fmset G bs"
+    assume "X \<in># fmset G as \<inter># fmset G bs"
     then have "X \<in># fmset G as" by simp
     then have "X \<in> set (map (assocs G) as)"
       by (simp add: fmset_def)
@@ -2328,7 +2328,7 @@
     where ccarr: "c \<in> carrier G"
       and cscarr: "set cs \<subseteq> carrier G"
       and csirr: "wfactors G cs c"
-      and csmset: "fmset G cs = fmset G as #\<inter> fmset G bs"
+      and csmset: "fmset G cs = fmset G as \<inter># fmset G bs"
     by auto
 
   have "c gcdof a b"
--- a/src/HOL/Analysis/Bounded_Linear_Function.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Bounded_Linear_Function.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -352,7 +352,7 @@
   apply (rule norm_blinfun_bound)
    apply (simp add: setsum_nonneg)
   apply (subst euclidean_representation[symmetric, where 'a='a])
-  apply (simp only: blinfun.bilinear_simps setsum_left_distrib)
+  apply (simp only: blinfun.bilinear_simps setsum_distrib_right)
   apply (rule order.trans[OF norm_setsum setsum_mono])
   apply (simp add: abs_mult mult_right_mono ac_simps Basis_le_norm)
   done
@@ -406,7 +406,7 @@
   "norm (blinfun_of_matrix a) \<le> (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>a i j\<bar>)"
   apply (rule norm_blinfun_bound)
    apply (simp add: setsum_nonneg)
-  apply (simp only: blinfun_of_matrix_apply setsum_left_distrib)
+  apply (simp only: blinfun_of_matrix_apply setsum_distrib_right)
   apply (rule order_trans[OF norm_setsum setsum_mono])
   apply (rule order_trans[OF norm_setsum setsum_mono])
   apply (simp add: abs_mult mult_right_mono ac_simps Basis_le_norm)
--- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -119,8 +119,8 @@
   val ss1 =
     simpset_of (put_simpset HOL_basic_ss @{context}
       addsimps [@{thm setsum.distrib} RS sym,
-      @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib},
-      @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym])
+      @{thm setsum_subtractf} RS sym, @{thm setsum_distrib_left},
+      @{thm setsum_distrib_right}, @{thm setsum_negf} RS sym])
   val ss2 =
     simpset_of (@{context} addsimps
              [@{thm plus_vec_def}, @{thm times_vec_def},
@@ -326,7 +326,7 @@
 lemma setsum_cmul:
   fixes f:: "'c \<Rightarrow> ('a::semiring_1)^'n"
   shows "setsum (\<lambda>x. c *s f x) S = c *s setsum f S"
-  by (simp add: vec_eq_iff setsum_right_distrib)
+  by (simp add: vec_eq_iff setsum_distrib_left)
 
 lemma setsum_norm_allsubsets_bound_cart:
   fixes f:: "'a \<Rightarrow> real ^'n"
@@ -517,14 +517,14 @@
   done
 
 lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"
-  apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult.assoc)
+  apply (vector matrix_matrix_mult_def setsum_distrib_left setsum_distrib_right mult.assoc)
   apply (subst setsum.commute)
   apply simp
   done
 
 lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x"
   apply (vector matrix_matrix_mult_def matrix_vector_mult_def
-    setsum_right_distrib setsum_left_distrib mult.assoc)
+    setsum_distrib_left setsum_distrib_right mult.assoc)
   apply (subst setsum.commute)
   apply simp
   done
@@ -555,7 +555,7 @@
   by (simp add: matrix_vector_mult_def inner_vec_def)
 
 lemma dot_lmul_matrix: "((x::real ^_) v* A) \<bullet> y = x \<bullet> (A *v y)"
-  apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib ac_simps)
+  apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def setsum_distrib_right setsum_distrib_left ac_simps)
   apply (subst setsum.commute)
   apply simp
   done
@@ -630,7 +630,7 @@
 
 lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::real ^ _))"
   by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff
-      field_simps setsum_right_distrib setsum.distrib)
+      field_simps setsum_distrib_left setsum.distrib)
 
 lemma matrix_works:
   assumes lf: "linear f"
@@ -660,7 +660,7 @@
 lemma adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)"
   apply (rule adjoint_unique)
   apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def
-    setsum_left_distrib setsum_right_distrib)
+    setsum_distrib_right setsum_distrib_left)
   apply (subst setsum.commute)
   apply (auto simp add: ac_simps)
   done
--- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -6215,7 +6215,7 @@
         by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w)
       have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) * (u - w) - f u)
             = norm ((\<Sum>k<N. (((w - z) / (u - z)) ^ k)) * f u * (u - w) / (u - z) - f u)"
-        unfolding setsum_left_distrib setsum_divide_distrib power_divide by (simp add: algebra_simps)
+        unfolding setsum_distrib_right setsum_divide_distrib power_divide by (simp add: algebra_simps)
       also have "... = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)"
         using \<open>0 < B\<close>
         apply (auto simp: geometric_sum [OF wzu_not1])
--- a/src/HOL/Analysis/Complex_Analysis_Basics.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -951,7 +951,7 @@
       then have "cmod h * cmod ((\<Sum>i<n. f' i y) - g' y) \<le> cmod h * e"
         by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2)
       then show "cmod ((\<Sum>i<n. h * f' i y) - g' y * h) \<le> e * cmod h"
-        by (simp add: norm_mult [symmetric] field_simps setsum_right_distrib)
+        by (simp add: norm_mult [symmetric] field_simps setsum_distrib_left)
     qed
   } note ** = this
   show ?thesis
--- a/src/HOL/Analysis/Complex_Transcendental.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Complex_Transcendental.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -597,7 +597,7 @@
 text\<open>32-bit Approximation to e\<close>
 lemma e_approx_32: "\<bar>exp(1) - 5837465777 / 2147483648\<bar> \<le> (inverse(2 ^ 32)::real)"
   using Taylor_exp [of 1 14] exp_le
-  apply (simp add: setsum_left_distrib in_Reals_norm Re_exp atMost_nat_numeral fact_numeral)
+  apply (simp add: setsum_distrib_right in_Reals_norm Re_exp atMost_nat_numeral fact_numeral)
   apply (simp only: pos_le_divide_eq [symmetric], linarith)
   done
 
--- a/src/HOL/Analysis/Conformal_Mappings.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Conformal_Mappings.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -2846,7 +2846,7 @@
         qed
     qed
   also have "... = c * (\<Sum>p\<in>pts. winding_number g p * residue f p)"
-    by (simp add: setsum_right_distrib algebra_simps)
+    by (simp add: setsum_distrib_left algebra_simps)
   finally show ?thesis unfolding c_def .
 qed
 
@@ -3459,7 +3459,7 @@
             qed
         qed
       then have "(\<Sum>p\<in>zeros. w p * cont p) = c * (\<Sum>p\<in>zeros.  w p *  h p * zorder f p)"
-        apply (subst setsum_right_distrib)
+        apply (subst setsum_distrib_left)
         by (simp add:algebra_simps)
       moreover have "(\<Sum>p\<in>poles. w p * cont p) = (\<Sum>p\<in>poles.  - c * w p *  h p * porder f p)"
         proof (rule setsum.cong[of poles poles,simplified])
@@ -3479,7 +3479,7 @@
             qed
         qed
       then have "(\<Sum>p\<in>poles. w p * cont p) = - c * (\<Sum>p\<in>poles. w p *  h p * porder f p)"
-        apply (subst setsum_right_distrib)
+        apply (subst setsum_distrib_left)
         by (simp add:algebra_simps)
       ultimately show ?thesis by (simp add: right_diff_distrib)
     qed
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -523,7 +523,7 @@
         qed auto
         then show ?thesis
           apply (rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
-          unfolding setsum_right_distrib[symmetric]
+          unfolding setsum_distrib_left[symmetric]
           using as and *** and True
           apply auto
           done
@@ -536,7 +536,7 @@
         then show ?thesis
           using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]]
           using *** *(2) and \<open>s \<subseteq> V\<close>
-          unfolding setsum_right_distrib
+          unfolding setsum_distrib_left
           by (auto simp add: setsum_clauses(2))
       qed
       then have "u x + (1 - u x) = 1 \<Longrightarrow>
@@ -619,7 +619,7 @@
       unfolding scaleR_left_distrib setsum.distrib if_smult scaleR_zero_left
         ** setsum.inter_restrict[OF xy, symmetric]
       unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.setsum [symmetric]
-        and setsum_right_distrib[symmetric]
+        and setsum_distrib_left[symmetric]
       unfolding x y
       using x(1-3) y(1-3) uv
       apply simp
@@ -1323,7 +1323,7 @@
     apply (rule_tac x="s - {v}" in exI)
     apply (rule_tac x="\<lambda>x. - (1 / u v) * u x" in exI)
     unfolding scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric]
-    unfolding setsum_right_distrib[symmetric] and setsum_diff1[OF as(1)]
+    unfolding setsum_distrib_left[symmetric] and setsum_diff1[OF as(1)]
     using as
     apply auto
     done
@@ -1793,7 +1793,7 @@
     apply rule
     unfolding * and setsum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and
       setsum.reindex[OF inj] and o_def Collect_mem_eq
-    unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] setsum_right_distrib[symmetric]
+    unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] setsum_distrib_left[symmetric]
   proof -
     fix i
     assume i: "i \<in> {1..k1+k2}"
@@ -1844,7 +1844,7 @@
   }
   moreover
   have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1"
-    unfolding setsum.distrib and setsum_right_distrib[symmetric] and ux(2) uy(2)
+    unfolding setsum.distrib and setsum_distrib_left[symmetric] and ux(2) uy(2)
     using uv(3) by auto
   moreover
   have "(\<Sum>x\<in>s. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
@@ -3306,7 +3306,7 @@
     have *: "\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'b::ab_group_add)"
       unfolding setsum.remove[OF obt(1) \<open>a\<in>s\<close>] by auto
     have "(\<Sum>v\<in>s. u v + t * w v) = 1"
-      unfolding setsum.distrib wv(1) setsum_right_distrib[symmetric] obt(5) by auto
+      unfolding setsum.distrib wv(1) setsum_distrib_left[symmetric] obt(5) by auto
     moreover have "(\<Sum>v\<in>s. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y"
       unfolding setsum.distrib obt(6) scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] wv(4)
       using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp
@@ -5279,7 +5279,7 @@
     apply (rule_tac x="{v \<in> c. u v < 0}" in exI)
     apply (rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * - u y" in exI)
     using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def
-    apply (auto simp add: setsum_negf setsum_right_distrib[symmetric])
+    apply (auto simp add: setsum_negf setsum_distrib_left[symmetric])
     done
   moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * u x"
     apply rule
@@ -5294,7 +5294,7 @@
     using assms(1)
     unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def
     using *
-    apply (auto simp add: setsum_negf setsum_right_distrib[symmetric])
+    apply (auto simp add: setsum_negf setsum_distrib_left[symmetric])
     done
   ultimately show ?thesis
     apply (rule_tac x="{v\<in>c. u v \<le> 0}" in exI)
@@ -5683,7 +5683,7 @@
     unfolding convex_hull_indexed mem_Collect_eq by auto
   have "(\<Sum>i = 1..k. u i * f (v i)) \<le> b"
     using setsum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"]
-    unfolding setsum_left_distrib[symmetric] obt(2) mult_1
+    unfolding setsum_distrib_right[symmetric] obt(2) mult_1
     apply (drule_tac meta_mp)
     apply (rule mult_left_mono)
     using assms(2) obt(1)
@@ -9200,9 +9200,9 @@
     have ge0: "\<forall>i\<in>I. e i \<ge> 0"
       using e_def xc yc uv by simp
     have "setsum (\<lambda>i. u * c i) I = u * setsum c I"
-      by (simp add: setsum_right_distrib)
+      by (simp add: setsum_distrib_left)
     moreover have "setsum (\<lambda>i. v * d i) I = v * setsum d I"
-      by (simp add: setsum_right_distrib)
+      by (simp add: setsum_distrib_left)
     ultimately have sum1: "setsum e I = 1"
       using e_def xc yc uv by (simp add: setsum.distrib)
     define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)"
@@ -11857,7 +11857,7 @@
   have sum_dd0: "setsum dd S = 0"
     unfolding dd_def  using S
     by (simp add: sumSS' comm_monoid_add_class.setsum.distrib setsum_subtractf
-                  algebra_simps setsum_left_distrib [symmetric] b1)
+                  algebra_simps setsum_distrib_right [symmetric] b1)
   have "(\<Sum>v\<in>S. (b v * x) *\<^sub>R v) = x *\<^sub>R (\<Sum>v\<in>S. b v *\<^sub>R v)" for x
     by (simp add: pth_5 real_vector.scale_setsum_right mult.commute)
   then have *: "(\<Sum>v\<in>S. (b v * x) *\<^sub>R v) = x *\<^sub>R a" for x
--- a/src/HOL/Analysis/Derivative.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Derivative.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -2244,7 +2244,7 @@
     {
       fix n :: nat and x h :: 'a assume nx: "n \<ge> N" "x \<in> s"
       have "norm ((\<Sum>i<n. f' i x * h) - g' x * h) = norm ((\<Sum>i<n. f' i x) - g' x) * norm h"
-        by (simp add: norm_mult [symmetric] ring_distribs setsum_left_distrib)
+        by (simp add: norm_mult [symmetric] ring_distribs setsum_distrib_right)
       also from N[OF nx] have "norm ((\<Sum>i<n. f' i x) - g' x) \<le> e" by simp
       hence "norm ((\<Sum>i<n. f' i x) - g' x) * norm h \<le> e * norm h"
         by (intro mult_right_mono) simp_all
--- a/src/HOL/Analysis/Determinants.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Determinants.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -226,7 +226,7 @@
   fixes A :: "'a::comm_ring_1^'n^'n"
   assumes p: "p permutes (UNIV :: 'n::finite set)"
   shows "det (\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A"
-  apply (simp add: det_def setsum_right_distrib mult.assoc[symmetric])
+  apply (simp add: det_def setsum_distrib_left mult.assoc[symmetric])
   apply (subst sum_permutations_compose_right[OF p])
 proof (rule setsum.cong)
   let ?U = "UNIV :: 'n set"
@@ -372,7 +372,7 @@
   fixes a b :: "'n::finite \<Rightarrow> _ ^ 'n"
   shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) =
     c * det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)"
-  unfolding det_def vec_lambda_beta setsum_right_distrib
+  unfolding det_def vec_lambda_beta setsum_distrib_left
 proof (rule setsum.cong)
   let ?U = "UNIV :: 'n set"
   let ?pU = "{p. p permutes ?U}"
@@ -645,7 +645,7 @@
 lemma det_rows_mul:
   "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n) =
     setprod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)"
-proof (simp add: det_def setsum_right_distrib cong add: setprod.cong, rule setsum.cong)
+proof (simp add: det_def setsum_distrib_left cong add: setprod.cong, rule setsum.cong)
   let ?U = "UNIV :: 'n set"
   let ?PU = "{p. p permutes ?U}"
   fix p
--- a/src/HOL/Analysis/Finite_Cartesian_Product.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -461,7 +461,7 @@
     by (simp add: inner_add_left setsum.distrib)
   show "inner (scaleR r x) y = r * inner x y"
     unfolding inner_vec_def
-    by (simp add: setsum_right_distrib)
+    by (simp add: setsum_distrib_left)
   show "0 \<le> inner x x"
     unfolding inner_vec_def
     by (simp add: setsum_nonneg)
--- a/src/HOL/Analysis/Gamma_Function.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Gamma_Function.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -450,7 +450,7 @@
       (\<Sum>k=Suc m..n. 2 * (norm t + (norm t)\<^sup>2) / (real_of_nat k)\<^sup>2)" using t_nz N(2) mn norm_t
       by (intro order.trans[OF norm_setsum setsum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all
     also have "... \<le> 2 * (norm t + norm t^2) * (\<Sum>k=Suc m..n. 1 / (of_nat k)\<^sup>2)"
-      by (simp add: setsum_right_distrib)
+      by (simp add: setsum_distrib_left)
     also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz
       by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all
     also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')"
@@ -543,7 +543,7 @@
       by (subst atLeast0LessThan [symmetric], subst setsum_shift_bounds_Suc_ivl [symmetric],
           subst atLeastLessThanSuc_atLeastAtMost) simp_all
     also have "\<dots> = z * of_real (harm n) - (\<Sum>k=1..n. ln (1 + z / of_nat k))"
-      by (simp add: harm_def setsum_subtractf setsum_right_distrib divide_inverse)
+      by (simp add: harm_def setsum_subtractf setsum_distrib_left divide_inverse)
     also from n have "\<dots> - ?g n = 0"
       by (simp add: ln_Gamma_series_def setsum_subtractf algebra_simps Ln_of_nat)
     finally show "(\<Sum>k<n. ?f k) - ?g n = 0" .
@@ -944,7 +944,7 @@
       using z' n by (intro uniformly_convergent_mult Polygamma_converges) (simp_all add: n'_def)
     thus "uniformly_convergent_on (ball z d)
               (\<lambda>k z. \<Sum>i<k. - of_nat n' * inverse ((z + of_nat i :: 'a) ^ (n'+1)))"
-      by (subst (asm) setsum_right_distrib) simp
+      by (subst (asm) setsum_distrib_left) simp
   qed (insert Polygamma_converges'[OF z' n'] d, simp_all)
   also have "(\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) =
                (- of_nat n') * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))"
@@ -2573,7 +2573,7 @@
 proof -
   have "(\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) =
           exp (z * of_real (\<Sum>k = 1..n. ln (1 + 1 / real_of_nat k)))"
-    by (subst exp_setsum [symmetric]) (simp_all add: setsum_right_distrib)
+    by (subst exp_setsum [symmetric]) (simp_all add: setsum_distrib_left)
   also have "(\<Sum>k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\<Prod>k=1..n. 1 + 1 / real_of_nat k)"
     by (subst ln_setprod [symmetric]) (auto intro!: add_pos_nonneg)
   also have "(\<Prod>k=1..n. 1 + 1 / of_nat k :: real) = (\<Prod>k=1..n. (of_nat k + 1) / of_nat k)"
--- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -3756,7 +3756,7 @@
   have "norm (setsum (\<lambda>l. content l *\<^sub>R c) p) \<le> (\<Sum>i\<in>p. norm (content i *\<^sub>R c))"
     using norm_setsum by blast
   also have "...  \<le> e * (\<Sum>i\<in>p. \<bar>content i\<bar>)"
-    by (simp add: setsum_right_distrib[symmetric] mult.commute assms(2) mult_right_mono setsum_nonneg)
+    by (simp add: setsum_distrib_left[symmetric] mult.commute assms(2) mult_right_mono setsum_nonneg)
   also have "... \<le> e * content (cbox a b)"
     apply (rule mult_left_mono [OF _ e])
     apply (simp add: sumeq)
@@ -3792,7 +3792,7 @@
     apply (rule order_trans[OF setsum_mono])
     apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e])
     apply (metis norm)
-    unfolding setsum_left_distrib[symmetric]
+    unfolding setsum_distrib_right[symmetric]
     using con setsum_le
     apply (auto simp: mult.commute intro: mult_left_mono [OF _ e])
     done
@@ -4697,7 +4697,7 @@
         done
       have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) \<le> setsum (\<lambda>i. (real i + 1) *
         norm (setsum (\<lambda>(x,k). content k *\<^sub>R indicator s x :: real) (q i))) {..N+1}"
-        unfolding real_norm_def setsum_right_distrib abs_of_nonneg[OF *] diff_0_right
+        unfolding real_norm_def setsum_distrib_left abs_of_nonneg[OF *] diff_0_right
         apply (rule order_trans)
         apply (rule norm_setsum)
         apply (subst sum_sum_product)
@@ -4775,7 +4775,7 @@
           done
       qed
       also have "\<dots> < e * inverse 2 * 2"
-        unfolding divide_inverse setsum_right_distrib[symmetric]
+        unfolding divide_inverse setsum_distrib_left[symmetric]
         apply (rule mult_strict_left_mono)
         unfolding power_inverse [symmetric] lessThan_Suc_atMost[symmetric]
         apply (subst geometric_sum)
@@ -5313,7 +5313,7 @@
     show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content (cbox a b)"
       unfolding content_real[OF assms(1), simplified box_real[symmetric]] additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of f,symmetric]
       unfolding additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of "\<lambda>x. x",symmetric]
-      unfolding setsum_right_distrib
+      unfolding setsum_distrib_left
       defer
       unfolding setsum_subtractf[symmetric]
     proof (rule setsum_norm_le,safe)
@@ -6479,7 +6479,7 @@
       by arith
     show ?case
       unfolding content_real[OF assms(1)] and *[of "\<lambda>x. x"] *[of f] setsum_subtractf[symmetric] split_minus
-      unfolding setsum_right_distrib
+      unfolding setsum_distrib_left
       apply (subst(2) pA)
       apply (subst pA)
       unfolding setsum.union_disjoint[OF pA(2-)]
@@ -6548,7 +6548,7 @@
         apply (unfold split_paired_all split_conv)
         defer
         unfolding setsum.union_disjoint[OF pA(2-),symmetric] pA(1)[symmetric]
-        unfolding setsum_right_distrib[symmetric]
+        unfolding setsum_distrib_left[symmetric]
         apply (subst additive_tagged_division_1[OF _ as(1)])
         apply (rule assms)
       proof -
@@ -8856,7 +8856,7 @@
       apply rule
       apply (drule qq)
       defer
-      unfolding divide_inverse setsum_left_distrib[symmetric]
+      unfolding divide_inverse setsum_distrib_right[symmetric]
       unfolding divide_inverse[symmetric]
       using * apply (auto simp add: field_simps)
       done
@@ -8976,7 +8976,7 @@
     done
   have th: "op ^ x \<circ> op + m = (\<lambda>i. x^m * x^i)"
     by (rule ext) (simp add: power_add power_mult)
-  from setsum.reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]
+  from setsum.reindex[OF i, of "op ^ x", unfolded f th setsum_distrib_left[symmetric]]
   have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})"
     by simp
   then show ?thesis
@@ -9164,7 +9164,7 @@
           apply (rule norm_setsum)
           apply (rule setsum_mono)
           unfolding split_paired_all split_conv
-          unfolding split_def setsum_left_distrib[symmetric] scaleR_diff_right[symmetric]
+          unfolding split_def setsum_distrib_right[symmetric] scaleR_diff_right[symmetric]
           unfolding additive_content_tagged_division[OF p(1), unfolded split_def]
         proof -
           fix x k
@@ -9201,7 +9201,7 @@
           proof
             show "(\<Sum>i = 0..s. e / 2 ^ (i + 2)) < e / 2"
               unfolding power_add divide_inverse inverse_mult_distrib
-              unfolding setsum_right_distrib[symmetric] setsum_left_distrib[symmetric]
+              unfolding setsum_distrib_left[symmetric] setsum_distrib_right[symmetric]
               unfolding power_inverse [symmetric] sum_gp
               apply(rule mult_strict_left_mono[OF _ e])
               unfolding power2_eq_square
@@ -10350,7 +10350,7 @@
               by auto
           qed
           finally show "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar> * norm (f x)) = content l *\<^sub>R norm (f x)"
-            unfolding setsum_left_distrib[symmetric] real_scaleR_def
+            unfolding setsum_distrib_right[symmetric] real_scaleR_def
             apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]])
             using xl(2)[unfolded uv]
             unfolding uv
@@ -10614,7 +10614,7 @@
   proof goal_cases
     case prems: (2 d)
     have "(\<Sum>k\<in>d. norm (integral k (h \<circ> f))) \<le> setsum (\<lambda>k. norm(integral k f)) d * b"
-      unfolding setsum_left_distrib
+      unfolding setsum_distrib_right
       apply (rule setsum_mono)
     proof goal_cases
       case (1 k)
--- a/src/HOL/Analysis/Homeomorphism.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Homeomorphism.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -1007,7 +1007,7 @@
   have gf[simp]: "g (f x) = x" for x
     apply (rule euclidean_eqI)
     apply (simp add: f_def g_def inner_setsum_left scaleR_setsum_left algebra_simps)
-    apply (simp add: Groups_Big.setsum_right_distrib [symmetric] *)
+    apply (simp add: Groups_Big.setsum_distrib_left [symmetric] *)
     done
   then have "inj f" by (metis injI)
   have gfU: "g ` f ` U = U"
--- a/src/HOL/Analysis/L2_Norm.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/L2_Norm.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -58,7 +58,7 @@
   "0 \<le> r \<Longrightarrow> r * setL2 f A = setL2 (\<lambda>x. r * f x) A"
   unfolding setL2_def
   apply (simp add: power_mult_distrib)
-  apply (simp add: setsum_right_distrib [symmetric])
+  apply (simp add: setsum_distrib_left [symmetric])
   apply (simp add: real_sqrt_mult setsum_nonneg)
   done
 
@@ -66,7 +66,7 @@
   "0 \<le> r \<Longrightarrow> setL2 f A * r = setL2 (\<lambda>x. f x * r) A"
   unfolding setL2_def
   apply (simp add: power_mult_distrib)
-  apply (simp add: setsum_left_distrib [symmetric])
+  apply (simp add: setsum_distrib_right [symmetric])
   apply (simp add: real_sqrt_mult setsum_nonneg)
   done
 
--- a/src/HOL/Analysis/Lebesgue_Measure.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Lebesgue_Measure.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -235,7 +235,7 @@
       by auto
     also have "... = (\<Sum>i \<in> S. F(r i) - F(l i)) +
         (epsilon / 4) * (\<Sum>i \<in> S. (1 / 2)^i)" (is "_ = ?t + _")
-      by (subst setsum.distrib) (simp add: field_simps setsum_right_distrib)
+      by (subst setsum.distrib) (simp add: field_simps setsum_distrib_left)
     also have "... \<le> ?t + (epsilon / 4) * (\<Sum> i < Suc n. (1 / 2)^i)"
       apply (rule add_left_mono)
       apply (rule mult_left_mono)
--- a/src/HOL/Analysis/Linear_Algebra.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Linear_Algebra.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -1960,7 +1960,7 @@
     qed
     from setsum_norm_le[of _ ?g, OF th]
     show "norm (f x) \<le> ?B * norm x"
-      unfolding th0 setsum_left_distrib by metis
+      unfolding th0 setsum_distrib_right by metis
   qed
 qed
 
@@ -2021,7 +2021,7 @@
     unfolding bilinear_setsum[OF bh finite_Basis finite_Basis] ..
   finally have th: "norm (h x y) = \<dots>" .
   show "norm (h x y) \<le> (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. norm (h i j)) * norm x * norm y"
-    apply (auto simp add: setsum_left_distrib th setsum.cartesian_product)
+    apply (auto simp add: setsum_distrib_right th setsum.cartesian_product)
     apply (rule setsum_norm_le)
     apply simp
     apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh]
--- a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -550,7 +550,7 @@
   qed
   also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M.
       if \<exists>x\<in>space M. y = f x \<and> z = g x then y * emeasure M {x\<in>space M. g x = z} else 0))"
-    by (auto intro!: setsum.cong simp: setsum_right_distrib)
+    by (auto intro!: setsum.cong simp: setsum_distrib_left)
   also have "\<dots> = ?r"
     by (subst setsum.commute)
        (auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_right[symmetric] eq)
@@ -592,7 +592,7 @@
     using f by (intro simple_function_partition) auto
   also have "\<dots> = c * integral\<^sup>S M f"
     using f unfolding simple_integral_def
-    by (subst setsum_right_distrib) (auto simp: mult.assoc Int_def conj_commute)
+    by (subst setsum_distrib_left) (auto simp: mult.assoc Int_def conj_commute)
   finally show ?thesis .
 qed
 
--- a/src/HOL/Analysis/Poly_Roots.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Poly_Roots.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -24,7 +24,7 @@
 lemma setsum_power_add:
   fixes x :: "'a::{comm_ring,monoid_mult}"
   shows "(\<Sum>i\<in>I. x^(m+i)) = x^m * (\<Sum>i\<in>I. x^i)"
-  by (simp add: setsum_right_distrib power_add)
+  by (simp add: setsum_distrib_left power_add)
 
 lemma setsum_power_shift:
   fixes x :: "'a::{comm_ring,monoid_mult}"
@@ -32,7 +32,7 @@
   shows "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i\<le>n-m. x^i)"
 proof -
   have "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i=m..n. x^(i-m))"
-    by (simp add: setsum_right_distrib power_add [symmetric])
+    by (simp add: setsum_distrib_left power_add [symmetric])
   also have "(\<Sum>i=m..n. x^(i-m)) = (\<Sum>i\<le>n-m. x^i)"
     using \<open>m \<le> n\<close> by (intro setsum.reindex_bij_witness[where j="\<lambda>i. i - m" and i="\<lambda>i. i + m"]) auto
   finally show ?thesis .
@@ -88,7 +88,7 @@
   also have "... = (\<Sum>i\<le>n. a i * (x - y) * (\<Sum>j<i. y^(i - Suc j) * x^j))"
     by (simp add: power_diff_sumr2 ac_simps)
   also have "... = (x - y) * (\<Sum>i\<le>n. (\<Sum>j<i. a i * y^(i - Suc j) * x^j))"
-    by (simp add: setsum_right_distrib ac_simps)
+    by (simp add: setsum_distrib_left ac_simps)
   also have "... = (x - y) * (\<Sum>j<n. (\<Sum>i=Suc j..n. a i * y^(i - Suc j) * x^j))"
     by (simp add: nested_setsum_swap')
   finally show ?thesis .
@@ -115,7 +115,7 @@
   { fix z
     have "(\<Sum>i\<le>n. c i * z^i) - (\<Sum>i\<le>n. c i * a^i) =
           (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)"
-      by (simp add: sub_polyfun setsum_left_distrib)
+      by (simp add: sub_polyfun setsum_distrib_right)
     then have "(\<Sum>i\<le>n. c i * z^i) =
           (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)
           + (\<Sum>i\<le>n. c i * a^i)"
--- a/src/HOL/Analysis/Polytope.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Polytope.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -303,7 +303,7 @@
     have cge0: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> c i"
       using a b u01 by (simp add: c_def)
     have sumc1: "setsum c S = 1"
-      by (simp add: c_def setsum.distrib setsum_right_distrib [symmetric] asum bsum)
+      by (simp add: c_def setsum.distrib setsum_distrib_left [symmetric] asum bsum)
     have sumci_xy: "(\<Sum>i\<in>S. c i *\<^sub>R i) = (1 - u) *\<^sub>R x + u *\<^sub>R y"
       apply (simp add: c_def setsum.distrib scaleR_left_distrib)
       by (simp only: scaleR_scaleR [symmetric] Real_Vector_Spaces.scaleR_right.setsum [symmetric] aeqx beqy)
@@ -357,7 +357,7 @@
           apply (rule_tac x="\<lambda>i. inverse (1-k) * c i" in exI)
           apply auto
           apply (metis sumcf cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg S(2) setsum_nonneg subsetCE)
-          apply (metis False mult.commute right_inverse right_minus_eq setsum_right_distrib sumcf)
+          apply (metis False mult.commute right_inverse right_minus_eq setsum_distrib_left sumcf)
           by (metis (mono_tags, lifting) scaleR_right.setsum scaleR_scaleR setsum.cong)
         with \<open>0 < k\<close>  have "inverse(k) *\<^sub>R (w - setsum (\<lambda>i. c i *\<^sub>R i) T) \<in> affine hull T"
           by (simp add: affine_diff_divide [OF affine_affine_hull] False waff convex_hull_subset_affine_hull [THEN subsetD])
@@ -365,7 +365,7 @@
           apply (simp add: weq_sumsum convex_hull_finite fin)
           apply (rule_tac x="\<lambda>i. inverse k * c i" in exI)
           using \<open>k > 0\<close> cge0
-          apply (auto simp: scaleR_right.setsum setsum_right_distrib [symmetric] k_def [symmetric])
+          apply (auto simp: scaleR_right.setsum setsum_distrib_left [symmetric] k_def [symmetric])
           done
         ultimately show ?thesis
           using disj by blast
--- a/src/HOL/Analysis/Summation_Tests.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Summation_Tests.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -288,7 +288,7 @@
     qed
     from this and A have "Bseq (\<lambda>n. \<Sum>k<n. 2^k * f (2^Suc k))" by (rule Bseq_eventually_mono)
     from Bseq_mult[OF Bfun_const[of 2] this] have "Bseq (\<lambda>n. \<Sum>k<n. 2^Suc k * f (2^Suc k))"
-      by (simp add: setsum_right_distrib setsum_left_distrib mult_ac)
+      by (simp add: setsum_distrib_left setsum_distrib_right mult_ac)
     hence "Bseq (\<lambda>n. (\<Sum>k=Suc 0..<Suc n. 2^k * f (2^k)) + f 1)"
       by (intro Bseq_add, subst setsum_shift_bounds_Suc_ivl) (simp add: atLeast0LessThan)
     hence "Bseq (\<lambda>n. (\<Sum>k=0..<Suc n. 2^k * f (2^k)))"
@@ -424,7 +424,7 @@
     have n: "n > m" by (simp add: n_def)
 
     from r have "r * norm (\<Sum>k\<le>n. f k) = norm (\<Sum>k\<le>n. r * f k)"
-      by (simp add: setsum_right_distrib[symmetric] abs_mult)
+      by (simp add: setsum_distrib_left[symmetric] abs_mult)
     also from n have "{..n} = {..m} \<union> {Suc m..n}" by auto
     hence "(\<Sum>k\<le>n. r * f k) = (\<Sum>k\<in>{..m} \<union> {Suc m..n}. r * f k)" by (simp only:)
     also have "\<dots> = (\<Sum>k\<le>m. r * f k) + (\<Sum>k=Suc m..n. r * f k)"
--- a/src/HOL/Analysis/Weierstrass_Theorems.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/Weierstrass_Theorems.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -38,7 +38,7 @@
 
 lemma sum_k_Bernstein [simp]: "(\<Sum>k = 0..n. real k * Bernstein n k x) = of_nat n * x"
   apply (subst binomial_deriv1 [of n x "1-x", simplified, symmetric])
-  apply (simp add: setsum_left_distrib)
+  apply (simp add: setsum_distrib_right)
   apply (auto simp: Bernstein_def algebra_simps realpow_num_eq_if intro!: setsum.cong)
   done
 
@@ -46,7 +46,7 @@
 proof -
   have "(\<Sum> k = 0..n. real k * (real k - 1) * Bernstein n k x) = real_of_nat n * real_of_nat (n - Suc 0) * x\<^sup>2"
     apply (subst binomial_deriv2 [of n x "1-x", simplified, symmetric])
-    apply (simp add: setsum_left_distrib)
+    apply (simp add: setsum_distrib_right)
     apply (rule setsum.cong [OF refl])
     apply (simp add: Bernstein_def power2_eq_square algebra_simps)
     apply (rename_tac k)
@@ -98,7 +98,7 @@
         by (simp add: algebra_simps power2_eq_square)
       have "(\<Sum> k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x) = n * x * (1 - x)"
         apply (simp add: * setsum.distrib)
-        apply (simp add: setsum_right_distrib [symmetric] mult.assoc)
+        apply (simp add: setsum_distrib_left [symmetric] mult.assoc)
         apply (simp add: algebra_simps power2_eq_square)
         done
       then have "(\<Sum> k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x)/n^2 = x * (1 - x) / n"
@@ -138,14 +138,14 @@
         qed
     } note * = this
     have "\<bar>f x - (\<Sum> k = 0..n. f(k / n) * Bernstein n k x)\<bar> \<le> \<bar>\<Sum> k = 0..n. (f x - f(k / n)) * Bernstein n k x\<bar>"
-      by (simp add: setsum_subtractf setsum_right_distrib [symmetric] algebra_simps)
+      by (simp add: setsum_subtractf setsum_distrib_left [symmetric] algebra_simps)
     also have "... \<le> (\<Sum> k = 0..n. (e/2 + (2 * M / d\<^sup>2) * (x - k / n)\<^sup>2) * Bernstein n k x)"
       apply (rule order_trans [OF setsum_abs setsum_mono])
       using *
       apply (simp add: abs_mult Bernstein_nonneg x mult_right_mono)
       done
     also have "... \<le> e/2 + (2 * M) / (d\<^sup>2 * n)"
-      apply (simp only: setsum.distrib Rings.semiring_class.distrib_right setsum_right_distrib [symmetric] mult.assoc sum_bern)
+      apply (simp only: setsum.distrib Rings.semiring_class.distrib_right setsum_distrib_left [symmetric] mult.assoc sum_bern)
       using \<open>d>0\<close> x
       apply (simp add: divide_simps Mge0 mult_le_one mult_left_le)
       done
--- a/src/HOL/Analysis/ex/Approximations.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Analysis/ex/Approximations.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -30,7 +30,7 @@
       by (subst atLeast0LessThan [symmetric], subst setsum_head_upt_Suc) simp_all
     also have "(\<Sum>k=Suc 0..<Suc m. f k * x^k) = (\<Sum>k<m. f (k+1) * x^k) * x"
       by (subst setsum_shift_bounds_Suc_ivl)
-         (simp add: setsum_left_distrib algebra_simps atLeast0LessThan power_commutes)
+         (simp add: setsum_distrib_right algebra_simps atLeast0LessThan power_commutes)
     finally have "(\<Sum>k<Suc m. f k * x ^ k) = f 0 + (\<Sum>k<m. f (k + 1) * x ^ k) * x" .
   }
   from this[of "pred_numeral n"]
@@ -199,7 +199,7 @@
 lemma euler_approx_aux_Suc:
   "euler_approx_aux (Suc m) = 1 + Suc m * euler_approx_aux m"
   unfolding euler_approx_aux_def
-  by (subst setsum_right_distrib) (simp add: atLeastAtMostSuc_conv)
+  by (subst setsum_distrib_left) (simp add: atLeastAtMostSuc_conv)
 
 lemma eval_euler_approx_aux:
   "euler_approx_aux 0 = 1"
@@ -209,7 +209,7 @@
 proof -
   have A: "euler_approx_aux (Suc m) = 1 + Suc m * euler_approx_aux m" for m :: nat
     unfolding euler_approx_aux_def
-    by (subst setsum_right_distrib) (simp add: atLeastAtMostSuc_conv)
+    by (subst setsum_distrib_left) (simp add: atLeastAtMostSuc_conv)
   show ?th by (subst numeral_eq_Suc, subst A, subst numeral_eq_Suc [symmetric]) simp
 qed (simp_all add: euler_approx_aux_def)
 
@@ -281,7 +281,7 @@
                    y_def [symmetric] d_def [symmetric])
   also have "2 * y * (\<Sum>k<n. inverse (real (2 * k + 1)) * y\<^sup>2 ^ k) = 
                (\<Sum>k<n. 2 * y^(2*k+1) / (real (2 * k + 1)))"
-    by (subst setsum_right_distrib, simp, subst power_mult) 
+    by (subst setsum_distrib_left, simp, subst power_mult) 
        (simp_all add: divide_simps mult_ac power_mult)
   finally show ?case by (simp only: d_def y_def approx_def) 
 qed
@@ -380,7 +380,7 @@
   from sums_split_initial_segment[OF this, of n]
     have "(\<lambda>i. x * ((- x\<^sup>2) ^ (i + n) / real (2 * (i + n) + 1))) sums
             (arctan x - arctan_approx n x)"
-    by (simp add: arctan_approx_def setsum_right_distrib)
+    by (simp add: arctan_approx_def setsum_distrib_left)
   from sums_group[OF this, of 2] assms
     have sums: "(\<lambda>k. x * (x\<^sup>2)^n * (x^4)^k * c k) sums (arctan x - arctan_approx n x)"
     by (simp add: algebra_simps power_add power_mult [symmetric] c_def)
@@ -423,7 +423,7 @@
       by (subst atLeast0LessThan [symmetric], subst setsum_head_upt_Suc) simp_all
     also have "(\<Sum>k=Suc 0..<Suc m. inverse (f k * x^k)) = (\<Sum>k<m. inverse (f (k+1) * x^k)) / x"
       by (subst setsum_shift_bounds_Suc_ivl)
-         (simp add: setsum_right_distrib divide_inverse algebra_simps
+         (simp add: setsum_distrib_left divide_inverse algebra_simps
                     atLeast0LessThan power_commutes)
     finally have "(\<Sum>k<Suc m. inverse (f k) * inverse (x ^ k)) =
                       inverse (f 0) + (\<Sum>k<m. inverse (f (k + 1)) * inverse (x ^ k)) / x" by simp
--- a/src/HOL/Binomial.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Binomial.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -341,7 +341,7 @@
     by (rule distrib_right)
   also have "\<dots> = (\<Sum>k=0..n. of_nat (n choose k) * a^(k+1) * b^(n-k)) +
       (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n - k + 1))"
-    by (auto simp add: setsum_right_distrib ac_simps)
+    by (auto simp add: setsum_distrib_left ac_simps)
   also have "\<dots> = (\<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 field_simps del: setsum_cl_ivl_Suc)
@@ -463,7 +463,7 @@
   also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) + (-1) ^ i * of_nat (n choose i))"
     by (simp add: setsum.distrib)
   also have "\<dots> = 2 * (\<Sum>i\<le>n. if even i then of_nat (n choose i) else 0)"
-    by (subst setsum_right_distrib, intro setsum.cong) simp_all
+    by (subst setsum_distrib_left, intro setsum.cong) simp_all
   finally show ?thesis ..
 qed
 
@@ -477,7 +477,7 @@
   also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) - (-1) ^ i * of_nat (n choose i))"
     by (simp add: setsum_subtractf)
   also have "\<dots> = 2 * (\<Sum>i\<le>n. if odd i then of_nat (n choose i) else 0)"
-    by (subst setsum_right_distrib, intro setsum.cong) simp_all
+    by (subst setsum_distrib_left, intro setsum.cong) simp_all
   finally show ?thesis ..
 qed
 
@@ -914,7 +914,7 @@
   have "(\<Sum>i\<le>n. i * (n choose i)) = (\<Sum>i\<le>Suc m. i * (Suc m choose i))"
     by (simp add: Suc)
   also have "\<dots> = Suc m * 2 ^ m"
-    by (simp only: setsum_atMost_Suc_shift Suc_times_binomial setsum_right_distrib[symmetric])
+    by (simp only: setsum_atMost_Suc_shift Suc_times_binomial setsum_distrib_left[symmetric])
        (simp add: choose_row_sum')
   finally show ?thesis
     using Suc by simp
@@ -934,9 +934,9 @@
       (\<Sum>i\<le>Suc m. (-1) ^ i * of_nat i * of_nat (Suc m choose i))"
     by (simp add: Suc)
   also have "\<dots> = (\<Sum>i\<le>m. (-1) ^ (Suc i) * of_nat (Suc i * (Suc m choose Suc i)))"
-    by (simp only: setsum_atMost_Suc_shift setsum_right_distrib[symmetric] mult_ac of_nat_mult) simp
+    by (simp only: setsum_atMost_Suc_shift setsum_distrib_left[symmetric] mult_ac of_nat_mult) simp
   also have "\<dots> = - of_nat (Suc m) * (\<Sum>i\<le>m. (-1) ^ i * of_nat (m choose i))"
-    by (subst setsum_right_distrib, rule setsum.cong[OF refl], subst Suc_times_binomial)
+    by (subst setsum_distrib_left, rule setsum.cong[OF refl], subst Suc_times_binomial)
        (simp add: algebra_simps)
   also have "(\<Sum>i\<le>m. (-1 :: 'a) ^ i * of_nat ((m choose i))) = 0"
     using choose_alternating_sum[OF \<open>m > 0\<close>] by simp
@@ -978,7 +978,7 @@
     by (subst setsum_atMost_Suc_shift) (simp add: ring_distribs setsum.distrib)
   also have "(\<Sum>i\<le>n. of_nat (n choose i) * pochhammer a (Suc i) * pochhammer b (n - i)) =
       a * pochhammer ((a + 1) + b) n"
-    by (subst Suc) (simp add: setsum_right_distrib pochhammer_rec mult_ac)
+    by (subst Suc) (simp add: setsum_distrib_left pochhammer_rec mult_ac)
   also have "(\<Sum>i\<le>n. of_nat (n choose Suc i) * pochhammer a (Suc i) * pochhammer b (n - i)) +
         pochhammer b (Suc n) =
       (\<Sum>i=0..Suc n. of_nat (n choose i) * pochhammer a i * pochhammer b (Suc n - i))"
@@ -992,7 +992,7 @@
   also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) * pochhammer a i * pochhammer b (Suc (n - i)))"
     by (intro setsum.cong) (simp_all add: Suc_diff_le)
   also have "\<dots> = b * pochhammer (a + (b + 1)) n"
-    by (subst Suc) (simp add: setsum_right_distrib mult_ac pochhammer_rec)
+    by (subst Suc) (simp add: setsum_distrib_left mult_ac pochhammer_rec)
   also have "a * pochhammer ((a + 1) + b) n + b * pochhammer (a + (b + 1)) n =
       pochhammer (a + b) (Suc n)"
     by (simp add: pochhammer_rec algebra_simps)
@@ -1263,9 +1263,9 @@
       by (simp only:)
   qed
   also have "\<dots> + ?B = y * (\<Sum>k=1..mm. (G mm k)) + (of_nat mm + r gchoose (Suc mm)) * x^(Suc mm)"
-    unfolding G_def by (subst setsum_right_distrib) (simp add: algebra_simps)
+    unfolding G_def by (subst setsum_distrib_left) (simp add: algebra_simps)
   also have "(\<Sum>k=0..mm. (of_nat mm + r gchoose k) * x^(Suc k) * y^(mm - k)) = x * (S mm)"
-    unfolding S_def by (subst setsum_right_distrib) (simp add: atLeast0AtMost algebra_simps)
+    unfolding S_def by (subst setsum_distrib_left) (simp add: atLeast0AtMost algebra_simps)
   also have "(G (Suc mm) 0) = y * (G mm 0)"
     by (simp add: G_def)
   finally have "S (Suc mm) =
@@ -1283,7 +1283,7 @@
   also have "(x + y) * S mm + \<dots> = (x + y) * ?rhs mm + (- r gchoose (Suc mm)) * (- x)^Suc mm"
     unfolding S_def by (subst Suc.IH) simp
   also have "(x + y) * ?rhs mm = (\<Sum>n\<le>mm. ((- r gchoose n) * (- x) ^ n * (x + y) ^ (Suc mm - n)))"
-    by (subst setsum_right_distrib, rule setsum.cong) (simp_all add: Suc_diff_le)
+    by (subst setsum_distrib_left, rule setsum.cong) (simp_all add: Suc_diff_le)
   also have "\<dots> + (-r gchoose (Suc mm)) * (-x)^Suc mm =
       (\<Sum>n\<le>Suc mm. (- r gchoose n) * (- x) ^ n * (x + y) ^ (Suc mm - n))"
     by simp
@@ -1345,7 +1345,7 @@
     using gbinomial_partial_sum_poly_xpos[where x="1" and y="1" and r="of_nat m + 1" and m="m"]
     by (simp add: add_ac)
   also have "\<dots> = 2 ^ m * (\<Sum>k\<le>m. (of_nat (m + k) gchoose k) / 2 ^ k)"
-    by (subst setsum_right_distrib) (simp add: algebra_simps power_diff)
+    by (subst setsum_distrib_left) (simp add: algebra_simps power_diff)
   finally show ?thesis
     by (subst (asm) mult_left_cancel) simp_all
 qed
@@ -1444,7 +1444,7 @@
     by simp
   also have "\<dots> = nat (\<Sum>I | I \<subseteq> A \<and> I \<noteq> {}. (\<Sum>_\<in>\<Inter>I. (- 1) ^ (card I + 1)))"
     (is "_ = nat ?rhs")
-    by (subst setsum_right_distrib) simp
+    by (subst setsum_distrib_left) simp
   also have "?rhs = (\<Sum>(I, _)\<in>Sigma {I. I \<subseteq> A \<and> I \<noteq> {}} Inter. (- 1) ^ (card I + 1))"
     using assms by (subst setsum.Sigma) auto
   also have "\<dots> = (\<Sum>(x, I)\<in>(SIGMA x:UNIV. {I. I \<subseteq> A \<and> I \<noteq> {} \<and> x \<in> \<Inter>I}). (- 1) ^ (card I + 1))"
@@ -1474,7 +1474,7 @@
     also have "\<dots> = (\<Sum>i=1..card A. (\<Sum>I|I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. (- 1) ^ (i + 1)))"
       using assms by (subst setsum.Sigma) auto
     also have "\<dots> = (\<Sum>i=1..card A. (- 1) ^ (i + 1) * (\<Sum>I|I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. 1))"
-      by (subst setsum_right_distrib) simp
+      by (subst setsum_distrib_left) simp
     also have "\<dots> = (\<Sum>i=1..card K. (- 1) ^ (i + 1) * (\<Sum>I|I \<subseteq> K \<and> card I = i. 1))"
       (is "_ = ?rhs")
     proof (rule setsum.mono_neutral_cong_right[rule_format])
@@ -1508,7 +1508,7 @@
     also have "\<dots> = (\<Sum>i = 0..card K. (- 1) * ((- 1) ^ i * int (card K choose i))) + 1"
       using K by (subst n_subsets[symmetric]) simp_all
     also have "\<dots> = - (\<Sum>i = 0..card K. (- 1) ^ i * int (card K choose i)) + 1"
-      by (subst setsum_right_distrib[symmetric]) simp
+      by (subst setsum_distrib_left[symmetric]) simp
     also have "\<dots> =  - ((-1 + 1) ^ card K) + 1"
       by (subst binomial_ring) (simp add: ac_simps)
     also have "\<dots> = 1"
--- a/src/HOL/Decision_Procs/Approximation.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -31,7 +31,7 @@
   have shift_pow: "\<And>i. - (x * ((-1)^i * a (Suc i) * x ^ i)) = (-1)^(Suc i) * a (Suc i) * x ^ (Suc i)"
     by auto
   show ?thesis
-    unfolding setsum_right_distrib shift_pow uminus_add_conv_diff [symmetric] setsum_negf[symmetric]
+    unfolding setsum_distrib_left shift_pow uminus_add_conv_diff [symmetric] setsum_negf[symmetric]
     setsum_head_upt_Suc[OF zero_less_Suc]
     setsum.reindex[OF inj_Suc, unfolded comp_def, symmetric, of "\<lambda> n. (-1)^n  *a n * x^n"] by auto
 qed
@@ -514,7 +514,7 @@
   proof -
     have "(sqrt y * lb_arctan_horner prec n 1 y) \<le> ?S n"
       using bounds(1) \<open>0 \<le> sqrt y\<close>
-      apply (simp only: power_add power_one_right mult.assoc[symmetric] setsum_left_distrib[symmetric])
+      apply (simp only: power_add power_one_right mult.assoc[symmetric] setsum_distrib_right[symmetric])
       apply (simp only: mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult)
       apply (auto intro!: mult_left_mono)
       done
@@ -527,7 +527,7 @@
     have "arctan (sqrt y) \<le> ?S (Suc n)" using arctan_bounds ..
     also have "\<dots> \<le> (sqrt y * ub_arctan_horner prec (Suc n) 1 y)"
       using bounds(2)[of "Suc n"] \<open>0 \<le> sqrt y\<close>
-      apply (simp only: power_add power_one_right mult.assoc[symmetric] setsum_left_distrib[symmetric])
+      apply (simp only: power_add power_one_right mult.assoc[symmetric] setsum_distrib_right[symmetric])
       apply (simp only: mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult)
       apply (auto intro!: mult_left_mono)
       done
@@ -1212,7 +1212,7 @@
   from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0,
     OF \<open>0 \<le> real_of_float (x * x)\<close> f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps]
   show "?lb" and "?ub" using \<open>0 \<le> real_of_float x\<close>
-    apply (simp_all only: power_add power_one_right mult.assoc[symmetric] setsum_left_distrib[symmetric])
+    apply (simp_all only: power_add power_one_right mult.assoc[symmetric] setsum_distrib_right[symmetric])
     apply (simp_all only: mult.commute[where 'a=real] of_nat_fact)
     apply (auto intro!: mult_left_mono simp add: power_mult power2_eq_square[of "real_of_float x"])
     done
@@ -2193,7 +2193,7 @@
   let "?s n" = "(- 1) ^ n * (1 / real (1 + n)) * (real_of_float x)^(Suc n)"
 
   have "?lb \<le> setsum ?s {0 ..< 2 * ev}"
-    unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric]
+    unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_distrib_right[symmetric]
     unfolding mult.commute[of "real_of_float x"] ev 
     using horner_bounds(1)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" 
                     and lb="\<lambda>n i k x. lb_ln_horner prec n k x" 
@@ -2208,7 +2208,7 @@
   have "?ln \<le> setsum ?s {0 ..< 2 * od + 1}"
     using ln_bounds(2)[OF \<open>0 \<le> real_of_float x\<close> \<open>real_of_float x < 1\<close>] by auto
   also have "\<dots> \<le> ?ub"
-    unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric]
+    unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_distrib_right[symmetric]
     unfolding mult.commute[of "real_of_float x"] od
     using horner_bounds(2)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" and lb="\<lambda>n i k x. lb_ln_horner prec n k x" and ub="\<lambda>n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*od+1",
       OF \<open>0 \<le> real_of_float x\<close> refl lb_ln_horner.simps ub_ln_horner.simps] \<open>0 \<le> real_of_float x\<close>
@@ -4019,7 +4019,7 @@
                inverse (real (\<Prod> j \<in> {k..<k+Suc n}. j)) * ?f (Suc n) t * (xs!x - c)^Suc n" (is "_ = ?T")
         unfolding funpow_Suc C_def[symmetric] setsum_move0 setprod_head_Suc
         by (auto simp add: algebra_simps)
-          (simp only: mult.left_commute [of _ "inverse (real k)"] setsum_right_distrib [symmetric])
+          (simp only: mult.left_commute [of _ "inverse (real k)"] setsum_distrib_left [symmetric])
       finally have "?T \<in> {l .. u}" .
     }
     thus ?thesis using DERIV by blast
--- a/src/HOL/Deriv.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Deriv.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -368,7 +368,7 @@
     using insert by (intro has_derivative_mult) auto
   also have "?P = (\<lambda>y. \<Sum>i'\<in>insert i I. f' i' y * (\<Prod>j\<in>insert i I - {i'}. f j x))"
     using insert(1,2)
-    by (auto simp add: setsum_right_distrib insert_Diff_if intro!: ext setsum.cong)
+    by (auto simp add: setsum_distrib_left insert_Diff_if intro!: ext setsum.cong)
   finally show ?case
     using insert by simp
 qed
@@ -845,7 +845,7 @@
   "(\<And> n. n \<in> S \<Longrightarrow> ((\<lambda>x. f x n) has_field_derivative (f' x n)) F) \<Longrightarrow>
     ((\<lambda>x. setsum (f x) S) has_field_derivative setsum (f' x) S) F"
   by (rule has_derivative_imp_has_field_derivative [OF has_derivative_setsum])
-     (auto simp: setsum_right_distrib mult_commute_abs dest: has_field_derivative_imp_has_derivative)
+     (auto simp: setsum_distrib_left mult_commute_abs dest: has_field_derivative_imp_has_derivative)
 
 lemma DERIV_inverse'[derivative_intros]:
   assumes "(f has_field_derivative D) (at x within s)"
--- a/src/HOL/Groups_Big.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Groups_Big.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -729,7 +729,7 @@
   "finite F \<Longrightarrow> (setsum f F = 0) = (\<forall>a\<in>F. f a = 0)"
   by (intro ballI setsum_nonneg_eq_0_iff zero_le)
 
-lemma setsum_right_distrib: "r * setsum f A = setsum (\<lambda>n. r * f n) A"
+lemma setsum_distrib_left: "r * setsum f A = setsum (\<lambda>n. r * f n) A"
   for f :: "'a \<Rightarrow> 'b::semiring_0"
 proof (induct A rule: infinite_finite_induct)
   case infinite
@@ -742,7 +742,7 @@
   then show ?case by (simp add: distrib_left)
 qed
 
-lemma setsum_left_distrib: "setsum f A * r = (\<Sum>n\<in>A. f n * r)"
+lemma setsum_distrib_right: "setsum f A * r = (\<Sum>n\<in>A. f n * r)"
   for r :: "'a::semiring_0"
 proof (induct A rule: infinite_finite_induct)
   case infinite
@@ -811,7 +811,7 @@
 lemma setsum_product:
   fixes f :: "'a \<Rightarrow> 'b::semiring_0"
   shows "setsum f A * setsum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"
-  by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum.commute)
+  by (simp add: setsum_distrib_left setsum_distrib_right) (rule setsum.commute)
 
 lemma setsum_mult_setsum_if_inj:
   fixes f :: "'a \<Rightarrow> 'b::semiring_0"
--- a/src/HOL/Inequalities.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Inequalities.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -59,7 +59,7 @@
   let ?S = "(\<Sum>j=0..<n. (\<Sum>k=0..<n. (a j - a k) * (b j - b k)))"
   have "2 * (of_nat n * (\<Sum>j=0..<n. (a j * b j)) - (\<Sum>j=0..<n. b j) * (\<Sum>k=0..<n. a k)) = ?S"
     by (simp only: one_add_one[symmetric] algebra_simps)
-      (simp add: algebra_simps setsum_subtractf setsum.distrib setsum.commute[of "\<lambda>i j. a i * b j"] setsum_right_distrib)
+      (simp add: algebra_simps setsum_subtractf setsum.distrib setsum.commute[of "\<lambda>i j. a i * b j"] setsum_distrib_left)
   also
   { fix i j::nat assume "i<n" "j<n"
     hence "a i - a j \<le> 0 \<and> b i - b j \<ge> 0 \<or> a i - a j \<ge> 0 \<and> b i - b j \<le> 0"
--- a/src/HOL/Library/BigO.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Library/BigO.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -556,7 +556,7 @@
   apply (subst abs_of_nonneg) back back
    apply (rule setsum_nonneg)
    apply force
-  apply (subst setsum_right_distrib)
+  apply (subst setsum_distrib_left)
   apply (rule allI)
   apply (rule order_trans)
    apply (rule setsum_abs)
--- a/src/HOL/Library/Convex.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Library/Convex.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -633,7 +633,7 @@
             OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"]
       by simp
     also have "\<dots> = (\<Sum> j \<in> s. (1 - a i) * ?a j * f (y j)) + a i * f (y i)"
-      unfolding setsum_right_distrib[of "1 - a i" "\<lambda> j. ?a j * f (y j)"]
+      unfolding setsum_distrib_left[of "1 - a i" "\<lambda> j. ?a j * f (y j)"]
       using i0 by auto
     also have "\<dots> = (\<Sum> j \<in> s. a j * f (y j)) + a i * f (y i)"
       using i0 by auto
--- a/src/HOL/Library/Extended_Real.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Library/Extended_Real.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -1098,7 +1098,7 @@
   "(\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> setsum f A * r = (\<Sum>n\<in>A. f n * r :: ereal)"
   using setsum_ereal_right_distrib[of A f r] by (simp add: mult_ac)
 
-lemma setsum_left_distrib_ereal:
+lemma setsum_distrib_right_ereal:
   "c \<ge> 0 \<Longrightarrow> setsum f A * ereal c = (\<Sum>x\<in>A. f x * c :: ereal)"
 by(subst setsum_comp_morphism[where h="\<lambda>x. x * ereal c", symmetric])(simp_all add: distrib_left_ereal_nn)
 
--- a/src/HOL/Library/Formal_Power_Series.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Library/Formal_Power_Series.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -141,7 +141,7 @@
           (\<Sum>j=0..n. \<Sum>i=0..n - j. a$j * b$i * c$(n - j - i))"
       by (rule fps_mult_assoc_lemma)
     then show "((a * b) * c) $ n = (a * (b * c)) $ n"
-      by (simp add: fps_mult_nth setsum_right_distrib setsum_left_distrib mult.assoc)
+      by (simp add: fps_mult_nth setsum_distrib_left setsum_distrib_right mult.assoc)
   qed
 qed
 
@@ -1562,7 +1562,7 @@
     also have "\<dots> = (fps_deriv (f * g)) $ n"
       apply (simp only: fps_deriv_nth fps_mult_nth setsum.distrib)
       unfolding s0 s1
-      unfolding setsum.distrib[symmetric] setsum_right_distrib
+      unfolding setsum.distrib[symmetric] setsum_distrib_left
       apply (rule setsum.cong)
       apply (auto simp add: of_nat_diff field_simps)
       done
@@ -2320,7 +2320,7 @@
       apply auto
       apply (rule setsum.cong)
       apply (rule refl)
-      unfolding setsum_left_distrib
+      unfolding setsum_distrib_right
       apply (rule sym)
       apply (rule_tac l = "\<lambda>xs. xs @ [n - x]" in setsum.reindex_cong)
       apply (simp add: inj_on_def)
@@ -3082,7 +3082,7 @@
   have "(fps_deriv (a oo b))$n = (((fps_deriv a) oo b) * (fps_deriv b)) $n" for n
   proof -
     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 field_simps setsum_right_distrib del: of_nat_Suc)
+      by (simp add: fps_compose_def field_simps setsum_distrib_left 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: 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}"
@@ -3102,7 +3102,7 @@
     have "(((fps_deriv a) oo b) * (fps_deriv b))$n = setsum (\<lambda>i. (fps_deriv b)$ (n - i) * ((fps_deriv a) oo b)$i) {0..n}"
       unfolding fps_mult_nth by (simp add: ac_simps)
     also have "\<dots> = setsum (\<lambda>i. setsum (\<lambda>j. of_nat (n - i +1) * b$(n - i + 1) * of_nat (j + 1) * a$(j+1) * (b^j)$i) {0..n}) {0..n}"
-      unfolding fps_deriv_nth fps_compose_nth setsum_right_distrib mult.assoc
+      unfolding fps_deriv_nth fps_compose_nth setsum_distrib_left mult.assoc
       apply (rule setsum.cong)
       apply (rule refl)
       apply (rule setsum.mono_neutral_left)
@@ -3114,7 +3114,7 @@
       apply simp
       done
     also have "\<dots> = setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
-      unfolding setsum_right_distrib
+      unfolding setsum_distrib_left
       apply (subst setsum.commute)
       apply (rule setsum.cong, rule refl)+
       apply simp
@@ -3295,7 +3295,7 @@
     apply auto
     done
   have "?r =  setsum (\<lambda>i. setsum (\<lambda>(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \<le> n}) {0..n}"
-    apply (simp add: fps_mult_nth setsum_right_distrib)
+    apply (simp add: fps_mult_nth setsum_distrib_left)
     apply (subst setsum.commute)
     apply (rule setsum.cong)
     apply (auto simp add: field_simps)
@@ -3377,7 +3377,7 @@
   assumes c0: "c $ 0 = (0::'a::idom)"
   shows "(a * b) oo c = (a oo c) * (b oo c)"
   apply (simp add: fps_eq_iff fps_compose_mult_distrib_lemma [OF c0])
-  apply (simp add: fps_compose_nth fps_mult_nth setsum_left_distrib)
+  apply (simp add: fps_compose_nth fps_mult_nth setsum_distrib_right)
   done
 
 lemma fps_compose_setprod_distrib:
@@ -3498,7 +3498,7 @@
 qed
 
 lemma fps_const_mult_apply_left: "fps_const c * (a oo b) = (fps_const c * a) oo b"
-  by (simp add: fps_eq_iff fps_compose_nth setsum_right_distrib mult.assoc)
+  by (simp add: fps_eq_iff fps_compose_nth setsum_distrib_left mult.assoc)
 
 lemma fps_const_mult_apply_right:
   "(a oo b) * fps_const (c::'a::comm_semiring_1) = (fps_const c * a) oo b"
@@ -3513,11 +3513,11 @@
   proof -
     have "?l$n = (setsum (\<lambda>i. (fps_const (a$i) * b^i) oo c) {0..n})$n"
       by (simp add: fps_compose_nth fps_compose_power[OF c0] fps_const_mult_apply_left
-        setsum_right_distrib mult.assoc fps_setsum_nth)
+        setsum_distrib_left mult.assoc fps_setsum_nth)
     also have "\<dots> = ((setsum (\<lambda>i. fps_const (a$i) * b^i) {0..n}) oo c)$n"
       by (simp add: fps_compose_setsum_distrib)
     also have "\<dots> = ?r$n"
-      apply (simp add: fps_compose_nth fps_setsum_nth setsum_left_distrib mult.assoc)
+      apply (simp add: fps_compose_nth fps_setsum_nth setsum_distrib_right mult.assoc)
       apply (rule setsum.cong)
       apply (rule refl)
       apply (rule setsum.mono_neutral_right)
@@ -4224,7 +4224,7 @@
     apply (simp add: th00)
     unfolding gbinomial_pochhammer
     using bn0
-    apply (simp add: setsum_left_distrib setsum_right_distrib field_simps)
+    apply (simp add: setsum_distrib_right setsum_distrib_left field_simps)
     done
   finally show ?thesis by simp
 qed
@@ -4253,7 +4253,7 @@
     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_simps setsum_right_distrib)
+    using nz by (simp add: field_simps setsum_distrib_left)
 qed
 
 
--- a/src/HOL/Library/Groups_Big_Fun.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Library/Groups_Big_Fun.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -240,7 +240,7 @@
   note assms
   moreover have "{a. g a * r \<noteq> 0} \<subseteq> {a. g a \<noteq> 0}" by auto
   ultimately show ?thesis
-    by (simp add: setsum_left_distrib Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
+    by (simp add: setsum_distrib_right Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
 qed  
 
 lemma Sum_any_right_distrib:
@@ -251,7 +251,7 @@
   note assms
   moreover have "{a. r * g a \<noteq> 0} \<subseteq> {a. g a \<noteq> 0}" by auto
   ultimately show ?thesis
-    by (simp add: setsum_right_distrib Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
+    by (simp add: setsum_distrib_left Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
 qed
 
 lemma Sum_any_product:
--- a/src/HOL/Library/Multiset.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Library/Multiset.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -656,40 +656,40 @@
 
 subsubsection \<open>Intersection\<close>
 
-definition inf_subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<inter>" 70) where
+definition inf_subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "\<inter>#" 70) where
   multiset_inter_def: "inf_subset_mset A B = A - (A - B)"
 
 interpretation subset_mset: semilattice_inf inf_subset_mset "op \<subseteq>#" "op \<subset>#"
 proof -
   have [simp]: "m \<le> n \<Longrightarrow> m \<le> q \<Longrightarrow> m \<le> n - (n - q)" for m n q :: nat
     by arith
-  show "class.semilattice_inf op #\<inter> op \<subseteq># op \<subset>#"
+  show "class.semilattice_inf op \<inter># op \<subseteq># op \<subset>#"
     by standard (auto simp add: multiset_inter_def subseteq_mset_def)
 qed
   \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
 
 lemma multiset_inter_count [simp]:
   fixes A B :: "'a multiset"
-  shows "count (A #\<inter> B) x = min (count A x) (count B x)"
+  shows "count (A \<inter># B) x = min (count A x) (count B x)"
   by (simp add: multiset_inter_def)
 
 lemma set_mset_inter [simp]:
-  "set_mset (A #\<inter> B) = set_mset A \<inter> set_mset B"
+  "set_mset (A \<inter># B) = set_mset A \<inter> set_mset B"
   by (simp only: set_eq_iff count_greater_zero_iff [symmetric] multiset_inter_count) simp
 
 lemma diff_intersect_left_idem [simp]:
-  "M - M #\<inter> N = M - N"
+  "M - M \<inter># N = M - N"
   by (simp add: multiset_eq_iff min_def)
 
 lemma diff_intersect_right_idem [simp]:
-  "M - N #\<inter> M = M - N"
+  "M - N \<inter># M = M - N"
   by (simp add: multiset_eq_iff min_def)
 
-lemma multiset_inter_single[simp]: "a \<noteq> b \<Longrightarrow> {#a#} #\<inter> {#b#} = {#}"
+lemma multiset_inter_single[simp]: "a \<noteq> b \<Longrightarrow> {#a#} \<inter># {#b#} = {#}"
   by (rule multiset_eqI) auto
 
 lemma multiset_union_diff_commute:
-  assumes "B #\<inter> C = {#}"
+  assumes "B \<inter># C = {#}"
   shows "A + B - C = A - C + B"
 proof (rule multiset_eqI)
   fix x
@@ -702,7 +702,7 @@
 qed
 
 lemma disjunct_not_in:
-  "A #\<inter> B = {#} \<longleftrightarrow> (\<forall>a. a \<notin># A \<or> a \<notin># B)" (is "?P \<longleftrightarrow> ?Q")
+  "A \<inter># B = {#} \<longleftrightarrow> (\<forall>a. a \<notin># A \<or> a \<notin># B)" (is "?P \<longleftrightarrow> ?Q")
 proof
   assume ?P
   show ?Q
@@ -722,46 +722,46 @@
     fix a
     from \<open>?Q\<close> have "count A a = 0 \<or> count B a = 0"
       by (auto simp add: not_in_iff)
-    then show "count (A #\<inter> B) a = count {#} a"
+    then show "count (A \<inter># B) a = count {#} a"
       by auto
   qed
 qed
 
 lemma add_mset_inter_add_mset[simp]:
-  "add_mset a A #\<inter> add_mset a B = add_mset a (A #\<inter> B)"
+  "add_mset a A \<inter># add_mset a B = add_mset a (A \<inter># B)"
   by (metis add_mset_add_single add_mset_diff_bothsides diff_subset_eq_self multiset_inter_def
       subset_mset.diff_add_assoc2)
 
 lemma add_mset_disjoint [simp]:
-  "add_mset a A #\<inter> B = {#} \<longleftrightarrow> a \<notin># B \<and> A #\<inter> B = {#}"
-  "{#} = add_mset a A #\<inter> B \<longleftrightarrow> a \<notin># B \<and> {#} = A #\<inter> B"
+  "add_mset a A \<inter># B = {#} \<longleftrightarrow> a \<notin># B \<and> A \<inter># B = {#}"
+  "{#} = add_mset a A \<inter># B \<longleftrightarrow> a \<notin># B \<and> {#} = A \<inter># B"
   by (auto simp: disjunct_not_in)
 
 lemma disjoint_add_mset [simp]:
-  "B #\<inter> add_mset a A = {#} \<longleftrightarrow> a \<notin># B \<and> B #\<inter> A = {#}"
-  "{#} = A #\<inter> add_mset b B \<longleftrightarrow> b \<notin># A \<and> {#} = A #\<inter> B"
+  "B \<inter># add_mset a A = {#} \<longleftrightarrow> a \<notin># B \<and> B \<inter># A = {#}"
+  "{#} = A \<inter># add_mset b B \<longleftrightarrow> b \<notin># A \<and> {#} = A \<inter># B"
   by (auto simp: disjunct_not_in)
 
-lemma empty_inter[simp]: "{#} #\<inter> M = {#}"
+lemma empty_inter[simp]: "{#} \<inter># M = {#}"
   by (simp add: multiset_eq_iff)
 
-lemma inter_empty[simp]: "M #\<inter> {#} = {#}"
+lemma inter_empty[simp]: "M \<inter># {#} = {#}"
   by (simp add: multiset_eq_iff)
 
-lemma inter_add_left1: "\<not> x \<in># N \<Longrightarrow> (add_mset x M) #\<inter> N = M #\<inter> N"
+lemma inter_add_left1: "\<not> x \<in># N \<Longrightarrow> (add_mset x M) \<inter># N = M \<inter># N"
   by (simp add: multiset_eq_iff not_in_iff)
 
-lemma inter_add_left2: "x \<in># N \<Longrightarrow> (add_mset x M) #\<inter> N = add_mset x (M #\<inter> (N - {#x#}))"
+lemma inter_add_left2: "x \<in># N \<Longrightarrow> (add_mset x M) \<inter># N = add_mset x (M \<inter># (N - {#x#}))"
   by (auto simp add: multiset_eq_iff elim: mset_add)
 
-lemma inter_add_right1: "\<not> x \<in># N \<Longrightarrow> N #\<inter> (add_mset x M) = N #\<inter> M"
+lemma inter_add_right1: "\<not> x \<in># N \<Longrightarrow> N \<inter># (add_mset x M) = N \<inter># M"
   by (simp add: multiset_eq_iff not_in_iff)
 
-lemma inter_add_right2: "x \<in># N \<Longrightarrow> N #\<inter> (add_mset x M) = add_mset x ((N - {#x#}) #\<inter> M)"
+lemma inter_add_right2: "x \<in># N \<Longrightarrow> N \<inter># (add_mset x M) = add_mset x ((N - {#x#}) \<inter># M)"
   by (auto simp add: multiset_eq_iff elim: mset_add)
 
 lemma disjunct_set_mset_diff:
-  assumes "M #\<inter> N = {#}"
+  assumes "M \<inter># N = {#}"
   shows "set_mset (M - N) = set_mset M"
 proof (rule set_eqI)
   fix a
@@ -787,85 +787,85 @@
 qed
 
 lemma inter_iff:
-  "a \<in># A #\<inter> B \<longleftrightarrow> a \<in># A \<and> a \<in># B"
+  "a \<in># A \<inter># B \<longleftrightarrow> a \<in># A \<and> a \<in># B"
   by simp
 
 lemma inter_union_distrib_left:
-  "A #\<inter> B + C = (A + C) #\<inter> (B + C)"
+  "A \<inter># B + C = (A + C) \<inter># (B + C)"
   by (simp add: multiset_eq_iff min_add_distrib_left)
 
 lemma inter_union_distrib_right:
-  "C + A #\<inter> B = (C + A) #\<inter> (C + B)"
+  "C + A \<inter># B = (C + A) \<inter># (C + B)"
   using inter_union_distrib_left [of A B C] by (simp add: ac_simps)
 
 lemma inter_subset_eq_union:
-  "A #\<inter> B \<subseteq># A + B"
+  "A \<inter># B \<subseteq># A + B"
   by (auto simp add: subseteq_mset_def)
 
 
 subsubsection \<open>Bounded union\<close>
 
-definition sup_subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"(infixl "#\<union>" 70)
+definition sup_subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"(infixl "\<union>#" 70)
   where "sup_subset_mset A B = A + (B - A)" \<comment> \<open>FIXME irregular fact name\<close>
 
 interpretation subset_mset: semilattice_sup sup_subset_mset "op \<subseteq>#" "op \<subset>#"
 proof -
   have [simp]: "m \<le> n \<Longrightarrow> q \<le> n \<Longrightarrow> m + (q - m) \<le> n" for m n q :: nat
     by arith
-  show "class.semilattice_sup op #\<union> op \<subseteq># op \<subset>#"
+  show "class.semilattice_sup op \<union># op \<subseteq># op \<subset>#"
     by standard (auto simp add: sup_subset_mset_def subseteq_mset_def)
 qed
   \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
 
-interpretation subset_mset: bounded_lattice_bot "op #\<inter>" "op \<subseteq>#" "op \<subset>#"
-  "op #\<union>" "{#}"
+interpretation subset_mset: bounded_lattice_bot "op \<inter>#" "op \<subseteq>#" "op \<subset>#"
+  "op \<union>#" "{#}"
   by standard auto
 
 lemma sup_subset_mset_count [simp]: \<comment> \<open>FIXME irregular fact name\<close>
-  "count (A #\<union> B) x = max (count A x) (count B x)"
+  "count (A \<union># B) x = max (count A x) (count B x)"
   by (simp add: sup_subset_mset_def)
 
 lemma set_mset_sup [simp]:
-  "set_mset (A #\<union> B) = set_mset A \<union> set_mset B"
+  "set_mset (A \<union># B) = set_mset A \<union> set_mset B"
   by (simp only: set_eq_iff count_greater_zero_iff [symmetric] sup_subset_mset_count)
     (auto simp add: not_in_iff elim: mset_add)
 
-lemma empty_sup: "{#} #\<union> M = M"
+lemma empty_sup: "{#} \<union># M = M"
   by (fact subset_mset.sup_bot.left_neutral)
 
-lemma sup_empty: "M #\<union> {#} = M"
+lemma sup_empty: "M \<union># {#} = M"
   by (fact subset_mset.sup_bot.right_neutral)
 
-lemma sup_union_left1 [simp]: "\<not> x \<in># N \<Longrightarrow> (add_mset x M) #\<union> N = add_mset x (M #\<union> N)"
+lemma sup_union_left1 [simp]: "\<not> x \<in># N \<Longrightarrow> (add_mset x M) \<union># N = add_mset x (M \<union># N)"
   by (simp add: multiset_eq_iff not_in_iff)
 
-lemma sup_union_left2: "x \<in># N \<Longrightarrow> (add_mset x M) #\<union> N = add_mset x (M #\<union> (N - {#x#}))"
+lemma sup_union_left2: "x \<in># N \<Longrightarrow> (add_mset x M) \<union># N = add_mset x (M \<union># (N - {#x#}))"
   by (simp add: multiset_eq_iff)
 
-lemma sup_union_right1 [simp]: "\<not> x \<in># N \<Longrightarrow> N #\<union> (add_mset x M) = add_mset x (N #\<union> M)"
+lemma sup_union_right1 [simp]: "\<not> x \<in># N \<Longrightarrow> N \<union># (add_mset x M) = add_mset x (N \<union># M)"
   by (simp add: multiset_eq_iff not_in_iff)
 
-lemma sup_union_right2: "x \<in># N \<Longrightarrow> N #\<union> (add_mset x M) = add_mset x ((N - {#x#}) #\<union> M)"
+lemma sup_union_right2: "x \<in># N \<Longrightarrow> N \<union># (add_mset x M) = add_mset x ((N - {#x#}) \<union># M)"
   by (simp add: multiset_eq_iff)
 
 lemma sup_union_distrib_left:
-  "A #\<union> B + C = (A + C) #\<union> (B + C)"
+  "A \<union># B + C = (A + C) \<union># (B + C)"
   by (simp add: multiset_eq_iff max_add_distrib_left)
 
 lemma union_sup_distrib_right:
-  "C + A #\<union> B = (C + A) #\<union> (C + B)"
+  "C + A \<union># B = (C + A) \<union># (C + B)"
   using sup_union_distrib_left [of A B C] by (simp add: ac_simps)
 
 lemma union_diff_inter_eq_sup:
-  "A + B - A #\<inter> B = A #\<union> B"
+  "A + B - A \<inter># B = A \<union># B"
   by (auto simp add: multiset_eq_iff)
 
 lemma union_diff_sup_eq_inter:
-  "A + B - A #\<union> B = A #\<inter> B"
+  "A + B - A \<union># B = A \<inter># B"
   by (auto simp add: multiset_eq_iff)
 
 lemma add_mset_union:
-  \<open>add_mset a A #\<union> add_mset a B = add_mset a (A #\<union> B)\<close>
+  \<open>add_mset a A \<union># add_mset a B = add_mset a (A \<union># B)\<close>
   by (auto simp: multiset_eq_iff max_def)
 
 
@@ -1096,7 +1096,7 @@
   using assms by (simp add: Sup_multiset_def Abs_multiset_inverse Sup_multiset_in_multiset)
 
 
-interpretation subset_mset: conditionally_complete_lattice Inf Sup "op #\<inter>" "op \<subseteq>#" "op \<subset>#" "op #\<union>"
+interpretation subset_mset: conditionally_complete_lattice Inf Sup "op \<inter>#" "op \<subseteq>#" "op \<subset>#" "op \<union>#"
 proof
   fix X :: "'a multiset" and A
   assume "X \<in> A"
@@ -1220,10 +1220,10 @@
   with assms show ?thesis by (simp add: in_Sup_multiset_iff)
 qed
 
-interpretation subset_mset: distrib_lattice "op #\<inter>" "op \<subseteq>#" "op \<subset>#" "op #\<union>"
+interpretation subset_mset: distrib_lattice "op \<inter>#" "op \<subseteq>#" "op \<subset>#" "op \<union>#"
 proof
   fix A B C :: "'a multiset"
-  show "A #\<union> (B #\<inter> C) = A #\<union> B #\<inter> (A #\<union> C)"
+  show "A \<union># (B \<inter># C) = A \<union># B \<inter># (A \<union># C)"
     by (intro multiset_eqI) simp_all
 qed
 
@@ -1263,10 +1263,10 @@
 lemma filter_diff_mset [simp]: "filter_mset P (M - N) = filter_mset P M - filter_mset P N"
   by (rule multiset_eqI) simp
 
-lemma filter_inter_mset [simp]: "filter_mset P (M #\<inter> N) = filter_mset P M #\<inter> filter_mset P N"
+lemma filter_inter_mset [simp]: "filter_mset P (M \<inter># N) = filter_mset P M \<inter># filter_mset P N"
   by (rule multiset_eqI) simp
 
-lemma filter_sup_mset[simp]: "filter_mset P (A #\<union> B) = filter_mset P A #\<union> filter_mset P B"
+lemma filter_sup_mset[simp]: "filter_mset P (A \<union># B) = filter_mset P A \<union># filter_mset P B"
   by (rule multiset_eqI) simp
 
 lemma filter_mset_add_mset [simp]:
@@ -2702,10 +2702,10 @@
 
 lemma mult_cancel_max:
   assumes "trans s" and "irrefl s"
-  shows "(X, Y) \<in> mult s \<longleftrightarrow> (X - X #\<inter> Y, Y - X #\<inter> Y) \<in> mult s" (is "?L \<longleftrightarrow> ?R")
+  shows "(X, Y) \<in> mult s \<longleftrightarrow> (X - X \<inter># Y, Y - X \<inter># Y) \<in> mult s" (is "?L \<longleftrightarrow> ?R")
 proof -
-  have "X - X #\<inter> Y + X #\<inter> Y = X" "Y - X #\<inter> Y + X #\<inter> Y = Y" by (auto simp: count_inject[symmetric])
-  thus ?thesis using mult_cancel[OF assms, of "X - X #\<inter> Y"  "X #\<inter> Y" "Y - X #\<inter> Y"] by auto
+  have "X - X \<inter># Y + X \<inter># Y = X" "Y - X \<inter># Y + X \<inter># Y = Y" by (auto simp: count_inject[symmetric])
+  thus ?thesis using mult_cancel[OF assms, of "X - X \<inter># Y"  "X \<inter># Y" "Y - X \<inter># Y"] by auto
 qed
 
 
@@ -2714,37 +2714,37 @@
 text \<open>
   Predicate variants of \<open>mult\<close> and the reflexive closure of \<open>mult\<close>, which are
   executable whenever the given predicate \<open>P\<close> is. Together with the standard
-  code equations for \<open>op #\<inter>\<close> and \<open>op -\<close> this should yield quadratic
+  code equations for \<open>op \<inter>#\<close> and \<open>op -\<close> this should yield quadratic
   (with respect to calls to \<open>P\<close>) implementations of \<open>multp\<close> and \<open>multeqp\<close>.
 \<close>
 
 definition multp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
   "multp P N M =
-    (let Z = M #\<inter> N; X = M - Z in
+    (let Z = M \<inter># N; X = M - Z in
     X \<noteq> {#} \<and> (let Y = N - Z in (\<forall>y \<in> set_mset Y. \<exists>x \<in> set_mset X. P y x)))"
 
 definition multeqp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
   "multeqp P N M =
-    (let Z = M #\<inter> N; X = M - Z; Y = N - Z in
+    (let Z = M \<inter># N; X = M - Z; Y = N - Z in
     (\<forall>y \<in> set_mset Y. \<exists>x \<in> set_mset X. P y x))"
 
 lemma multp_iff:
   assumes "irrefl R" and "trans R" and [simp]: "\<And>x y. P x y \<longleftrightarrow> (x, y) \<in> R"
   shows "multp P N M \<longleftrightarrow> (N, M) \<in> mult R" (is "?L \<longleftrightarrow> ?R")
 proof -
-  have *: "M #\<inter> N + (N - M #\<inter> N) = N" "M #\<inter> N + (M - M #\<inter> N) = M"
-    "(M - M #\<inter> N) #\<inter> (N - M #\<inter> N) = {#}" by (auto simp: count_inject[symmetric])
+  have *: "M \<inter># N + (N - M \<inter># N) = N" "M \<inter># N + (M - M \<inter># N) = M"
+    "(M - M \<inter># N) \<inter># (N - M \<inter># N) = {#}" by (auto simp: count_inject[symmetric])
   show ?thesis
   proof
     assume ?L thus ?R
-      using one_step_implies_mult[of "M - M #\<inter> N" "N - M #\<inter> N" R "M #\<inter> N"] *
+      using one_step_implies_mult[of "M - M \<inter># N" "N - M \<inter># N" R "M \<inter># N"] *
       by (auto simp: multp_def Let_def)
   next
-    { fix I J K :: "'a multiset" assume "(I + J) #\<inter> (I + K) = {#}"
+    { fix I J K :: "'a multiset" assume "(I + J) \<inter># (I + K) = {#}"
       then have "I = {#}" by (metis inter_union_distrib_right union_eq_empty)
     } note [dest!] = this
     assume ?R thus ?L
-      using mult_implies_one_step[OF assms(2), of "N - M #\<inter> N" "M - M #\<inter> N"]
+      using mult_implies_one_step[OF assms(2), of "N - M \<inter># N" "M - M \<inter># N"]
         mult_cancel_max[OF assms(2,1), of "N" "M"] * by (auto simp: multp_def)
   qed
 qed
@@ -2753,9 +2753,9 @@
   assumes "irrefl R" and "trans R" and "\<And>x y. P x y \<longleftrightarrow> (x, y) \<in> R"
   shows "multeqp P N M \<longleftrightarrow> (N, M) \<in> (mult R)\<^sup>="
 proof -
-  { assume "N \<noteq> M" "M - M #\<inter> N = {#}"
+  { assume "N \<noteq> M" "M - M \<inter># N = {#}"
     then obtain y where "count N y \<noteq> count M y" by (auto simp: count_inject[symmetric])
-    then have "\<exists>y. count M y < count N y" using `M - M #\<inter> N = {#}`
+    then have "\<exists>y. count M y < count N y" using `M - M \<inter># N = {#}`
       by (auto simp: count_inject[symmetric] dest!: le_neq_implies_less fun_cong[of _ _ y])
   }
   then have "multeqp P N M \<longleftrightarrow> multp P N M \<or> N = M"
@@ -3035,13 +3035,13 @@
 lemma mset_subset_trans: "(M::'a multiset) \<subset># K \<Longrightarrow> K \<subset># N \<Longrightarrow> M \<subset># N"
   by (fact subset_mset.less_trans)
 
-lemma multiset_inter_commute: "A #\<inter> B = B #\<inter> A"
+lemma multiset_inter_commute: "A \<inter># B = B \<inter># A"
   by (fact subset_mset.inf.commute)
 
-lemma multiset_inter_assoc: "A #\<inter> (B #\<inter> C) = A #\<inter> B #\<inter> C"
+lemma multiset_inter_assoc: "A \<inter># (B \<inter># C) = A \<inter># B \<inter># C"
   by (fact subset_mset.inf.assoc [symmetric])
 
-lemma multiset_inter_left_commute: "A #\<inter> (B #\<inter> C) = B #\<inter> (A #\<inter> C)"
+lemma multiset_inter_left_commute: "A \<inter># (B \<inter># C) = B \<inter># (A \<inter># C)"
   by (fact subset_mset.inf.left_commute)
 
 lemmas multiset_inter_ac =
@@ -3112,24 +3112,24 @@
   by (rule sym, induct ys arbitrary: xs) (simp_all add: diff_add diff_right_commute diff_diff_add)
 
 lemma [code]:
-  "mset xs #\<inter> mset ys =
+  "mset xs \<inter># mset ys =
     mset (snd (fold (\<lambda>x (ys, zs).
       if x \<in> set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, [])))"
 proof -
   have "\<And>zs. mset (snd (fold (\<lambda>x (ys, zs).
     if x \<in> set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, zs))) =
-      (mset xs #\<inter> mset ys) + mset zs"
+      (mset xs \<inter># mset ys) + mset zs"
     by (induct xs arbitrary: ys)
       (auto simp add: inter_add_right1 inter_add_right2 ac_simps)
   then show ?thesis by simp
 qed
 
 lemma [code]:
-  "mset xs #\<union> mset ys =
+  "mset xs \<union># mset ys =
     mset (case_prod append (fold (\<lambda>x (ys, zs). (remove1 x ys, x # zs)) xs (ys, [])))"
 proof -
   have "\<And>zs. mset (case_prod append (fold (\<lambda>x (ys, zs). (remove1 x ys, x # zs)) xs (ys, zs))) =
-      (mset xs #\<union> mset ys) + mset zs"
+      (mset xs \<union># mset ys) + mset zs"
     by (induct xs arbitrary: ys) (simp_all add: multiset_eq_iff)
   then show ?thesis by simp
 qed
--- a/src/HOL/Library/Polynomial.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Library/Polynomial.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -507,7 +507,7 @@
       also note pCons.IH
       also have "a + x * (\<Sum>i\<le>degree p. coeff p i * x ^ i) =
                  coeff ?p' 0 * x^0 + (\<Sum>i\<le>degree p. coeff ?p' (Suc i) * x^Suc i)"
-          by (simp add: field_simps setsum_right_distrib coeff_pCons)
+          by (simp add: field_simps setsum_distrib_left coeff_pCons)
       also note setsum_atMost_Suc_shift[symmetric]
       also note degree_pCons_eq[OF \<open>p \<noteq> 0\<close>, of a, symmetric]
       finally show ?thesis .
@@ -3412,7 +3412,7 @@
       by (subst setprod.insert, insert insert, auto)
   } note id2 = this
   show ?case
-    unfolding id pderiv_mult insert(3) setsum_right_distrib
+    unfolding id pderiv_mult insert(3) setsum_distrib_left
     by (auto simp add: ac_simps id2 intro!: setsum.cong)
 qed auto
 
--- a/src/HOL/Library/Stirling.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Library/Stirling.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -110,7 +110,7 @@
     also have "\<dots> = Suc n * (\<Sum>k=1..n. fact n div k) + Suc n * fact n div Suc n"
       by (metis nat.distinct(1) nonzero_mult_divide_cancel_left)
     also have "\<dots> = (\<Sum>k=1..n. fact (Suc n) div k) + fact (Suc n) div Suc n"
-      by (simp add: setsum_right_distrib div_mult_swap dvd_fact)
+      by (simp add: setsum_distrib_left div_mult_swap dvd_fact)
     also have "\<dots> = (\<Sum>k=1..Suc n. fact (Suc n) div k)"
       by simp
     finally show ?thesis .
@@ -162,7 +162,7 @@
   also have "\<dots> = (\<Sum>k\<le>n. n * stirling n (Suc k) + stirling n k)"
     by simp
   also have "\<dots> = n * (\<Sum>k\<le>n. stirling n (Suc k)) + (\<Sum>k\<le>n. stirling n k)"
-    by (simp add: setsum.distrib setsum_right_distrib)
+    by (simp add: setsum.distrib setsum_distrib_left)
   also have "\<dots> = n * fact n + fact n"
   proof -
     have "n * (\<Sum>k\<le>n. stirling n (Suc k)) = n * ((\<Sum>k\<le>Suc n. stirling n k) - stirling n 0)"
@@ -195,7 +195,7 @@
     by (subst setsum_atMost_Suc_shift) (simp add: setsum.distrib ring_distribs)
   also have "\<dots> = pochhammer x (Suc n)"
     by (subst setsum_atMost_Suc_shift [symmetric])
-      (simp add: algebra_simps setsum.distrib setsum_right_distrib pochhammer_Suc Suc [symmetric])
+      (simp add: algebra_simps setsum.distrib setsum_distrib_left pochhammer_Suc Suc [symmetric])
   finally show ?case .
 qed
 
--- a/src/HOL/Metis_Examples/Big_O.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Metis_Examples/Big_O.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -499,7 +499,7 @@
 apply (subst abs_of_nonneg) back back
  apply (rule setsum_nonneg)
  apply force
-apply (subst setsum_right_distrib)
+apply (subst setsum_distrib_left)
 apply (rule allI)
 apply (rule order_trans)
  apply (rule setsum_abs)
--- a/src/HOL/Nonstandard_Analysis/HSeries.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Nonstandard_Analysis/HSeries.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -59,7 +59,7 @@
 
 lemma sumhr_mult:
   "!!m n. hypreal_of_real r * sumhr(m,n,f) = sumhr(m,n,%n. r * f n)"
-unfolding sumhr_app by transfer (rule setsum_right_distrib)
+unfolding sumhr_app by transfer (rule setsum_distrib_left)
 
 lemma sumhr_split_add:
   "!!n p. n < p ==> sumhr(0,n,f) + sumhr(n,p,f) = sumhr(0,p,f)"
--- a/src/HOL/Number_Theory/Factorial_Ring.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Number_Theory/Factorial_Ring.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -541,7 +541,7 @@
     by simp
   with A B C have subset: "A \<subseteq># B + C"
     by (subst (asm) prod_mset_dvd_prod_mset_primes_iff) auto
-  define A1 and A2 where "A1 = A #\<inter> B" and "A2 = A - A1"
+  define A1 and A2 where "A1 = A \<inter># B" and "A2 = A - A1"
   have "A = A1 + A2" unfolding A1_def A2_def
     by (rule sym, intro subset_mset.add_diff_inverse) simp_all
   from subset have "A1 \<subseteq># B" "A2 \<subseteq># C"
@@ -1440,10 +1440,10 @@
 
 definition "gcd_factorial a b = (if a = 0 then normalize b
      else if b = 0 then normalize a
-     else prod_mset (prime_factorization a #\<inter> prime_factorization b))"
+     else prod_mset (prime_factorization a \<inter># prime_factorization b))"
 
 definition "lcm_factorial a b = (if a = 0 \<or> b = 0 then 0
-     else prod_mset (prime_factorization a #\<union> prime_factorization b))"
+     else prod_mset (prime_factorization a \<union># prime_factorization b))"
 
 definition "Gcd_factorial A =
   (if A \<subseteq> {0} then 0 else prod_mset (Inf (prime_factorization ` (A - {0}))))"
@@ -1457,24 +1457,24 @@
 
 lemma prime_factorization_gcd_factorial:
   assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
-  shows   "prime_factorization (gcd_factorial a b) = prime_factorization a #\<inter> prime_factorization b"
+  shows   "prime_factorization (gcd_factorial a b) = prime_factorization a \<inter># prime_factorization b"
 proof -
   have "prime_factorization (gcd_factorial a b) =
-          prime_factorization (prod_mset (prime_factorization a #\<inter> prime_factorization b))"
+          prime_factorization (prod_mset (prime_factorization a \<inter># prime_factorization b))"
     by (simp add: gcd_factorial_def)
-  also have "\<dots> = prime_factorization a #\<inter> prime_factorization b"
+  also have "\<dots> = prime_factorization a \<inter># prime_factorization b"
     by (subst prime_factorization_prod_mset_primes) auto
   finally show ?thesis .
 qed
 
 lemma prime_factorization_lcm_factorial:
   assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
-  shows   "prime_factorization (lcm_factorial a b) = prime_factorization a #\<union> prime_factorization b"
+  shows   "prime_factorization (lcm_factorial a b) = prime_factorization a \<union># prime_factorization b"
 proof -
   have "prime_factorization (lcm_factorial a b) =
-          prime_factorization (prod_mset (prime_factorization a #\<union> prime_factorization b))"
+          prime_factorization (prod_mset (prime_factorization a \<union># prime_factorization b))"
     by (simp add: lcm_factorial_def)
-  also have "\<dots> = prime_factorization a #\<union> prime_factorization b"
+  also have "\<dots> = prime_factorization a \<union># prime_factorization b"
     by (subst prime_factorization_prod_mset_primes) auto
   finally show ?thesis .
 qed
@@ -1527,8 +1527,8 @@
 
 lemma normalize_gcd_factorial: "normalize (gcd_factorial a b) = gcd_factorial a b"
 proof -
-  have "normalize (prod_mset (prime_factorization a #\<inter> prime_factorization b)) =
-          prod_mset (prime_factorization a #\<inter> prime_factorization b)"
+  have "normalize (prod_mset (prime_factorization a \<inter># prime_factorization b)) =
+          prod_mset (prime_factorization a \<inter># prime_factorization b)"
     by (intro normalize_prod_mset_primes) auto
   thus ?thesis by (simp add: gcd_factorial_def)
 qed
@@ -1541,7 +1541,7 @@
   from that False have "?p c \<subseteq># ?p a" "?p c \<subseteq># ?p b"
     by (simp_all add: prime_factorization_subset_iff_dvd)
   hence "prime_factorization c \<subseteq>#
-           prime_factorization (prod_mset (prime_factorization a #\<inter> prime_factorization b))"
+           prime_factorization (prod_mset (prime_factorization a \<inter># prime_factorization b))"
     using False by (subst prime_factorization_prod_mset_primes) auto
   with False show ?thesis
     by (auto simp: gcd_factorial_def prime_factorization_subset_iff_dvd [symmetric])
@@ -1553,12 +1553,12 @@
   case False
   let ?p = "prime_factorization"
   from False have "prod_mset (?p (a * b)) div gcd_factorial a b =
-                     prod_mset (?p a + ?p b - ?p a #\<inter> ?p b)"
+                     prod_mset (?p a + ?p b - ?p a \<inter># ?p b)"
     by (subst prod_mset_diff) (auto simp: lcm_factorial_def gcd_factorial_def
                                 prime_factorization_mult subset_mset.le_infI1)
   also from False have "prod_mset (?p (a * b)) = normalize (a * b)"
     by (intro prod_mset_prime_factorization) simp_all
-  also from False have "prod_mset (?p a + ?p b - ?p a #\<inter> ?p b) = lcm_factorial a b"
+  also from False have "prod_mset (?p a + ?p b - ?p a \<inter># ?p b) = lcm_factorial a b"
     by (simp add: union_diff_inter_eq_sup lcm_factorial_def)
   finally show ?thesis ..
 qed (auto simp: lcm_factorial_def)
@@ -1750,12 +1750,12 @@
 
 lemma prime_factorization_gcd:
   assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
-  shows   "prime_factorization (gcd a b) = prime_factorization a #\<inter> prime_factorization b"
+  shows   "prime_factorization (gcd a b) = prime_factorization a \<inter># prime_factorization b"
   by (simp add: gcd_eq_gcd_factorial prime_factorization_gcd_factorial)
 
 lemma prime_factorization_lcm:
   assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
-  shows   "prime_factorization (lcm a b) = prime_factorization a #\<union> prime_factorization b"
+  shows   "prime_factorization (lcm a b) = prime_factorization a \<union># prime_factorization b"
   by (simp add: lcm_eq_lcm_factorial prime_factorization_lcm_factorial)
 
 lemma prime_factorization_Gcd:
--- a/src/HOL/Probability/Distributions.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Probability/Distributions.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -149,7 +149,7 @@
   assume "0 \<le> x"
   have "(\<Sum>n\<le>k. (l * x) ^ n * exp (- (l * x)) / fact n) =
     exp (- (l * x)) * (\<Sum>n\<le>k. (l * x) ^ n / fact n)"
-    unfolding setsum_right_distrib by (intro setsum.cong) (auto simp: field_simps)
+    unfolding setsum_distrib_left by (intro setsum.cong) (auto simp: field_simps)
   also have "\<dots> = (\<Sum>n\<le>k. (l * x) ^ n / fact n) / exp (l * x)"
     by (simp add: exp_minus field_simps)
   also have "\<dots> \<le> 1"
--- a/src/HOL/Probability/Probability_Mass_Function.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Probability/Probability_Mass_Function.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -1610,7 +1610,7 @@
 
 lemma nn_integral_pmf_of_set: "nn_integral (measure_pmf pmf_of_set) f = setsum f S / card S"
   by (subst nn_integral_measure_pmf_finite)
-     (simp_all add: setsum_left_distrib[symmetric] card_gt_0_iff S_not_empty S_finite divide_ennreal_def
+     (simp_all add: setsum_distrib_right[symmetric] card_gt_0_iff S_not_empty S_finite divide_ennreal_def
                 divide_ennreal[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric] ennreal_times_divide)
 
 lemma integral_pmf_of_set: "integral\<^sup>L (measure_pmf pmf_of_set) f = setsum f S / card S"
--- a/src/HOL/Probability/Projective_Limit.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Probability/Projective_Limit.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -290,7 +290,7 @@
       finally show "\<mu>G (Z i - Z' i) \<le> (2 powr - real i) * ?a" .
     qed
     also have "\<dots> = ennreal ((\<Sum> i\<in>{1..n}. (2 powr -enn2real i)) * enn2real ?a)"
-      using r by (simp add: setsum_left_distrib ennreal_mult[symmetric])
+      using r by (simp add: setsum_distrib_right ennreal_mult[symmetric])
     also have "\<dots> < ennreal (1 * enn2real ?a)"
     proof (intro ennreal_lessI mult_strict_right_mono)
       have "(\<Sum>i = 1..n. 2 powr - real i) = (\<Sum>i = 1..<Suc n. (1/2) ^ i)"
--- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -100,7 +100,7 @@
   show ?case unfolding *
     using Suc[of "\<lambda>i. P (Suc i)"]
     by (simp add: setsum.reindex split_conv setsum_cartesian_product'
-      lessThan_Suc_eq_insert_0 setprod.reindex setsum_left_distrib[symmetric] setsum_right_distrib[symmetric])
+      lessThan_Suc_eq_insert_0 setprod.reindex setsum_distrib_right[symmetric] setsum_distrib_left[symmetric])
 qed
 
 declare space_point_measure[simp]
@@ -158,8 +158,8 @@
 
   have [simp]: "\<And>A. inj_on (\<lambda>(xs, n). n # xs) A" by (force intro!: inj_onI)
 
-  note setsum_right_distrib[symmetric, simp]
-  note setsum_left_distrib[symmetric, simp]
+  note setsum_distrib_left[symmetric, simp]
+  note setsum_distrib_right[symmetric, simp]
   note setsum_cartesian_product'[simp]
 
   have "(\<Sum>ms | set ms \<subseteq> messages \<and> length ms = n. \<Prod>x<length ms. M (ms ! x)) = 1"
@@ -256,7 +256,7 @@
     apply (simp add: * P_def)
     apply (simp add: setsum_cartesian_product')
     using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\<lambda>x x. True"] \<open>k \<in> keys\<close>
-    by (auto simp add: setsum_right_distrib[symmetric] subset_eq setprod.neutral_const)
+    by (auto simp add: setsum_distrib_left[symmetric] subset_eq setprod.neutral_const)
 qed
 
 lemma fst_image_msgs[simp]: "fst`msgs = keys"
@@ -323,7 +323,7 @@
   also have "\<dots> = - (\<Sum>y\<in>Y`msgs. (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))))"
     by (subst setsum.commute) rule
   also have "\<dots> = -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))"
-    by (auto simp add: setsum_right_distrib vimage_def intro!: setsum.cong prob_conj_imp1)
+    by (auto simp add: setsum_distrib_left vimage_def intro!: setsum.cong prob_conj_imp1)
   finally show "- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))) =
     -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))" .
 qed simp_all
@@ -402,7 +402,7 @@
     also have "\<P>(t\<circ>OB | fst) {(t obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(t\<circ>OB|fst) {(t obs, k')} * \<P>(fst) {k'}) =
       \<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'})"
       using CP_t_K[OF \<open>k\<in>keys\<close> obs] CP_t_K[OF _ obs] \<open>real (card ?S) \<noteq> 0\<close>
-      by (simp only: setsum_right_distrib[symmetric] ac_simps
+      by (simp only: setsum_distrib_left[symmetric] ac_simps
           mult_divide_mult_cancel_left[OF \<open>real (card ?S) \<noteq> 0\<close>]
         cong: setsum.cong)
     also have "(\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'}) = \<P>(OB) {obs}"
@@ -450,7 +450,7 @@
     apply (safe intro!: setsum.cong)
     using P_t_sum_P_O
     by (simp add: setsum_divide_distrib[symmetric] field_simps **
-                  setsum_right_distrib[symmetric] setsum_left_distrib[symmetric])
+                  setsum_distrib_left[symmetric] setsum_distrib_right[symmetric])
   also have "\<dots> = \<H>(fst | t\<circ>OB)"
     unfolding conditional_entropy_eq_ce_with_hypothesis
     by (simp add: comp_def image_image[symmetric])
--- a/src/HOL/Quotient_Examples/FSet.thy	Mon Sep 19 17:37:22 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1163 +0,0 @@
-(*  Title:      HOL/Quotient_Examples/FSet.thy
-    Author:     Cezary Kaliszyk, TU Munich
-    Author:     Christian Urban, TU Munich
-
-Type of finite sets.
-*)
-
-(********************************************************************
-  WARNING: There is a formalization of 'a fset as a subtype of sets in
-  HOL/Library/FSet.thy using Lifting/Transfer. The user should use
-  that file rather than this file unless there are some very specific
-  reasons.
-*********************************************************************)
-
-theory FSet
-imports "~~/src/HOL/Library/Multiset" "~~/src/HOL/Library/Quotient_List"
-begin
-
-text \<open>
-  The type of finite sets is created by a quotient construction
-  over lists. The definition of the equivalence:
-\<close>
-
-definition
-  list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50)
-where
-  [simp]: "xs \<approx> ys \<longleftrightarrow> set xs = set ys"
-
-lemma list_eq_reflp:
-  "reflp list_eq"
-  by (auto intro: reflpI)
-
-lemma list_eq_symp:
-  "symp list_eq"
-  by (auto intro: sympI)
-
-lemma list_eq_transp:
-  "transp list_eq"
-  by (auto intro: transpI)
-
-lemma list_eq_equivp:
-  "equivp list_eq"
-  by (auto intro: equivpI list_eq_reflp list_eq_symp list_eq_transp)
-
-text \<open>The \<open>fset\<close> type\<close>
-
-quotient_type
-  'a fset = "'a list" / "list_eq"
-  by (rule list_eq_equivp)
-
-text \<open>
-  Definitions for sublist, cardinality, 
-  intersection, difference and respectful fold over 
-  lists.
-\<close>
-
-declare List.member_def [simp]
-
-definition
-  sub_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
-where 
-  [simp]: "sub_list xs ys \<longleftrightarrow> set xs \<subseteq> set ys"
-
-definition
-  card_list :: "'a list \<Rightarrow> nat"
-where
-  [simp]: "card_list xs = card (set xs)"
-
-definition
-  inter_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where
-  [simp]: "inter_list xs ys = [x \<leftarrow> xs. x \<in> set xs \<and> x \<in> set ys]"
-
-definition
-  diff_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where
-  [simp]: "diff_list xs ys = [x \<leftarrow> xs. x \<notin> set ys]"
-
-definition
-  rsp_fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
-where
-  "rsp_fold f \<longleftrightarrow> (\<forall>u v. f u \<circ> f v = f v \<circ> f u)"
-
-lemma rsp_foldI:
-  "(\<And>u v. f u \<circ> f v = f v \<circ> f u) \<Longrightarrow> rsp_fold f"
-  by (simp add: rsp_fold_def)
-
-lemma rsp_foldE:
-  assumes "rsp_fold f"
-  obtains "f u \<circ> f v = f v \<circ> f u"
-  using assms by (simp add: rsp_fold_def)
-
-definition
-  fold_once :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b"
-where
-  "fold_once f xs = (if rsp_fold f then fold f (remdups xs) else id)"
-
-lemma fold_once_default [simp]:
-  "\<not> rsp_fold f \<Longrightarrow> fold_once f xs = id"
-  by (simp add: fold_once_def)
-
-lemma fold_once_fold_remdups:
-  "rsp_fold f \<Longrightarrow> fold_once f xs = fold f (remdups xs)"
-  by (simp add: fold_once_def)
-
-
-section \<open>Quotient composition lemmas\<close>
-
-lemma list_all2_refl':
-  assumes q: "equivp R"
-  shows "(list_all2 R) r r"
-  by (rule list_all2_refl) (metis equivp_def q)
-
-lemma compose_list_refl:
-  assumes q: "equivp R"
-  shows "(list_all2 R OOO op \<approx>) r r"
-proof
-  have *: "r \<approx> r" by (rule equivp_reflp[OF fset_equivp])
-  show "list_all2 R r r" by (rule list_all2_refl'[OF q])
-  with * show "(op \<approx> OO list_all2 R) r r" ..
-qed
-
-lemma map_list_eq_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba"
-  by (simp only: list_eq_def set_map)
-
-lemma quotient_compose_list_g:
-  assumes q: "Quotient3 R Abs Rep"
-  and     e: "equivp R"
-  shows  "Quotient3 ((list_all2 R) OOO (op \<approx>))
-    (abs_fset \<circ> (map Abs)) ((map Rep) \<circ> rep_fset)"
-  unfolding Quotient3_def comp_def
-proof (intro conjI allI)
-  fix a r s
-  show "abs_fset (map Abs (map Rep (rep_fset a))) = a"
-    by (simp add: abs_o_rep[OF q] Quotient3_abs_rep[OF Quotient3_fset] List.map.id)
-  have b: "list_all2 R (map Rep (rep_fset a)) (map Rep (rep_fset a))"
-    by (rule list_all2_refl'[OF e])
-  have c: "(op \<approx> OO list_all2 R) (map Rep (rep_fset a)) (map Rep (rep_fset a))"
-    by (rule, rule equivp_reflp[OF fset_equivp]) (rule b)
-  show "(list_all2 R OOO op \<approx>) (map Rep (rep_fset a)) (map Rep (rep_fset a))"
-    by (rule, rule list_all2_refl'[OF e]) (rule c)
-  show "(list_all2 R OOO op \<approx>) r s = ((list_all2 R OOO op \<approx>) r r \<and>
-        (list_all2 R OOO op \<approx>) s s \<and> abs_fset (map Abs r) = abs_fset (map Abs s))"
-  proof (intro iffI conjI)
-    show "(list_all2 R OOO op \<approx>) r r" by (rule compose_list_refl[OF e])
-    show "(list_all2 R OOO op \<approx>) s s" by (rule compose_list_refl[OF e])
-  next
-    assume a: "(list_all2 R OOO op \<approx>) r s"
-    then have b: "map Abs r \<approx> map Abs s"
-    proof (elim relcomppE)
-      fix b ba
-      assume c: "list_all2 R r b"
-      assume d: "b \<approx> ba"
-      assume e: "list_all2 R ba s"
-      have f: "map Abs r = map Abs b"
-        using Quotient3_rel[OF list_quotient3[OF q]] c by blast
-      have "map Abs ba = map Abs s"
-        using Quotient3_rel[OF list_quotient3[OF q]] e by blast
-      then have g: "map Abs s = map Abs ba" by simp
-      then show "map Abs r \<approx> map Abs s" using d f map_list_eq_cong by simp
-    qed
-    then show "abs_fset (map Abs r) = abs_fset (map Abs s)"
-      using Quotient3_rel[OF Quotient3_fset] by blast
-  next
-    assume a: "(list_all2 R OOO op \<approx>) r r \<and> (list_all2 R OOO op \<approx>) s s
-      \<and> abs_fset (map Abs r) = abs_fset (map Abs s)"
-    then have s: "(list_all2 R OOO op \<approx>) s s" by simp
-    have d: "map Abs r \<approx> map Abs s"
-      by (subst Quotient3_rel [OF Quotient3_fset, symmetric]) (simp add: a)
-    have b: "map Rep (map Abs r) \<approx> map Rep (map Abs s)"
-      by (rule map_list_eq_cong[OF d])
-    have y: "list_all2 R (map Rep (map Abs s)) s"
-      by (fact rep_abs_rsp_left[OF list_quotient3[OF q], OF list_all2_refl'[OF e, of s]])
-    have c: "(op \<approx> OO list_all2 R) (map Rep (map Abs r)) s"
-      by (rule relcomppI) (rule b, rule y)
-    have z: "list_all2 R r (map Rep (map Abs r))"
-      by (fact rep_abs_rsp[OF list_quotient3[OF q], OF list_all2_refl'[OF e, of r]])
-    then show "(list_all2 R OOO op \<approx>) r s"
-      using a c relcomppI by simp
-  qed
-qed
-
-lemma quotient_compose_list[quot_thm]:
-  shows  "Quotient3 ((list_all2 op \<approx>) OOO (op \<approx>))
-    (abs_fset \<circ> (map abs_fset)) ((map rep_fset) \<circ> rep_fset)"
-  by (rule quotient_compose_list_g, rule Quotient3_fset, rule list_eq_equivp)
-
-
-section \<open>Quotient definitions for fsets\<close>
-
-
-subsection \<open>Finite sets are a bounded, distributive lattice with minus\<close>
-
-instantiation fset :: (type) "{bounded_lattice_bot, distrib_lattice, minus}"
-begin
-
-quotient_definition
-  "bot :: 'a fset" 
-  is "Nil :: 'a list" done
-
-abbreviation
-  empty_fset  ("{||}")
-where
-  "{||} \<equiv> bot :: 'a fset"
-
-quotient_definition
-  "less_eq_fset :: ('a fset \<Rightarrow> 'a fset \<Rightarrow> bool)"
-  is "sub_list :: ('a list \<Rightarrow> 'a list \<Rightarrow> bool)" by simp
-
-abbreviation
-  subset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50)
-where
-  "xs |\<subseteq>| ys \<equiv> xs \<le> ys"
-
-definition
-  less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool"
-where  
-  "xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)"
-
-abbreviation
-  psubset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50)
-where
-  "xs |\<subset>| ys \<equiv> xs < ys"
-
-quotient_definition
-  "sup :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-  is "append :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" by simp
-
-abbreviation
-  union_fset (infixl "|\<union>|" 65)
-where
-  "xs |\<union>| ys \<equiv> sup xs (ys::'a fset)"
-
-quotient_definition
-  "inf :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-  is "inter_list :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" by simp
-
-abbreviation
-  inter_fset (infixl "|\<inter>|" 65)
-where
-  "xs |\<inter>| ys \<equiv> inf xs (ys::'a fset)"
-
-quotient_definition
-  "minus :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-  is "diff_list :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" by fastforce
-
-instance
-proof
-  fix x y z :: "'a fset"
-  show "x |\<subset>| y \<longleftrightarrow> x |\<subseteq>| y \<and> \<not> y |\<subseteq>| x"
-    by (unfold less_fset_def, descending) auto
-  show "x |\<subseteq>| x" by (descending) (simp)
-  show "{||} |\<subseteq>| x" by (descending) (simp)
-  show "x |\<subseteq>| x |\<union>| y" by (descending) (simp)
-  show "y |\<subseteq>| x |\<union>| y" by (descending) (simp)
-  show "x |\<inter>| y |\<subseteq>| x" by (descending) (auto)
-  show "x |\<inter>| y |\<subseteq>| y" by (descending) (auto)
-  show "x |\<union>| (y |\<inter>| z) = x |\<union>| y |\<inter>| (x |\<union>| z)" 
-    by (descending) (auto)
-next
-  fix x y z :: "'a fset"
-  assume a: "x |\<subseteq>| y"
-  assume b: "y |\<subseteq>| z"
-  show "x |\<subseteq>| z" using a b by (descending) (simp)
-next
-  fix x y :: "'a fset"
-  assume a: "x |\<subseteq>| y"
-  assume b: "y |\<subseteq>| x"
-  show "x = y" using a b by (descending) (auto)
-next
-  fix x y z :: "'a fset"
-  assume a: "y |\<subseteq>| x"
-  assume b: "z |\<subseteq>| x"
-  show "y |\<union>| z |\<subseteq>| x" using a b by (descending) (simp)
-next
-  fix x y z :: "'a fset"
-  assume a: "x |\<subseteq>| y"
-  assume b: "x |\<subseteq>| z"
-  show "x |\<subseteq>| y |\<inter>| z" using a b by (descending) (auto)
-qed
-
-end
-
-
-subsection \<open>Other constants for fsets\<close>
-
-quotient_definition
-  "insert_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-  is "Cons" by auto
-
-syntax
-  "_insert_fset"     :: "args => 'a fset"  ("{|(_)|}")
-
-translations
-  "{|x, xs|}" == "CONST insert_fset x {|xs|}"
-  "{|x|}"     == "CONST insert_fset x {||}"
-
-quotient_definition
-  fset_member
-where
-  "fset_member :: 'a fset \<Rightarrow> 'a \<Rightarrow> bool" is "List.member" by fastforce
-
-abbreviation
-  in_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<in>|" 50)
-where
-  "x |\<in>| S \<equiv> fset_member S x"
-
-abbreviation
-  notin_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<notin>|" 50)
-where
-  "x |\<notin>| S \<equiv> \<not> (x |\<in>| S)"
-
-
-subsection \<open>Other constants on the Quotient Type\<close>
-
-quotient_definition
-  "card_fset :: 'a fset \<Rightarrow> nat"
-  is card_list by simp
-
-quotient_definition
-  "map_fset :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
-  is map by simp
-
-quotient_definition
-  "remove_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-  is removeAll by simp
-
-quotient_definition
-  "fset :: 'a fset \<Rightarrow> 'a set"
-  is "set" by simp
-
-lemma fold_once_set_equiv:
-  assumes "xs \<approx> ys"
-  shows "fold_once f xs = fold_once f ys"
-proof (cases "rsp_fold f")
-  case False then show ?thesis by simp
-next
-  case True
-  then have "\<And>x y. x \<in> set (remdups xs) \<Longrightarrow> y \<in> set (remdups xs) \<Longrightarrow> f x \<circ> f y = f y \<circ> f x"
-    by (rule rsp_foldE)
-  moreover from assms have "mset (remdups xs) = mset (remdups ys)"
-    by (simp add: set_eq_iff_mset_remdups_eq)
-  ultimately have "fold f (remdups xs) = fold f (remdups ys)"
-    by (rule fold_multiset_equiv)
-  with True show ?thesis by (simp add: fold_once_fold_remdups)
-qed
-
-quotient_definition
-  "fold_fset :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b \<Rightarrow> 'b"
-  is fold_once by (rule fold_once_set_equiv)
-
-lemma concat_rsp_pre:
-  assumes a: "list_all2 op \<approx> x x'"
-  and     b: "x' \<approx> y'"
-  and     c: "list_all2 op \<approx> y' y"
-  and     d: "\<exists>x\<in>set x. xa \<in> set x"
-  shows "\<exists>x\<in>set y. xa \<in> set x"
-proof -
-  obtain xb where e: "xb \<in> set x" and f: "xa \<in> set xb" using d by auto
-  have "\<exists>y. y \<in> set x' \<and> xb \<approx> y" by (rule list_all2_find_element[OF e a])
-  then obtain ya where h: "ya \<in> set x'" and i: "xb \<approx> ya" by auto
-  have "ya \<in> set y'" using b h by simp
-  then have "\<exists>yb. yb \<in> set y \<and> ya \<approx> yb" using c by (rule list_all2_find_element)
-  then show ?thesis using f i by auto
-qed
-
-quotient_definition
-  "concat_fset :: ('a fset) fset \<Rightarrow> 'a fset"
-  is concat 
-proof (elim relcomppE)
-fix a b ba bb
-  assume a: "list_all2 op \<approx> a ba"
-  with list_symp [OF list_eq_symp] have a': "list_all2 op \<approx> ba a" by (rule sympE)
-  assume b: "ba \<approx> bb"
-  with list_eq_symp have b': "bb \<approx> ba" by (rule sympE)
-  assume c: "list_all2 op \<approx> bb b"
-  with list_symp [OF list_eq_symp] have c': "list_all2 op \<approx> b bb" by (rule sympE)
-  have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" 
-  proof
-    fix x
-    show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" 
-    proof
-      assume d: "\<exists>xa\<in>set a. x \<in> set xa"
-      show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d])
-    next
-      assume e: "\<exists>xa\<in>set b. x \<in> set xa"
-      show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e])
-    qed
-  qed
-  then show "concat a \<approx> concat b" by auto
-qed
-
-quotient_definition
-  "filter_fset :: ('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-  is filter by force
-
-
-subsection \<open>Compositional respectfulness and preservation lemmas\<close>
-
-lemma Nil_rsp2 [quot_respect]: 
-  shows "(list_all2 op \<approx> OOO op \<approx>) Nil Nil"
-  by (rule compose_list_refl, rule list_eq_equivp)
-
-lemma Cons_rsp2 [quot_respect]:
-  shows "(op \<approx> ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) Cons Cons"
-  apply (auto intro!: rel_funI)
-  apply (rule_tac b="x # b" in relcomppI)
-  apply auto
-  apply (rule_tac b="x # ba" in relcomppI)
-  apply auto
-  done
-
-lemma Nil_prs2 [quot_preserve]:
-  assumes "Quotient3 R Abs Rep"
-  shows "(Abs \<circ> map f) [] = Abs []"
-  by simp
-
-lemma Cons_prs2 [quot_preserve]:
-  assumes q: "Quotient3 R1 Abs1 Rep1"
-  and     r: "Quotient3 R2 Abs2 Rep2"
-  shows "(Rep1 ---> (map Rep1 \<circ> Rep2) ---> (Abs2 \<circ> map Abs1)) (op #) = (id ---> Rep2 ---> Abs2) (op #)"
-  by (auto simp add: fun_eq_iff comp_def Quotient3_abs_rep [OF q])
-
-lemma append_prs2 [quot_preserve]:
-  assumes q: "Quotient3 R1 Abs1 Rep1"
-  and     r: "Quotient3 R2 Abs2 Rep2"
-  shows "((map Rep1 \<circ> Rep2) ---> (map Rep1 \<circ> Rep2) ---> (Abs2 \<circ> map Abs1)) op @ =
-    (Rep2 ---> Rep2 ---> Abs2) op @"
-  by (simp add: fun_eq_iff abs_o_rep[OF q] List.map.id)
-
-lemma list_all2_app_l:
-  assumes a: "reflp R"
-  and b: "list_all2 R l r"
-  shows "list_all2 R (z @ l) (z @ r)"
-  using a b by (induct z) (auto elim: reflpE)
-
-lemma append_rsp2_pre0:
-  assumes a:"list_all2 op \<approx> x x'"
-  shows "list_all2 op \<approx> (x @ z) (x' @ z)"
-  using a apply (induct x x' rule: list_induct2')
-  by simp_all (rule list_all2_refl'[OF list_eq_equivp])
-
-lemma append_rsp2_pre1:
-  assumes a:"list_all2 op \<approx> x x'"
-  shows "list_all2 op \<approx> (z @ x) (z @ x')"
-  using a apply (induct x x' arbitrary: z rule: list_induct2')
-  apply (rule list_all2_refl'[OF list_eq_equivp])
-  apply (simp_all del: list_eq_def)
-  apply (rule list_all2_app_l)
-  apply (simp_all add: reflpI)
-  done
-
-lemma append_rsp2_pre:
-  assumes "list_all2 op \<approx> x x'"
-    and "list_all2 op \<approx> z z'"
-  shows "list_all2 op \<approx> (x @ z) (x' @ z')"
-  using assms by (rule list_all2_appendI)
-
-lemma compositional_rsp3:
-  assumes "(R1 ===> R2 ===> R3) C C" and "(R4 ===> R5 ===> R6) C C"
-  shows "(R1 OOO R4 ===> R2 OOO R5 ===> R3 OOO R6) C C"
-  by (auto intro!: rel_funI)
-     (metis (full_types) assms rel_funE relcomppI)
-
-lemma append_rsp2 [quot_respect]:
-  "(list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) append append"
-  by (intro compositional_rsp3)
-     (auto intro!: rel_funI simp add: append_rsp2_pre)
-
-lemma map_rsp2 [quot_respect]:
-  "((op \<approx> ===> op \<approx>) ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) map map"
-proof (auto intro!: rel_funI)
-  fix f f' :: "'a list \<Rightarrow> 'b list"
-  fix xa ya x y :: "'a list list"
-  assume fs: "(op \<approx> ===> op \<approx>) f f'" and x: "list_all2 op \<approx> xa x" and xy: "set x = set y" and y: "list_all2 op \<approx> y ya"
-  have a: "(list_all2 op \<approx>) (map f xa) (map f x)"
-    using x
-    by (induct xa x rule: list_induct2')
-       (simp_all, metis fs rel_funE list_eq_def)
-  have b: "set (map f x) = set (map f y)"
-    using xy fs
-    by (induct x y rule: list_induct2')
-       (simp_all, metis image_insert)
-  have c: "(list_all2 op \<approx>) (map f y) (map f' ya)"
-    using y fs
-    by (induct y ya rule: list_induct2')
-       (simp_all, metis apply_rsp' list_eq_def)
-  show "(list_all2 op \<approx> OOO op \<approx>) (map f xa) (map f' ya)"
-    by (metis a b c list_eq_def relcomppI)
-qed
-
-lemma map_prs2 [quot_preserve]:
-  shows "((abs_fset ---> rep_fset) ---> (map rep_fset \<circ> rep_fset) ---> abs_fset \<circ> map abs_fset) map = (id ---> rep_fset ---> abs_fset) map"
-  by (auto simp add: fun_eq_iff)
-     (simp only: map_map[symmetric] map_prs_aux[OF Quotient3_fset Quotient3_fset])
-
-section \<open>Lifted theorems\<close>
-
-subsection \<open>fset\<close>
-
-lemma fset_simps [simp]:
-  shows "fset {||} = {}"
-  and   "fset (insert_fset x S) = insert x (fset S)"
-  by (descending, simp)+
-
-lemma finite_fset [simp]: 
-  shows "finite (fset S)"
-  by (descending) (simp)
-
-lemma fset_cong:
-  shows "fset S = fset T \<longleftrightarrow> S = T"
-  by (descending) (simp)
-
-lemma filter_fset [simp]:
-  shows "fset (filter_fset P xs) = Collect P \<inter> fset xs"
-  by (descending) (auto)
-
-lemma remove_fset [simp]: 
-  shows "fset (remove_fset x xs) = fset xs - {x}"
-  by (descending) (simp)
-
-lemma inter_fset [simp]: 
-  shows "fset (xs |\<inter>| ys) = fset xs \<inter> fset ys"
-  by (descending) (auto)
-
-lemma union_fset [simp]: 
-  shows "fset (xs |\<union>| ys) = fset xs \<union> fset ys"
-  by (lifting set_append)
-
-lemma minus_fset [simp]: 
-  shows "fset (xs - ys) = fset xs - fset ys"
-  by (descending) (auto)
-
-
-subsection \<open>in_fset\<close>
-
-lemma in_fset: 
-  shows "x |\<in>| S \<longleftrightarrow> x \<in> fset S"
-  by descending simp
-
-lemma notin_fset: 
-  shows "x |\<notin>| S \<longleftrightarrow> x \<notin> fset S"
-  by (simp add: in_fset)
-
-lemma notin_empty_fset: 
-  shows "x |\<notin>| {||}"
-  by (simp add: in_fset)
-
-lemma fset_eq_iff:
-  shows "S = T \<longleftrightarrow> (\<forall>x. (x |\<in>| S) = (x |\<in>| T))"
-  by descending auto
-
-lemma none_in_empty_fset:
-  shows "(\<forall>x. x |\<notin>| S) \<longleftrightarrow> S = {||}"
-  by descending simp
-
-
-subsection \<open>insert_fset\<close>
-
-lemma in_insert_fset_iff [simp]:
-  shows "x |\<in>| insert_fset y S \<longleftrightarrow> x = y \<or> x |\<in>| S"
-  by descending simp
-
-lemma
-  shows insert_fsetI1: "x |\<in>| insert_fset x S"
-  and   insert_fsetI2: "x |\<in>| S \<Longrightarrow> x |\<in>| insert_fset y S"
-  by simp_all
-
-lemma insert_absorb_fset [simp]:
-  shows "x |\<in>| S \<Longrightarrow> insert_fset x S = S"
-  by (descending) (auto)
-
-lemma empty_not_insert_fset[simp]:
-  shows "{||} \<noteq> insert_fset x S"
-  and   "insert_fset x S \<noteq> {||}"
-  by (descending, simp)+
-
-lemma insert_fset_left_comm:
-  shows "insert_fset x (insert_fset y S) = insert_fset y (insert_fset x S)"
-  by (descending) (auto)
-
-lemma insert_fset_left_idem:
-  shows "insert_fset x (insert_fset x S) = insert_fset x S"
-  by (descending) (auto)
-
-lemma singleton_fset_eq[simp]:
-  shows "{|x|} = {|y|} \<longleftrightarrow> x = y"
-  by (descending) (auto)
-
-lemma in_fset_mdef:
-  shows "x |\<in>| F \<longleftrightarrow> x |\<notin>| (F - {|x|}) \<and> F = insert_fset x (F - {|x|})"
-  by (descending) (auto)
-
-
-subsection \<open>union_fset\<close>
-
-lemmas [simp] =
-  sup_bot_left[where 'a="'a fset"]
-  sup_bot_right[where 'a="'a fset"]
-
-lemma union_insert_fset [simp]:
-  shows "insert_fset x S |\<union>| T = insert_fset x (S |\<union>| T)"
-  by (lifting append.simps(2))
-
-lemma singleton_union_fset_left:
-  shows "{|a|} |\<union>| S = insert_fset a S"
-  by simp
-
-lemma singleton_union_fset_right:
-  shows "S |\<union>| {|a|} = insert_fset a S"
-  by (subst sup.commute) simp
-
-lemma in_union_fset:
-  shows "x |\<in>| S |\<union>| T \<longleftrightarrow> x |\<in>| S \<or> x |\<in>| T"
-  by (descending) (simp)
-
-
-subsection \<open>minus_fset\<close>
-
-lemma minus_in_fset: 
-  shows "x |\<in>| (xs - ys) \<longleftrightarrow> x |\<in>| xs \<and> x |\<notin>| ys"
-  by (descending) (simp)
-
-lemma minus_insert_fset: 
-  shows "insert_fset x xs - ys = (if x |\<in>| ys then xs - ys else insert_fset x (xs - ys))"
-  by (descending) (auto)
-
-lemma minus_insert_in_fset[simp]: 
-  shows "x |\<in>| ys \<Longrightarrow> insert_fset x xs - ys = xs - ys"
-  by (simp add: minus_insert_fset)
-
-lemma minus_insert_notin_fset[simp]: 
-  shows "x |\<notin>| ys \<Longrightarrow> insert_fset x xs - ys = insert_fset x (xs - ys)"
-  by (simp add: minus_insert_fset)
-
-lemma in_minus_fset: 
-  shows "x |\<in>| F - S \<Longrightarrow> x |\<notin>| S"
-  unfolding in_fset minus_fset
-  by blast
-
-lemma notin_minus_fset: 
-  shows "x |\<in>| S \<Longrightarrow> x |\<notin>| F - S"
-  unfolding in_fset minus_fset
-  by blast
-
-
-subsection \<open>remove_fset\<close>
-
-lemma in_remove_fset:
-  shows "x |\<in>| remove_fset y S \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y"
-  by (descending) (simp)
-
-lemma notin_remove_fset:
-  shows "x |\<notin>| remove_fset x S"
-  by (descending) (simp)
-
-lemma notin_remove_ident_fset:
-  shows "x |\<notin>| S \<Longrightarrow> remove_fset x S = S"
-  by (descending) (simp)
-
-lemma remove_fset_cases:
-  shows "S = {||} \<or> (\<exists>x. x |\<in>| S \<and> S = insert_fset x (remove_fset x S))"
-  by (descending) (auto simp add: insert_absorb)
-  
-
-subsection \<open>inter_fset\<close>
-
-lemma inter_empty_fset_l:
-  shows "{||} |\<inter>| S = {||}"
-  by simp
-
-lemma inter_empty_fset_r:
-  shows "S |\<inter>| {||} = {||}"
-  by simp
-
-lemma inter_insert_fset:
-  shows "insert_fset x S |\<inter>| T = (if x |\<in>| T then insert_fset x (S |\<inter>| T) else S |\<inter>| T)"
-  by (descending) (auto)
-
-lemma in_inter_fset:
-  shows "x |\<in>| (S |\<inter>| T) \<longleftrightarrow> x |\<in>| S \<and> x |\<in>| T"
-  by (descending) (simp)
-
-
-subsection \<open>subset_fset and psubset_fset\<close>
-
-lemma subset_fset: 
-  shows "xs |\<subseteq>| ys \<longleftrightarrow> fset xs \<subseteq> fset ys"
-  by (descending) (simp)
-
-lemma psubset_fset: 
-  shows "xs |\<subset>| ys \<longleftrightarrow> fset xs \<subset> fset ys"
-  unfolding less_fset_def 
-  by (descending) (auto)
-
-lemma subset_insert_fset:
-  shows "(insert_fset x xs) |\<subseteq>| ys \<longleftrightarrow> x |\<in>| ys \<and> xs |\<subseteq>| ys"
-  by (descending) (simp)
-
-lemma subset_in_fset: 
-  shows "xs |\<subseteq>| ys = (\<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys)"
-  by (descending) (auto)
-
-lemma subset_empty_fset:
-  shows "xs |\<subseteq>| {||} \<longleftrightarrow> xs = {||}"
-  by (descending) (simp)
-
-lemma not_psubset_empty_fset: 
-  shows "\<not> xs |\<subset>| {||}"
-  by (metis fset_simps(1) psubset_fset not_psubset_empty)
-
-
-subsection \<open>map_fset\<close>
-
-lemma map_fset_simps [simp]:
-   shows "map_fset f {||} = {||}"
-  and   "map_fset f (insert_fset x S) = insert_fset (f x) (map_fset f S)"
-  by (descending, simp)+
-
-lemma map_fset_image [simp]:
-  shows "fset (map_fset f S) = f ` (fset S)"
-  by (descending) (simp)
-
-lemma inj_map_fset_cong:
-  shows "inj f \<Longrightarrow> map_fset f S = map_fset f T \<longleftrightarrow> S = T"
-  by (descending) (metis inj_vimage_image_eq list_eq_def set_map)
-
-lemma map_union_fset: 
-  shows "map_fset f (S |\<union>| T) = map_fset f S |\<union>| map_fset f T"
-  by (descending) (simp)
-
-lemma in_fset_map_fset[simp]: "a |\<in>| map_fset f X = (\<exists>b. b |\<in>| X \<and> a = f b)"
-  by descending auto
-
-
-subsection \<open>card_fset\<close>
-
-lemma card_fset: 
-  shows "card_fset xs = card (fset xs)"
-  by (descending) (simp)
-
-lemma card_insert_fset_iff [simp]:
-  shows "card_fset (insert_fset x S) = (if x |\<in>| S then card_fset S else Suc (card_fset S))"
-  by (descending) (simp add: insert_absorb)
-
-lemma card_fset_0[simp]:
-  shows "card_fset S = 0 \<longleftrightarrow> S = {||}"
-  by (descending) (simp)
-
-lemma card_empty_fset[simp]:
-  shows "card_fset {||} = 0"
-  by (simp add: card_fset)
-
-lemma card_fset_1:
-  shows "card_fset S = 1 \<longleftrightarrow> (\<exists>x. S = {|x|})"
-  by (descending) (auto simp add: card_Suc_eq)
-
-lemma card_fset_gt_0:
-  shows "x \<in> fset S \<Longrightarrow> 0 < card_fset S"
-  by (descending) (auto simp add: card_gt_0_iff)
-  
-lemma card_notin_fset:
-  shows "(x |\<notin>| S) = (card_fset (insert_fset x S) = Suc (card_fset S))"
-  by simp
-
-lemma card_fset_Suc: 
-  shows "card_fset S = Suc n \<Longrightarrow> \<exists>x T. x |\<notin>| T \<and> S = insert_fset x T \<and> card_fset T = n"
-  apply(descending)
-  apply(auto dest!: card_eq_SucD)
-  by (metis Diff_insert_absorb set_removeAll)
-
-lemma card_remove_fset_iff [simp]:
-  shows "card_fset (remove_fset y S) = (if y |\<in>| S then card_fset S - 1 else card_fset S)"
-  by (descending) (simp)
-
-lemma card_Suc_exists_in_fset: 
-  shows "card_fset S = Suc n \<Longrightarrow> \<exists>a. a |\<in>| S"
-  by (drule card_fset_Suc) (auto)
-
-lemma in_card_fset_not_0: 
-  shows "a |\<in>| A \<Longrightarrow> card_fset A \<noteq> 0"
-  by (descending) (auto)
-
-lemma card_fset_mono: 
-  shows "xs |\<subseteq>| ys \<Longrightarrow> card_fset xs \<le> card_fset ys"
-  unfolding card_fset psubset_fset
-  by (simp add: card_mono subset_fset)
-
-lemma card_subset_fset_eq: 
-  shows "xs |\<subseteq>| ys \<Longrightarrow> card_fset ys \<le> card_fset xs \<Longrightarrow> xs = ys"
-  unfolding card_fset subset_fset
-  by (auto dest: card_seteq[OF finite_fset] simp add: fset_cong)
-
-lemma psubset_card_fset_mono: 
-  shows "xs |\<subset>| ys \<Longrightarrow> card_fset xs < card_fset ys"
-  unfolding card_fset subset_fset
-  by (metis finite_fset psubset_fset psubset_card_mono)
-
-lemma card_union_inter_fset: 
-  shows "card_fset xs + card_fset ys = card_fset (xs |\<union>| ys) + card_fset (xs |\<inter>| ys)"
-  unfolding card_fset union_fset inter_fset
-  by (rule card_Un_Int[OF finite_fset finite_fset])
-
-lemma card_union_disjoint_fset: 
-  shows "xs |\<inter>| ys = {||} \<Longrightarrow> card_fset (xs |\<union>| ys) = card_fset xs + card_fset ys"
-  unfolding card_fset union_fset 
-  apply (rule card_Un_disjoint[OF finite_fset finite_fset])
-  by (metis inter_fset fset_simps(1))
-
-lemma card_remove_fset_less1: 
-  shows "x |\<in>| xs \<Longrightarrow> card_fset (remove_fset x xs) < card_fset xs"
-  unfolding card_fset in_fset remove_fset 
-  by (rule card_Diff1_less[OF finite_fset])
-
-lemma card_remove_fset_less2: 
-  shows "x |\<in>| xs \<Longrightarrow> y |\<in>| xs \<Longrightarrow> card_fset (remove_fset y (remove_fset x xs)) < card_fset xs"
-  unfolding card_fset remove_fset in_fset
-  by (rule card_Diff2_less[OF finite_fset])
-
-lemma card_remove_fset_le1: 
-  shows "card_fset (remove_fset x xs) \<le> card_fset xs"
-  unfolding remove_fset card_fset
-  by (rule card_Diff1_le[OF finite_fset])
-
-lemma card_psubset_fset: 
-  shows "ys |\<subseteq>| xs \<Longrightarrow> card_fset ys < card_fset xs \<Longrightarrow> ys |\<subset>| xs"
-  unfolding card_fset psubset_fset subset_fset
-  by (rule card_psubset[OF finite_fset])
-
-lemma card_map_fset_le: 
-  shows "card_fset (map_fset f xs) \<le> card_fset xs"
-  unfolding card_fset map_fset_image
-  by (rule card_image_le[OF finite_fset])
-
-lemma card_minus_insert_fset[simp]:
-  assumes "a |\<in>| A" and "a |\<notin>| B"
-  shows "card_fset (A - insert_fset a B) = card_fset (A - B) - 1"
-  using assms 
-  unfolding in_fset card_fset minus_fset
-  by (simp add: card_Diff_insert[OF finite_fset])
-
-lemma card_minus_subset_fset:
-  assumes "B |\<subseteq>| A"
-  shows "card_fset (A - B) = card_fset A - card_fset B"
-  using assms 
-  unfolding subset_fset card_fset minus_fset
-  by (rule card_Diff_subset[OF finite_fset])
-
-lemma card_minus_fset:
-  shows "card_fset (A - B) = card_fset A - card_fset (A |\<inter>| B)"
-  unfolding inter_fset card_fset minus_fset
-  by (rule card_Diff_subset_Int) (simp)
-
-
-subsection \<open>concat_fset\<close>
-
-lemma concat_empty_fset [simp]:
-  shows "concat_fset {||} = {||}"
-  by descending simp
-
-lemma concat_insert_fset [simp]:
-  shows "concat_fset (insert_fset x S) = x |\<union>| concat_fset S"
-  by descending simp
-
-lemma concat_union_fset [simp]:
-  shows "concat_fset (xs |\<union>| ys) = concat_fset xs |\<union>| concat_fset ys"
-  by descending simp
-
-lemma map_concat_fset:
-  shows "map_fset f (concat_fset xs) = concat_fset (map_fset (map_fset f) xs)"
-  by (lifting map_concat)
-
-subsection \<open>filter_fset\<close>
-
-lemma subset_filter_fset: 
-  "filter_fset P xs |\<subseteq>| filter_fset Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)"
-  by descending auto
-
-lemma eq_filter_fset: 
-  "(filter_fset P xs = filter_fset Q xs) = (\<forall>x. x |\<in>| xs \<longrightarrow> P x = Q x)"
-  by descending auto
-
-lemma psubset_filter_fset:
-  "(\<And>x. x |\<in>| xs \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| xs & \<not> P x & Q x) \<Longrightarrow> 
-    filter_fset P xs |\<subset>| filter_fset Q xs"
-  unfolding less_fset_def by (auto simp add: subset_filter_fset eq_filter_fset)
-
-
-subsection \<open>fold_fset\<close>
-
-lemma fold_empty_fset: 
-  "fold_fset f {||} = id"
-  by descending (simp add: fold_once_def)
-
-lemma fold_insert_fset: "fold_fset f (insert_fset a A) =
-  (if rsp_fold f then if a |\<in>| A then fold_fset f A else fold_fset f A \<circ> f a else id)"
-  by descending (simp add: fold_once_fold_remdups)
-
-lemma remdups_removeAll:
-  "remdups (removeAll x xs) = remove1 x (remdups xs)"
-  by (induct xs) auto
-
-lemma member_commute_fold_once:
-  assumes "rsp_fold f"
-    and "x \<in> set xs"
-  shows "fold_once f xs = fold_once f (removeAll x xs) \<circ> f x"
-proof -
-  from assms have "fold f (remdups xs) = fold f (remove1 x (remdups xs)) \<circ> f x"
-    by (auto intro!: fold_remove1_split elim: rsp_foldE)
-  then show ?thesis using \<open>rsp_fold f\<close> by (simp add: fold_once_fold_remdups remdups_removeAll)
-qed
-
-lemma in_commute_fold_fset:
-  "rsp_fold f \<Longrightarrow> h |\<in>| b \<Longrightarrow> fold_fset f b = fold_fset f (remove_fset h b) \<circ> f h"
-  by descending (simp add: member_commute_fold_once)
-
-
-subsection \<open>Choice in fsets\<close>
-
-lemma fset_choice: 
-  assumes a: "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y)"
-  shows "\<exists>f. \<forall>x. x |\<in>| A \<longrightarrow> P x (f x)"
-  using a
-  apply(descending)
-  using finite_set_choice
-  by (auto simp add: Ball_def)
-
-
-section \<open>Induction and Cases rules for fsets\<close>
-
-lemma fset_exhaust [case_names empty insert, cases type: fset]:
-  assumes empty_fset_case: "S = {||} \<Longrightarrow> P" 
-  and     insert_fset_case: "\<And>x S'. S = insert_fset x S' \<Longrightarrow> P"
-  shows "P"
-  using assms by (lifting list.exhaust)
-
-lemma fset_induct [case_names empty insert]:
-  assumes empty_fset_case: "P {||}"
-  and     insert_fset_case: "\<And>x S. P S \<Longrightarrow> P (insert_fset x S)"
-  shows "P S"
-  using assms 
-  by (descending) (blast intro: list.induct)
-
-lemma fset_induct_stronger [case_names empty insert, induct type: fset]:
-  assumes empty_fset_case: "P {||}"
-  and     insert_fset_case: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (insert_fset x S)"
-  shows "P S"
-proof(induct S rule: fset_induct)
-  case empty
-  show "P {||}" using empty_fset_case by simp
-next
-  case (insert x S)
-  have "P S" by fact
-  then show "P (insert_fset x S)" using insert_fset_case 
-    by (cases "x |\<in>| S") (simp_all)
-qed
-
-lemma fset_card_induct:
-  assumes empty_fset_case: "P {||}"
-  and     card_fset_Suc_case: "\<And>S T. Suc (card_fset S) = (card_fset T) \<Longrightarrow> P S \<Longrightarrow> P T"
-  shows "P S"
-proof (induct S)
-  case empty
-  show "P {||}" by (rule empty_fset_case)
-next
-  case (insert x S)
-  have h: "P S" by fact
-  have "x |\<notin>| S" by fact
-  then have "Suc (card_fset S) = card_fset (insert_fset x S)" 
-    using card_fset_Suc by auto
-  then show "P (insert_fset x S)" 
-    using h card_fset_Suc_case by simp
-qed
-
-lemma fset_raw_strong_cases:
-  obtains "xs = []"
-    | ys x where "\<not> List.member ys x" and "xs \<approx> x # ys"
-proof (induct xs)
-  case Nil
-  then show thesis by simp
-next
-  case (Cons a xs)
-  have a: "\<lbrakk>xs = [] \<Longrightarrow> thesis; \<And>x ys. \<lbrakk>\<not> List.member ys x; xs \<approx> x # ys\<rbrakk> \<Longrightarrow> thesis\<rbrakk> \<Longrightarrow> thesis"
-    by (rule Cons(1))
-  have b: "\<And>x' ys'. \<lbrakk>\<not> List.member ys' x'; a # xs \<approx> x' # ys'\<rbrakk> \<Longrightarrow> thesis" by fact
-  have c: "xs = [] \<Longrightarrow> thesis" using b 
-    apply(simp)
-    by (metis list.set(1) emptyE empty_subsetI)
-  have "\<And>x ys. \<lbrakk>\<not> List.member ys x; xs \<approx> x # ys\<rbrakk> \<Longrightarrow> thesis"
-  proof -
-    fix x :: 'a
-    fix ys :: "'a list"
-    assume d:"\<not> List.member ys x"
-    assume e:"xs \<approx> x # ys"
-    show thesis
-    proof (cases "x = a")
-      assume h: "x = a"
-      then have f: "\<not> List.member ys a" using d by simp
-      have g: "a # xs \<approx> a # ys" using e h by auto
-      show thesis using b f g by simp
-    next
-      assume h: "x \<noteq> a"
-      then have f: "\<not> List.member (a # ys) x" using d by auto
-      have g: "a # xs \<approx> x # (a # ys)" using e h by auto
-      show thesis using b f g by (simp del: List.member_def) 
-    qed
-  qed
-  then show thesis using a c by blast
-qed
-
-
-lemma fset_strong_cases:
-  obtains "xs = {||}"
-    | ys x where "x |\<notin>| ys" and "xs = insert_fset x ys"
-  by (lifting fset_raw_strong_cases)
-
-
-lemma fset_induct2:
-  "P {||} {||} \<Longrightarrow>
-  (\<And>x xs. x |\<notin>| xs \<Longrightarrow> P (insert_fset x xs) {||}) \<Longrightarrow>
-  (\<And>y ys. y |\<notin>| ys \<Longrightarrow> P {||} (insert_fset y ys)) \<Longrightarrow>
-  (\<And>x xs y ys. \<lbrakk>P xs ys; x |\<notin>| xs; y |\<notin>| ys\<rbrakk> \<Longrightarrow> P (insert_fset x xs) (insert_fset y ys)) \<Longrightarrow>
-  P xsa ysa"
-  apply (induct xsa arbitrary: ysa)
-  apply (induct_tac x rule: fset_induct_stronger)
-  apply simp_all
-  apply (induct_tac xa rule: fset_induct_stronger)
-  apply simp_all
-  done
-
-text \<open>Extensionality\<close>
-
-lemma fset_eqI:
-  assumes "\<And>x. x \<in> fset A \<longleftrightarrow> x \<in> fset B"
-  shows "A = B"
-using assms proof (induct A arbitrary: B)
-  case empty then show ?case
-    by (auto simp add: in_fset none_in_empty_fset [symmetric] sym)
-next
-  case (insert x A)
-  from insert.prems insert.hyps(1) have "\<And>z. z \<in> fset A \<longleftrightarrow> z \<in> fset (B - {|x|})"
-    by (auto simp add: in_fset)
-  then have A: "A = B - {|x|}" by (rule insert.hyps(2))
-  with insert.prems [symmetric, of x] have "x |\<in>| B" by (simp add: in_fset)
-  with A show ?case by (metis in_fset_mdef)
-qed
-
-subsection \<open>alternate formulation with a different decomposition principle
-  and a proof of equivalence\<close>
-
-inductive
-  list_eq2 :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<approx>2 _")
-where
-  "(a # b # xs) \<approx>2 (b # a # xs)"
-| "[] \<approx>2 []"
-| "xs \<approx>2 ys \<Longrightarrow> ys \<approx>2 xs"
-| "(a # a # xs) \<approx>2 (a # xs)"
-| "xs \<approx>2 ys \<Longrightarrow> (a # xs) \<approx>2 (a # ys)"
-| "xs1 \<approx>2 xs2 \<Longrightarrow> xs2 \<approx>2 xs3 \<Longrightarrow> xs1 \<approx>2 xs3"
-
-lemma list_eq2_refl:
-  shows "xs \<approx>2 xs"
-  by (induct xs) (auto intro: list_eq2.intros)
-
-lemma cons_delete_list_eq2:
-  shows "(a # (removeAll a A)) \<approx>2 (if List.member A a then A else a # A)"
-  apply (induct A)
-  apply (simp add: list_eq2_refl)
-  apply (case_tac "List.member (aa # A) a")
-  apply (simp_all)
-  apply (case_tac [!] "a = aa")
-  apply (simp_all)
-  apply (case_tac "List.member A a")
-  apply (auto)[2]
-  apply (metis list_eq2.intros(3) list_eq2.intros(4) list_eq2.intros(5) list_eq2.intros(6))
-  apply (metis list_eq2.intros(1) list_eq2.intros(5) list_eq2.intros(6))
-  apply (auto simp add: list_eq2_refl)
-  done
-
-lemma member_delete_list_eq2:
-  assumes a: "List.member r e"
-  shows "(e # removeAll e r) \<approx>2 r"
-  using a cons_delete_list_eq2[of e r]
-  by simp
-
-lemma list_eq2_equiv:
-  "(l \<approx> r) \<longleftrightarrow> (list_eq2 l r)"
-proof
-  show "list_eq2 l r \<Longrightarrow> l \<approx> r" by (induct rule: list_eq2.induct) auto
-next
-  {
-    fix n
-    assume a: "card_list l = n" and b: "l \<approx> r"
-    have "l \<approx>2 r"
-      using a b
-    proof (induct n arbitrary: l r)
-      case 0
-      have "card_list l = 0" by fact
-      then have "\<forall>x. \<not> List.member l x" by auto
-      then have z: "l = []" by auto
-      then have "r = []" using \<open>l \<approx> r\<close> by simp
-      then show ?case using z list_eq2_refl by simp
-    next
-      case (Suc m)
-      have b: "l \<approx> r" by fact
-      have d: "card_list l = Suc m" by fact
-      then have "\<exists>a. List.member l a" 
-        apply(simp)
-        apply(drule card_eq_SucD)
-        apply(blast)
-        done
-      then obtain a where e: "List.member l a" by auto
-      then have e': "List.member r a" using list_eq_def [simplified List.member_def [symmetric], of l r] b 
-        by auto
-      have f: "card_list (removeAll a l) = m" using e d by (simp)
-      have g: "removeAll a l \<approx> removeAll a r" using remove_fset.rsp b by simp
-      have "(removeAll a l) \<approx>2 (removeAll a r)" by (rule Suc.hyps[OF f g])
-      then have h: "(a # removeAll a l) \<approx>2 (a # removeAll a r)" by (rule list_eq2.intros(5))
-      have i: "l \<approx>2 (a # removeAll a l)"
-        by (rule list_eq2.intros(3)[OF member_delete_list_eq2[OF e]])
-      have "l \<approx>2 (a # removeAll a r)" by (rule list_eq2.intros(6)[OF i h])
-      then show ?case using list_eq2.intros(6)[OF _ member_delete_list_eq2[OF e']] by simp
-    qed
-    }
-  then show "l \<approx> r \<Longrightarrow> l \<approx>2 r" by blast
-qed
-
-
-(* We cannot write it as "assumes .. shows" since Isabelle changes
-   the quantifiers to schematic variables and reintroduces them in
-   a different order *)
-lemma fset_eq_cases:
- "\<lbrakk>a1 = a2;
-   \<And>a b xs. \<lbrakk>a1 = insert_fset a (insert_fset b xs); a2 = insert_fset b (insert_fset a xs)\<rbrakk> \<Longrightarrow> P;
-   \<lbrakk>a1 = {||}; a2 = {||}\<rbrakk> \<Longrightarrow> P; \<And>xs ys. \<lbrakk>a1 = ys; a2 = xs; xs = ys\<rbrakk> \<Longrightarrow> P;
-   \<And>a xs. \<lbrakk>a1 = insert_fset a (insert_fset a xs); a2 = insert_fset a xs\<rbrakk> \<Longrightarrow> P;
-   \<And>xs ys a. \<lbrakk>a1 = insert_fset a xs; a2 = insert_fset a ys; xs = ys\<rbrakk> \<Longrightarrow> P;
-   \<And>xs1 xs2 xs3. \<lbrakk>a1 = xs1; a2 = xs3; xs1 = xs2; xs2 = xs3\<rbrakk> \<Longrightarrow> P\<rbrakk>
-  \<Longrightarrow> P"
-  by (lifting list_eq2.cases[simplified list_eq2_equiv[symmetric]])
-
-lemma fset_eq_induct:
-  assumes "x1 = x2"
-  and "\<And>a b xs. P (insert_fset a (insert_fset b xs)) (insert_fset b (insert_fset a xs))"
-  and "P {||} {||}"
-  and "\<And>xs ys. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P ys xs"
-  and "\<And>a xs. P (insert_fset a (insert_fset a xs)) (insert_fset a xs)"
-  and "\<And>xs ys a. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P (insert_fset a xs) (insert_fset a ys)"
-  and "\<And>xs1 xs2 xs3. \<lbrakk>xs1 = xs2; P xs1 xs2; xs2 = xs3; P xs2 xs3\<rbrakk> \<Longrightarrow> P xs1 xs3"
-  shows "P x1 x2"
-  using assms
-  by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]])
-
-ML \<open>
-fun dest_fsetT (Type (@{type_name fset}, [T])) = T
-  | dest_fsetT T = raise TYPE ("dest_fsetT: fset type expected", [T], []);
-\<close>
-
-no_notation
-  list_eq (infix "\<approx>" 50) and 
-  list_eq2 (infix "\<approx>2" 50)
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Quotient_Examples/Quotient_FSet.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -0,0 +1,1163 @@
+(*  Title:      HOL/Quotient_Examples/Quotient_FSet.thy
+    Author:     Cezary Kaliszyk, TU Munich
+    Author:     Christian Urban, TU Munich
+
+Type of finite sets.
+*)
+
+(********************************************************************
+  WARNING: There is a formalization of 'a fset as a subtype of sets in
+  HOL/Library/FSet.thy using Lifting/Transfer. The user should use
+  that file rather than this file unless there are some very specific
+  reasons.
+*********************************************************************)
+
+theory Quotient_FSet
+imports "~~/src/HOL/Library/Multiset" "~~/src/HOL/Library/Quotient_List"
+begin
+
+text \<open>
+  The type of finite sets is created by a quotient construction
+  over lists. The definition of the equivalence:
+\<close>
+
+definition
+  list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50)
+where
+  [simp]: "xs \<approx> ys \<longleftrightarrow> set xs = set ys"
+
+lemma list_eq_reflp:
+  "reflp list_eq"
+  by (auto intro: reflpI)
+
+lemma list_eq_symp:
+  "symp list_eq"
+  by (auto intro: sympI)
+
+lemma list_eq_transp:
+  "transp list_eq"
+  by (auto intro: transpI)
+
+lemma list_eq_equivp:
+  "equivp list_eq"
+  by (auto intro: equivpI list_eq_reflp list_eq_symp list_eq_transp)
+
+text \<open>The \<open>fset\<close> type\<close>
+
+quotient_type
+  'a fset = "'a list" / "list_eq"
+  by (rule list_eq_equivp)
+
+text \<open>
+  Definitions for sublist, cardinality, 
+  intersection, difference and respectful fold over 
+  lists.
+\<close>
+
+declare List.member_def [simp]
+
+definition
+  sub_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
+where 
+  [simp]: "sub_list xs ys \<longleftrightarrow> set xs \<subseteq> set ys"
+
+definition
+  card_list :: "'a list \<Rightarrow> nat"
+where
+  [simp]: "card_list xs = card (set xs)"
+
+definition
+  inter_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
+  [simp]: "inter_list xs ys = [x \<leftarrow> xs. x \<in> set xs \<and> x \<in> set ys]"
+
+definition
+  diff_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
+  [simp]: "diff_list xs ys = [x \<leftarrow> xs. x \<notin> set ys]"
+
+definition
+  rsp_fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
+where
+  "rsp_fold f \<longleftrightarrow> (\<forall>u v. f u \<circ> f v = f v \<circ> f u)"
+
+lemma rsp_foldI:
+  "(\<And>u v. f u \<circ> f v = f v \<circ> f u) \<Longrightarrow> rsp_fold f"
+  by (simp add: rsp_fold_def)
+
+lemma rsp_foldE:
+  assumes "rsp_fold f"
+  obtains "f u \<circ> f v = f v \<circ> f u"
+  using assms by (simp add: rsp_fold_def)
+
+definition
+  fold_once :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b"
+where
+  "fold_once f xs = (if rsp_fold f then fold f (remdups xs) else id)"
+
+lemma fold_once_default [simp]:
+  "\<not> rsp_fold f \<Longrightarrow> fold_once f xs = id"
+  by (simp add: fold_once_def)
+
+lemma fold_once_fold_remdups:
+  "rsp_fold f \<Longrightarrow> fold_once f xs = fold f (remdups xs)"
+  by (simp add: fold_once_def)
+
+
+section \<open>Quotient composition lemmas\<close>
+
+lemma list_all2_refl':
+  assumes q: "equivp R"
+  shows "(list_all2 R) r r"
+  by (rule list_all2_refl) (metis equivp_def q)
+
+lemma compose_list_refl:
+  assumes q: "equivp R"
+  shows "(list_all2 R OOO op \<approx>) r r"
+proof
+  have *: "r \<approx> r" by (rule equivp_reflp[OF fset_equivp])
+  show "list_all2 R r r" by (rule list_all2_refl'[OF q])
+  with * show "(op \<approx> OO list_all2 R) r r" ..
+qed
+
+lemma map_list_eq_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba"
+  by (simp only: list_eq_def set_map)
+
+lemma quotient_compose_list_g:
+  assumes q: "Quotient3 R Abs Rep"
+  and     e: "equivp R"
+  shows  "Quotient3 ((list_all2 R) OOO (op \<approx>))
+    (abs_fset \<circ> (map Abs)) ((map Rep) \<circ> rep_fset)"
+  unfolding Quotient3_def comp_def
+proof (intro conjI allI)
+  fix a r s
+  show "abs_fset (map Abs (map Rep (rep_fset a))) = a"
+    by (simp add: abs_o_rep[OF q] Quotient3_abs_rep[OF Quotient3_fset] List.map.id)
+  have b: "list_all2 R (map Rep (rep_fset a)) (map Rep (rep_fset a))"
+    by (rule list_all2_refl'[OF e])
+  have c: "(op \<approx> OO list_all2 R) (map Rep (rep_fset a)) (map Rep (rep_fset a))"
+    by (rule, rule equivp_reflp[OF fset_equivp]) (rule b)
+  show "(list_all2 R OOO op \<approx>) (map Rep (rep_fset a)) (map Rep (rep_fset a))"
+    by (rule, rule list_all2_refl'[OF e]) (rule c)
+  show "(list_all2 R OOO op \<approx>) r s = ((list_all2 R OOO op \<approx>) r r \<and>
+        (list_all2 R OOO op \<approx>) s s \<and> abs_fset (map Abs r) = abs_fset (map Abs s))"
+  proof (intro iffI conjI)
+    show "(list_all2 R OOO op \<approx>) r r" by (rule compose_list_refl[OF e])
+    show "(list_all2 R OOO op \<approx>) s s" by (rule compose_list_refl[OF e])
+  next
+    assume a: "(list_all2 R OOO op \<approx>) r s"
+    then have b: "map Abs r \<approx> map Abs s"
+    proof (elim relcomppE)
+      fix b ba
+      assume c: "list_all2 R r b"
+      assume d: "b \<approx> ba"
+      assume e: "list_all2 R ba s"
+      have f: "map Abs r = map Abs b"
+        using Quotient3_rel[OF list_quotient3[OF q]] c by blast
+      have "map Abs ba = map Abs s"
+        using Quotient3_rel[OF list_quotient3[OF q]] e by blast
+      then have g: "map Abs s = map Abs ba" by simp
+      then show "map Abs r \<approx> map Abs s" using d f map_list_eq_cong by simp
+    qed
+    then show "abs_fset (map Abs r) = abs_fset (map Abs s)"
+      using Quotient3_rel[OF Quotient3_fset] by blast
+  next
+    assume a: "(list_all2 R OOO op \<approx>) r r \<and> (list_all2 R OOO op \<approx>) s s
+      \<and> abs_fset (map Abs r) = abs_fset (map Abs s)"
+    then have s: "(list_all2 R OOO op \<approx>) s s" by simp
+    have d: "map Abs r \<approx> map Abs s"
+      by (subst Quotient3_rel [OF Quotient3_fset, symmetric]) (simp add: a)
+    have b: "map Rep (map Abs r) \<approx> map Rep (map Abs s)"
+      by (rule map_list_eq_cong[OF d])
+    have y: "list_all2 R (map Rep (map Abs s)) s"
+      by (fact rep_abs_rsp_left[OF list_quotient3[OF q], OF list_all2_refl'[OF e, of s]])
+    have c: "(op \<approx> OO list_all2 R) (map Rep (map Abs r)) s"
+      by (rule relcomppI) (rule b, rule y)
+    have z: "list_all2 R r (map Rep (map Abs r))"
+      by (fact rep_abs_rsp[OF list_quotient3[OF q], OF list_all2_refl'[OF e, of r]])
+    then show "(list_all2 R OOO op \<approx>) r s"
+      using a c relcomppI by simp
+  qed
+qed
+
+lemma quotient_compose_list[quot_thm]:
+  shows  "Quotient3 ((list_all2 op \<approx>) OOO (op \<approx>))
+    (abs_fset \<circ> (map abs_fset)) ((map rep_fset) \<circ> rep_fset)"
+  by (rule quotient_compose_list_g, rule Quotient3_fset, rule list_eq_equivp)
+
+
+section \<open>Quotient definitions for fsets\<close>
+
+
+subsection \<open>Finite sets are a bounded, distributive lattice with minus\<close>
+
+instantiation fset :: (type) "{bounded_lattice_bot, distrib_lattice, minus}"
+begin
+
+quotient_definition
+  "bot :: 'a fset" 
+  is "Nil :: 'a list" done
+
+abbreviation
+  empty_fset  ("{||}")
+where
+  "{||} \<equiv> bot :: 'a fset"
+
+quotient_definition
+  "less_eq_fset :: ('a fset \<Rightarrow> 'a fset \<Rightarrow> bool)"
+  is "sub_list :: ('a list \<Rightarrow> 'a list \<Rightarrow> bool)" by simp
+
+abbreviation
+  subset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50)
+where
+  "xs |\<subseteq>| ys \<equiv> xs \<le> ys"
+
+definition
+  less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool"
+where  
+  "xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)"
+
+abbreviation
+  psubset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50)
+where
+  "xs |\<subset>| ys \<equiv> xs < ys"
+
+quotient_definition
+  "sup :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+  is "append :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" by simp
+
+abbreviation
+  union_fset (infixl "|\<union>|" 65)
+where
+  "xs |\<union>| ys \<equiv> sup xs (ys::'a fset)"
+
+quotient_definition
+  "inf :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+  is "inter_list :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" by simp
+
+abbreviation
+  inter_fset (infixl "|\<inter>|" 65)
+where
+  "xs |\<inter>| ys \<equiv> inf xs (ys::'a fset)"
+
+quotient_definition
+  "minus :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+  is "diff_list :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" by fastforce
+
+instance
+proof
+  fix x y z :: "'a fset"
+  show "x |\<subset>| y \<longleftrightarrow> x |\<subseteq>| y \<and> \<not> y |\<subseteq>| x"
+    by (unfold less_fset_def, descending) auto
+  show "x |\<subseteq>| x" by (descending) (simp)
+  show "{||} |\<subseteq>| x" by (descending) (simp)
+  show "x |\<subseteq>| x |\<union>| y" by (descending) (simp)
+  show "y |\<subseteq>| x |\<union>| y" by (descending) (simp)
+  show "x |\<inter>| y |\<subseteq>| x" by (descending) (auto)
+  show "x |\<inter>| y |\<subseteq>| y" by (descending) (auto)
+  show "x |\<union>| (y |\<inter>| z) = x |\<union>| y |\<inter>| (x |\<union>| z)" 
+    by (descending) (auto)
+next
+  fix x y z :: "'a fset"
+  assume a: "x |\<subseteq>| y"
+  assume b: "y |\<subseteq>| z"
+  show "x |\<subseteq>| z" using a b by (descending) (simp)
+next
+  fix x y :: "'a fset"
+  assume a: "x |\<subseteq>| y"
+  assume b: "y |\<subseteq>| x"
+  show "x = y" using a b by (descending) (auto)
+next
+  fix x y z :: "'a fset"
+  assume a: "y |\<subseteq>| x"
+  assume b: "z |\<subseteq>| x"
+  show "y |\<union>| z |\<subseteq>| x" using a b by (descending) (simp)
+next
+  fix x y z :: "'a fset"
+  assume a: "x |\<subseteq>| y"
+  assume b: "x |\<subseteq>| z"
+  show "x |\<subseteq>| y |\<inter>| z" using a b by (descending) (auto)
+qed
+
+end
+
+
+subsection \<open>Other constants for fsets\<close>
+
+quotient_definition
+  "insert_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+  is "Cons" by auto
+
+syntax
+  "_insert_fset"     :: "args => 'a fset"  ("{|(_)|}")
+
+translations
+  "{|x, xs|}" == "CONST insert_fset x {|xs|}"
+  "{|x|}"     == "CONST insert_fset x {||}"
+
+quotient_definition
+  fset_member
+where
+  "fset_member :: 'a fset \<Rightarrow> 'a \<Rightarrow> bool" is "List.member" by fastforce
+
+abbreviation
+  in_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<in>|" 50)
+where
+  "x |\<in>| S \<equiv> fset_member S x"
+
+abbreviation
+  notin_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<notin>|" 50)
+where
+  "x |\<notin>| S \<equiv> \<not> (x |\<in>| S)"
+
+
+subsection \<open>Other constants on the Quotient Type\<close>
+
+quotient_definition
+  "card_fset :: 'a fset \<Rightarrow> nat"
+  is card_list by simp
+
+quotient_definition
+  "map_fset :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
+  is map by simp
+
+quotient_definition
+  "remove_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+  is removeAll by simp
+
+quotient_definition
+  "fset :: 'a fset \<Rightarrow> 'a set"
+  is "set" by simp
+
+lemma fold_once_set_equiv:
+  assumes "xs \<approx> ys"
+  shows "fold_once f xs = fold_once f ys"
+proof (cases "rsp_fold f")
+  case False then show ?thesis by simp
+next
+  case True
+  then have "\<And>x y. x \<in> set (remdups xs) \<Longrightarrow> y \<in> set (remdups xs) \<Longrightarrow> f x \<circ> f y = f y \<circ> f x"
+    by (rule rsp_foldE)
+  moreover from assms have "mset (remdups xs) = mset (remdups ys)"
+    by (simp add: set_eq_iff_mset_remdups_eq)
+  ultimately have "fold f (remdups xs) = fold f (remdups ys)"
+    by (rule fold_multiset_equiv)
+  with True show ?thesis by (simp add: fold_once_fold_remdups)
+qed
+
+quotient_definition
+  "fold_fset :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b \<Rightarrow> 'b"
+  is fold_once by (rule fold_once_set_equiv)
+
+lemma concat_rsp_pre:
+  assumes a: "list_all2 op \<approx> x x'"
+  and     b: "x' \<approx> y'"
+  and     c: "list_all2 op \<approx> y' y"
+  and     d: "\<exists>x\<in>set x. xa \<in> set x"
+  shows "\<exists>x\<in>set y. xa \<in> set x"
+proof -
+  obtain xb where e: "xb \<in> set x" and f: "xa \<in> set xb" using d by auto
+  have "\<exists>y. y \<in> set x' \<and> xb \<approx> y" by (rule list_all2_find_element[OF e a])
+  then obtain ya where h: "ya \<in> set x'" and i: "xb \<approx> ya" by auto
+  have "ya \<in> set y'" using b h by simp
+  then have "\<exists>yb. yb \<in> set y \<and> ya \<approx> yb" using c by (rule list_all2_find_element)
+  then show ?thesis using f i by auto
+qed
+
+quotient_definition
+  "concat_fset :: ('a fset) fset \<Rightarrow> 'a fset"
+  is concat 
+proof (elim relcomppE)
+fix a b ba bb
+  assume a: "list_all2 op \<approx> a ba"
+  with list_symp [OF list_eq_symp] have a': "list_all2 op \<approx> ba a" by (rule sympE)
+  assume b: "ba \<approx> bb"
+  with list_eq_symp have b': "bb \<approx> ba" by (rule sympE)
+  assume c: "list_all2 op \<approx> bb b"
+  with list_symp [OF list_eq_symp] have c': "list_all2 op \<approx> b bb" by (rule sympE)
+  have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" 
+  proof
+    fix x
+    show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" 
+    proof
+      assume d: "\<exists>xa\<in>set a. x \<in> set xa"
+      show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d])
+    next
+      assume e: "\<exists>xa\<in>set b. x \<in> set xa"
+      show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e])
+    qed
+  qed
+  then show "concat a \<approx> concat b" by auto
+qed
+
+quotient_definition
+  "filter_fset :: ('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+  is filter by force
+
+
+subsection \<open>Compositional respectfulness and preservation lemmas\<close>
+
+lemma Nil_rsp2 [quot_respect]: 
+  shows "(list_all2 op \<approx> OOO op \<approx>) Nil Nil"
+  by (rule compose_list_refl, rule list_eq_equivp)
+
+lemma Cons_rsp2 [quot_respect]:
+  shows "(op \<approx> ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) Cons Cons"
+  apply (auto intro!: rel_funI)
+  apply (rule_tac b="x # b" in relcomppI)
+  apply auto
+  apply (rule_tac b="x # ba" in relcomppI)
+  apply auto
+  done
+
+lemma Nil_prs2 [quot_preserve]:
+  assumes "Quotient3 R Abs Rep"
+  shows "(Abs \<circ> map f) [] = Abs []"
+  by simp
+
+lemma Cons_prs2 [quot_preserve]:
+  assumes q: "Quotient3 R1 Abs1 Rep1"
+  and     r: "Quotient3 R2 Abs2 Rep2"
+  shows "(Rep1 ---> (map Rep1 \<circ> Rep2) ---> (Abs2 \<circ> map Abs1)) (op #) = (id ---> Rep2 ---> Abs2) (op #)"
+  by (auto simp add: fun_eq_iff comp_def Quotient3_abs_rep [OF q])
+
+lemma append_prs2 [quot_preserve]:
+  assumes q: "Quotient3 R1 Abs1 Rep1"
+  and     r: "Quotient3 R2 Abs2 Rep2"
+  shows "((map Rep1 \<circ> Rep2) ---> (map Rep1 \<circ> Rep2) ---> (Abs2 \<circ> map Abs1)) op @ =
+    (Rep2 ---> Rep2 ---> Abs2) op @"
+  by (simp add: fun_eq_iff abs_o_rep[OF q] List.map.id)
+
+lemma list_all2_app_l:
+  assumes a: "reflp R"
+  and b: "list_all2 R l r"
+  shows "list_all2 R (z @ l) (z @ r)"
+  using a b by (induct z) (auto elim: reflpE)
+
+lemma append_rsp2_pre0:
+  assumes a:"list_all2 op \<approx> x x'"
+  shows "list_all2 op \<approx> (x @ z) (x' @ z)"
+  using a apply (induct x x' rule: list_induct2')
+  by simp_all (rule list_all2_refl'[OF list_eq_equivp])
+
+lemma append_rsp2_pre1:
+  assumes a:"list_all2 op \<approx> x x'"
+  shows "list_all2 op \<approx> (z @ x) (z @ x')"
+  using a apply (induct x x' arbitrary: z rule: list_induct2')
+  apply (rule list_all2_refl'[OF list_eq_equivp])
+  apply (simp_all del: list_eq_def)
+  apply (rule list_all2_app_l)
+  apply (simp_all add: reflpI)
+  done
+
+lemma append_rsp2_pre:
+  assumes "list_all2 op \<approx> x x'"
+    and "list_all2 op \<approx> z z'"
+  shows "list_all2 op \<approx> (x @ z) (x' @ z')"
+  using assms by (rule list_all2_appendI)
+
+lemma compositional_rsp3:
+  assumes "(R1 ===> R2 ===> R3) C C" and "(R4 ===> R5 ===> R6) C C"
+  shows "(R1 OOO R4 ===> R2 OOO R5 ===> R3 OOO R6) C C"
+  by (auto intro!: rel_funI)
+     (metis (full_types) assms rel_funE relcomppI)
+
+lemma append_rsp2 [quot_respect]:
+  "(list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) append append"
+  by (intro compositional_rsp3)
+     (auto intro!: rel_funI simp add: append_rsp2_pre)
+
+lemma map_rsp2 [quot_respect]:
+  "((op \<approx> ===> op \<approx>) ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) map map"
+proof (auto intro!: rel_funI)
+  fix f f' :: "'a list \<Rightarrow> 'b list"
+  fix xa ya x y :: "'a list list"
+  assume fs: "(op \<approx> ===> op \<approx>) f f'" and x: "list_all2 op \<approx> xa x" and xy: "set x = set y" and y: "list_all2 op \<approx> y ya"
+  have a: "(list_all2 op \<approx>) (map f xa) (map f x)"
+    using x
+    by (induct xa x rule: list_induct2')
+       (simp_all, metis fs rel_funE list_eq_def)
+  have b: "set (map f x) = set (map f y)"
+    using xy fs
+    by (induct x y rule: list_induct2')
+       (simp_all, metis image_insert)
+  have c: "(list_all2 op \<approx>) (map f y) (map f' ya)"
+    using y fs
+    by (induct y ya rule: list_induct2')
+       (simp_all, metis apply_rsp' list_eq_def)
+  show "(list_all2 op \<approx> OOO op \<approx>) (map f xa) (map f' ya)"
+    by (metis a b c list_eq_def relcomppI)
+qed
+
+lemma map_prs2 [quot_preserve]:
+  shows "((abs_fset ---> rep_fset) ---> (map rep_fset \<circ> rep_fset) ---> abs_fset \<circ> map abs_fset) map = (id ---> rep_fset ---> abs_fset) map"
+  by (auto simp add: fun_eq_iff)
+     (simp only: map_map[symmetric] map_prs_aux[OF Quotient3_fset Quotient3_fset])
+
+section \<open>Lifted theorems\<close>
+
+subsection \<open>fset\<close>
+
+lemma fset_simps [simp]:
+  shows "fset {||} = {}"
+  and   "fset (insert_fset x S) = insert x (fset S)"
+  by (descending, simp)+
+
+lemma finite_fset [simp]: 
+  shows "finite (fset S)"
+  by (descending) (simp)
+
+lemma fset_cong:
+  shows "fset S = fset T \<longleftrightarrow> S = T"
+  by (descending) (simp)
+
+lemma filter_fset [simp]:
+  shows "fset (filter_fset P xs) = Collect P \<inter> fset xs"
+  by (descending) (auto)
+
+lemma remove_fset [simp]: 
+  shows "fset (remove_fset x xs) = fset xs - {x}"
+  by (descending) (simp)
+
+lemma inter_fset [simp]: 
+  shows "fset (xs |\<inter>| ys) = fset xs \<inter> fset ys"
+  by (descending) (auto)
+
+lemma union_fset [simp]: 
+  shows "fset (xs |\<union>| ys) = fset xs \<union> fset ys"
+  by (lifting set_append)
+
+lemma minus_fset [simp]: 
+  shows "fset (xs - ys) = fset xs - fset ys"
+  by (descending) (auto)
+
+
+subsection \<open>in_fset\<close>
+
+lemma in_fset: 
+  shows "x |\<in>| S \<longleftrightarrow> x \<in> fset S"
+  by descending simp
+
+lemma notin_fset: 
+  shows "x |\<notin>| S \<longleftrightarrow> x \<notin> fset S"
+  by (simp add: in_fset)
+
+lemma notin_empty_fset: 
+  shows "x |\<notin>| {||}"
+  by (simp add: in_fset)
+
+lemma fset_eq_iff:
+  shows "S = T \<longleftrightarrow> (\<forall>x. (x |\<in>| S) = (x |\<in>| T))"
+  by descending auto
+
+lemma none_in_empty_fset:
+  shows "(\<forall>x. x |\<notin>| S) \<longleftrightarrow> S = {||}"
+  by descending simp
+
+
+subsection \<open>insert_fset\<close>
+
+lemma in_insert_fset_iff [simp]:
+  shows "x |\<in>| insert_fset y S \<longleftrightarrow> x = y \<or> x |\<in>| S"
+  by descending simp
+
+lemma
+  shows insert_fsetI1: "x |\<in>| insert_fset x S"
+  and   insert_fsetI2: "x |\<in>| S \<Longrightarrow> x |\<in>| insert_fset y S"
+  by simp_all
+
+lemma insert_absorb_fset [simp]:
+  shows "x |\<in>| S \<Longrightarrow> insert_fset x S = S"
+  by (descending) (auto)
+
+lemma empty_not_insert_fset[simp]:
+  shows "{||} \<noteq> insert_fset x S"
+  and   "insert_fset x S \<noteq> {||}"
+  by (descending, simp)+
+
+lemma insert_fset_left_comm:
+  shows "insert_fset x (insert_fset y S) = insert_fset y (insert_fset x S)"
+  by (descending) (auto)
+
+lemma insert_fset_left_idem:
+  shows "insert_fset x (insert_fset x S) = insert_fset x S"
+  by (descending) (auto)
+
+lemma singleton_fset_eq[simp]:
+  shows "{|x|} = {|y|} \<longleftrightarrow> x = y"
+  by (descending) (auto)
+
+lemma in_fset_mdef:
+  shows "x |\<in>| F \<longleftrightarrow> x |\<notin>| (F - {|x|}) \<and> F = insert_fset x (F - {|x|})"
+  by (descending) (auto)
+
+
+subsection \<open>union_fset\<close>
+
+lemmas [simp] =
+  sup_bot_left[where 'a="'a fset"]
+  sup_bot_right[where 'a="'a fset"]
+
+lemma union_insert_fset [simp]:
+  shows "insert_fset x S |\<union>| T = insert_fset x (S |\<union>| T)"
+  by (lifting append.simps(2))
+
+lemma singleton_union_fset_left:
+  shows "{|a|} |\<union>| S = insert_fset a S"
+  by simp
+
+lemma singleton_union_fset_right:
+  shows "S |\<union>| {|a|} = insert_fset a S"
+  by (subst sup.commute) simp
+
+lemma in_union_fset:
+  shows "x |\<in>| S |\<union>| T \<longleftrightarrow> x |\<in>| S \<or> x |\<in>| T"
+  by (descending) (simp)
+
+
+subsection \<open>minus_fset\<close>
+
+lemma minus_in_fset: 
+  shows "x |\<in>| (xs - ys) \<longleftrightarrow> x |\<in>| xs \<and> x |\<notin>| ys"
+  by (descending) (simp)
+
+lemma minus_insert_fset: 
+  shows "insert_fset x xs - ys = (if x |\<in>| ys then xs - ys else insert_fset x (xs - ys))"
+  by (descending) (auto)
+
+lemma minus_insert_in_fset[simp]: 
+  shows "x |\<in>| ys \<Longrightarrow> insert_fset x xs - ys = xs - ys"
+  by (simp add: minus_insert_fset)
+
+lemma minus_insert_notin_fset[simp]: 
+  shows "x |\<notin>| ys \<Longrightarrow> insert_fset x xs - ys = insert_fset x (xs - ys)"
+  by (simp add: minus_insert_fset)
+
+lemma in_minus_fset: 
+  shows "x |\<in>| F - S \<Longrightarrow> x |\<notin>| S"
+  unfolding in_fset minus_fset
+  by blast
+
+lemma notin_minus_fset: 
+  shows "x |\<in>| S \<Longrightarrow> x |\<notin>| F - S"
+  unfolding in_fset minus_fset
+  by blast
+
+
+subsection \<open>remove_fset\<close>
+
+lemma in_remove_fset:
+  shows "x |\<in>| remove_fset y S \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y"
+  by (descending) (simp)
+
+lemma notin_remove_fset:
+  shows "x |\<notin>| remove_fset x S"
+  by (descending) (simp)
+
+lemma notin_remove_ident_fset:
+  shows "x |\<notin>| S \<Longrightarrow> remove_fset x S = S"
+  by (descending) (simp)
+
+lemma remove_fset_cases:
+  shows "S = {||} \<or> (\<exists>x. x |\<in>| S \<and> S = insert_fset x (remove_fset x S))"
+  by (descending) (auto simp add: insert_absorb)
+  
+
+subsection \<open>inter_fset\<close>
+
+lemma inter_empty_fset_l:
+  shows "{||} |\<inter>| S = {||}"
+  by simp
+
+lemma inter_empty_fset_r:
+  shows "S |\<inter>| {||} = {||}"
+  by simp
+
+lemma inter_insert_fset:
+  shows "insert_fset x S |\<inter>| T = (if x |\<in>| T then insert_fset x (S |\<inter>| T) else S |\<inter>| T)"
+  by (descending) (auto)
+
+lemma in_inter_fset:
+  shows "x |\<in>| (S |\<inter>| T) \<longleftrightarrow> x |\<in>| S \<and> x |\<in>| T"
+  by (descending) (simp)
+
+
+subsection \<open>subset_fset and psubset_fset\<close>
+
+lemma subset_fset: 
+  shows "xs |\<subseteq>| ys \<longleftrightarrow> fset xs \<subseteq> fset ys"
+  by (descending) (simp)
+
+lemma psubset_fset: 
+  shows "xs |\<subset>| ys \<longleftrightarrow> fset xs \<subset> fset ys"
+  unfolding less_fset_def 
+  by (descending) (auto)
+
+lemma subset_insert_fset:
+  shows "(insert_fset x xs) |\<subseteq>| ys \<longleftrightarrow> x |\<in>| ys \<and> xs |\<subseteq>| ys"
+  by (descending) (simp)
+
+lemma subset_in_fset: 
+  shows "xs |\<subseteq>| ys = (\<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys)"
+  by (descending) (auto)
+
+lemma subset_empty_fset:
+  shows "xs |\<subseteq>| {||} \<longleftrightarrow> xs = {||}"
+  by (descending) (simp)
+
+lemma not_psubset_empty_fset: 
+  shows "\<not> xs |\<subset>| {||}"
+  by (metis fset_simps(1) psubset_fset not_psubset_empty)
+
+
+subsection \<open>map_fset\<close>
+
+lemma map_fset_simps [simp]:
+   shows "map_fset f {||} = {||}"
+  and   "map_fset f (insert_fset x S) = insert_fset (f x) (map_fset f S)"
+  by (descending, simp)+
+
+lemma map_fset_image [simp]:
+  shows "fset (map_fset f S) = f ` (fset S)"
+  by (descending) (simp)
+
+lemma inj_map_fset_cong:
+  shows "inj f \<Longrightarrow> map_fset f S = map_fset f T \<longleftrightarrow> S = T"
+  by (descending) (metis inj_vimage_image_eq list_eq_def set_map)
+
+lemma map_union_fset: 
+  shows "map_fset f (S |\<union>| T) = map_fset f S |\<union>| map_fset f T"
+  by (descending) (simp)
+
+lemma in_fset_map_fset[simp]: "a |\<in>| map_fset f X = (\<exists>b. b |\<in>| X \<and> a = f b)"
+  by descending auto
+
+
+subsection \<open>card_fset\<close>
+
+lemma card_fset: 
+  shows "card_fset xs = card (fset xs)"
+  by (descending) (simp)
+
+lemma card_insert_fset_iff [simp]:
+  shows "card_fset (insert_fset x S) = (if x |\<in>| S then card_fset S else Suc (card_fset S))"
+  by (descending) (simp add: insert_absorb)
+
+lemma card_fset_0[simp]:
+  shows "card_fset S = 0 \<longleftrightarrow> S = {||}"
+  by (descending) (simp)
+
+lemma card_empty_fset[simp]:
+  shows "card_fset {||} = 0"
+  by (simp add: card_fset)
+
+lemma card_fset_1:
+  shows "card_fset S = 1 \<longleftrightarrow> (\<exists>x. S = {|x|})"
+  by (descending) (auto simp add: card_Suc_eq)
+
+lemma card_fset_gt_0:
+  shows "x \<in> fset S \<Longrightarrow> 0 < card_fset S"
+  by (descending) (auto simp add: card_gt_0_iff)
+  
+lemma card_notin_fset:
+  shows "(x |\<notin>| S) = (card_fset (insert_fset x S) = Suc (card_fset S))"
+  by simp
+
+lemma card_fset_Suc: 
+  shows "card_fset S = Suc n \<Longrightarrow> \<exists>x T. x |\<notin>| T \<and> S = insert_fset x T \<and> card_fset T = n"
+  apply(descending)
+  apply(auto dest!: card_eq_SucD)
+  by (metis Diff_insert_absorb set_removeAll)
+
+lemma card_remove_fset_iff [simp]:
+  shows "card_fset (remove_fset y S) = (if y |\<in>| S then card_fset S - 1 else card_fset S)"
+  by (descending) (simp)
+
+lemma card_Suc_exists_in_fset: 
+  shows "card_fset S = Suc n \<Longrightarrow> \<exists>a. a |\<in>| S"
+  by (drule card_fset_Suc) (auto)
+
+lemma in_card_fset_not_0: 
+  shows "a |\<in>| A \<Longrightarrow> card_fset A \<noteq> 0"
+  by (descending) (auto)
+
+lemma card_fset_mono: 
+  shows "xs |\<subseteq>| ys \<Longrightarrow> card_fset xs \<le> card_fset ys"
+  unfolding card_fset psubset_fset
+  by (simp add: card_mono subset_fset)
+
+lemma card_subset_fset_eq: 
+  shows "xs |\<subseteq>| ys \<Longrightarrow> card_fset ys \<le> card_fset xs \<Longrightarrow> xs = ys"
+  unfolding card_fset subset_fset
+  by (auto dest: card_seteq[OF finite_fset] simp add: fset_cong)
+
+lemma psubset_card_fset_mono: 
+  shows "xs |\<subset>| ys \<Longrightarrow> card_fset xs < card_fset ys"
+  unfolding card_fset subset_fset
+  by (metis finite_fset psubset_fset psubset_card_mono)
+
+lemma card_union_inter_fset: 
+  shows "card_fset xs + card_fset ys = card_fset (xs |\<union>| ys) + card_fset (xs |\<inter>| ys)"
+  unfolding card_fset union_fset inter_fset
+  by (rule card_Un_Int[OF finite_fset finite_fset])
+
+lemma card_union_disjoint_fset: 
+  shows "xs |\<inter>| ys = {||} \<Longrightarrow> card_fset (xs |\<union>| ys) = card_fset xs + card_fset ys"
+  unfolding card_fset union_fset 
+  apply (rule card_Un_disjoint[OF finite_fset finite_fset])
+  by (metis inter_fset fset_simps(1))
+
+lemma card_remove_fset_less1: 
+  shows "x |\<in>| xs \<Longrightarrow> card_fset (remove_fset x xs) < card_fset xs"
+  unfolding card_fset in_fset remove_fset 
+  by (rule card_Diff1_less[OF finite_fset])
+
+lemma card_remove_fset_less2: 
+  shows "x |\<in>| xs \<Longrightarrow> y |\<in>| xs \<Longrightarrow> card_fset (remove_fset y (remove_fset x xs)) < card_fset xs"
+  unfolding card_fset remove_fset in_fset
+  by (rule card_Diff2_less[OF finite_fset])
+
+lemma card_remove_fset_le1: 
+  shows "card_fset (remove_fset x xs) \<le> card_fset xs"
+  unfolding remove_fset card_fset
+  by (rule card_Diff1_le[OF finite_fset])
+
+lemma card_psubset_fset: 
+  shows "ys |\<subseteq>| xs \<Longrightarrow> card_fset ys < card_fset xs \<Longrightarrow> ys |\<subset>| xs"
+  unfolding card_fset psubset_fset subset_fset
+  by (rule card_psubset[OF finite_fset])
+
+lemma card_map_fset_le: 
+  shows "card_fset (map_fset f xs) \<le> card_fset xs"
+  unfolding card_fset map_fset_image
+  by (rule card_image_le[OF finite_fset])
+
+lemma card_minus_insert_fset[simp]:
+  assumes "a |\<in>| A" and "a |\<notin>| B"
+  shows "card_fset (A - insert_fset a B) = card_fset (A - B) - 1"
+  using assms 
+  unfolding in_fset card_fset minus_fset
+  by (simp add: card_Diff_insert[OF finite_fset])
+
+lemma card_minus_subset_fset:
+  assumes "B |\<subseteq>| A"
+  shows "card_fset (A - B) = card_fset A - card_fset B"
+  using assms 
+  unfolding subset_fset card_fset minus_fset
+  by (rule card_Diff_subset[OF finite_fset])
+
+lemma card_minus_fset:
+  shows "card_fset (A - B) = card_fset A - card_fset (A |\<inter>| B)"
+  unfolding inter_fset card_fset minus_fset
+  by (rule card_Diff_subset_Int) (simp)
+
+
+subsection \<open>concat_fset\<close>
+
+lemma concat_empty_fset [simp]:
+  shows "concat_fset {||} = {||}"
+  by descending simp
+
+lemma concat_insert_fset [simp]:
+  shows "concat_fset (insert_fset x S) = x |\<union>| concat_fset S"
+  by descending simp
+
+lemma concat_union_fset [simp]:
+  shows "concat_fset (xs |\<union>| ys) = concat_fset xs |\<union>| concat_fset ys"
+  by descending simp
+
+lemma map_concat_fset:
+  shows "map_fset f (concat_fset xs) = concat_fset (map_fset (map_fset f) xs)"
+  by (lifting map_concat)
+
+subsection \<open>filter_fset\<close>
+
+lemma subset_filter_fset: 
+  "filter_fset P xs |\<subseteq>| filter_fset Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)"
+  by descending auto
+
+lemma eq_filter_fset: 
+  "(filter_fset P xs = filter_fset Q xs) = (\<forall>x. x |\<in>| xs \<longrightarrow> P x = Q x)"
+  by descending auto
+
+lemma psubset_filter_fset:
+  "(\<And>x. x |\<in>| xs \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| xs & \<not> P x & Q x) \<Longrightarrow> 
+    filter_fset P xs |\<subset>| filter_fset Q xs"
+  unfolding less_fset_def by (auto simp add: subset_filter_fset eq_filter_fset)
+
+
+subsection \<open>fold_fset\<close>
+
+lemma fold_empty_fset: 
+  "fold_fset f {||} = id"
+  by descending (simp add: fold_once_def)
+
+lemma fold_insert_fset: "fold_fset f (insert_fset a A) =
+  (if rsp_fold f then if a |\<in>| A then fold_fset f A else fold_fset f A \<circ> f a else id)"
+  by descending (simp add: fold_once_fold_remdups)
+
+lemma remdups_removeAll:
+  "remdups (removeAll x xs) = remove1 x (remdups xs)"
+  by (induct xs) auto
+
+lemma member_commute_fold_once:
+  assumes "rsp_fold f"
+    and "x \<in> set xs"
+  shows "fold_once f xs = fold_once f (removeAll x xs) \<circ> f x"
+proof -
+  from assms have "fold f (remdups xs) = fold f (remove1 x (remdups xs)) \<circ> f x"
+    by (auto intro!: fold_remove1_split elim: rsp_foldE)
+  then show ?thesis using \<open>rsp_fold f\<close> by (simp add: fold_once_fold_remdups remdups_removeAll)
+qed
+
+lemma in_commute_fold_fset:
+  "rsp_fold f \<Longrightarrow> h |\<in>| b \<Longrightarrow> fold_fset f b = fold_fset f (remove_fset h b) \<circ> f h"
+  by descending (simp add: member_commute_fold_once)
+
+
+subsection \<open>Choice in fsets\<close>
+
+lemma fset_choice: 
+  assumes a: "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y)"
+  shows "\<exists>f. \<forall>x. x |\<in>| A \<longrightarrow> P x (f x)"
+  using a
+  apply(descending)
+  using finite_set_choice
+  by (auto simp add: Ball_def)
+
+
+section \<open>Induction and Cases rules for fsets\<close>
+
+lemma fset_exhaust [case_names empty insert, cases type: fset]:
+  assumes empty_fset_case: "S = {||} \<Longrightarrow> P" 
+  and     insert_fset_case: "\<And>x S'. S = insert_fset x S' \<Longrightarrow> P"
+  shows "P"
+  using assms by (lifting list.exhaust)
+
+lemma fset_induct [case_names empty insert]:
+  assumes empty_fset_case: "P {||}"
+  and     insert_fset_case: "\<And>x S. P S \<Longrightarrow> P (insert_fset x S)"
+  shows "P S"
+  using assms 
+  by (descending) (blast intro: list.induct)
+
+lemma fset_induct_stronger [case_names empty insert, induct type: fset]:
+  assumes empty_fset_case: "P {||}"
+  and     insert_fset_case: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (insert_fset x S)"
+  shows "P S"
+proof(induct S rule: fset_induct)
+  case empty
+  show "P {||}" using empty_fset_case by simp
+next
+  case (insert x S)
+  have "P S" by fact
+  then show "P (insert_fset x S)" using insert_fset_case 
+    by (cases "x |\<in>| S") (simp_all)
+qed
+
+lemma fset_card_induct:
+  assumes empty_fset_case: "P {||}"
+  and     card_fset_Suc_case: "\<And>S T. Suc (card_fset S) = (card_fset T) \<Longrightarrow> P S \<Longrightarrow> P T"
+  shows "P S"
+proof (induct S)
+  case empty
+  show "P {||}" by (rule empty_fset_case)
+next
+  case (insert x S)
+  have h: "P S" by fact
+  have "x |\<notin>| S" by fact
+  then have "Suc (card_fset S) = card_fset (insert_fset x S)" 
+    using card_fset_Suc by auto
+  then show "P (insert_fset x S)" 
+    using h card_fset_Suc_case by simp
+qed
+
+lemma fset_raw_strong_cases:
+  obtains "xs = []"
+    | ys x where "\<not> List.member ys x" and "xs \<approx> x # ys"
+proof (induct xs)
+  case Nil
+  then show thesis by simp
+next
+  case (Cons a xs)
+  have a: "\<lbrakk>xs = [] \<Longrightarrow> thesis; \<And>x ys. \<lbrakk>\<not> List.member ys x; xs \<approx> x # ys\<rbrakk> \<Longrightarrow> thesis\<rbrakk> \<Longrightarrow> thesis"
+    by (rule Cons(1))
+  have b: "\<And>x' ys'. \<lbrakk>\<not> List.member ys' x'; a # xs \<approx> x' # ys'\<rbrakk> \<Longrightarrow> thesis" by fact
+  have c: "xs = [] \<Longrightarrow> thesis" using b 
+    apply(simp)
+    by (metis list.set(1) emptyE empty_subsetI)
+  have "\<And>x ys. \<lbrakk>\<not> List.member ys x; xs \<approx> x # ys\<rbrakk> \<Longrightarrow> thesis"
+  proof -
+    fix x :: 'a
+    fix ys :: "'a list"
+    assume d:"\<not> List.member ys x"
+    assume e:"xs \<approx> x # ys"
+    show thesis
+    proof (cases "x = a")
+      assume h: "x = a"
+      then have f: "\<not> List.member ys a" using d by simp
+      have g: "a # xs \<approx> a # ys" using e h by auto
+      show thesis using b f g by simp
+    next
+      assume h: "x \<noteq> a"
+      then have f: "\<not> List.member (a # ys) x" using d by auto
+      have g: "a # xs \<approx> x # (a # ys)" using e h by auto
+      show thesis using b f g by (simp del: List.member_def) 
+    qed
+  qed
+  then show thesis using a c by blast
+qed
+
+
+lemma fset_strong_cases:
+  obtains "xs = {||}"
+    | ys x where "x |\<notin>| ys" and "xs = insert_fset x ys"
+  by (lifting fset_raw_strong_cases)
+
+
+lemma fset_induct2:
+  "P {||} {||} \<Longrightarrow>
+  (\<And>x xs. x |\<notin>| xs \<Longrightarrow> P (insert_fset x xs) {||}) \<Longrightarrow>
+  (\<And>y ys. y |\<notin>| ys \<Longrightarrow> P {||} (insert_fset y ys)) \<Longrightarrow>
+  (\<And>x xs y ys. \<lbrakk>P xs ys; x |\<notin>| xs; y |\<notin>| ys\<rbrakk> \<Longrightarrow> P (insert_fset x xs) (insert_fset y ys)) \<Longrightarrow>
+  P xsa ysa"
+  apply (induct xsa arbitrary: ysa)
+  apply (induct_tac x rule: fset_induct_stronger)
+  apply simp_all
+  apply (induct_tac xa rule: fset_induct_stronger)
+  apply simp_all
+  done
+
+text \<open>Extensionality\<close>
+
+lemma fset_eqI:
+  assumes "\<And>x. x \<in> fset A \<longleftrightarrow> x \<in> fset B"
+  shows "A = B"
+using assms proof (induct A arbitrary: B)
+  case empty then show ?case
+    by (auto simp add: in_fset none_in_empty_fset [symmetric] sym)
+next
+  case (insert x A)
+  from insert.prems insert.hyps(1) have "\<And>z. z \<in> fset A \<longleftrightarrow> z \<in> fset (B - {|x|})"
+    by (auto simp add: in_fset)
+  then have A: "A = B - {|x|}" by (rule insert.hyps(2))
+  with insert.prems [symmetric, of x] have "x |\<in>| B" by (simp add: in_fset)
+  with A show ?case by (metis in_fset_mdef)
+qed
+
+subsection \<open>alternate formulation with a different decomposition principle
+  and a proof of equivalence\<close>
+
+inductive
+  list_eq2 :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<approx>2 _")
+where
+  "(a # b # xs) \<approx>2 (b # a # xs)"
+| "[] \<approx>2 []"
+| "xs \<approx>2 ys \<Longrightarrow> ys \<approx>2 xs"
+| "(a # a # xs) \<approx>2 (a # xs)"
+| "xs \<approx>2 ys \<Longrightarrow> (a # xs) \<approx>2 (a # ys)"
+| "xs1 \<approx>2 xs2 \<Longrightarrow> xs2 \<approx>2 xs3 \<Longrightarrow> xs1 \<approx>2 xs3"
+
+lemma list_eq2_refl:
+  shows "xs \<approx>2 xs"
+  by (induct xs) (auto intro: list_eq2.intros)
+
+lemma cons_delete_list_eq2:
+  shows "(a # (removeAll a A)) \<approx>2 (if List.member A a then A else a # A)"
+  apply (induct A)
+  apply (simp add: list_eq2_refl)
+  apply (case_tac "List.member (aa # A) a")
+  apply (simp_all)
+  apply (case_tac [!] "a = aa")
+  apply (simp_all)
+  apply (case_tac "List.member A a")
+  apply (auto)[2]
+  apply (metis list_eq2.intros(3) list_eq2.intros(4) list_eq2.intros(5) list_eq2.intros(6))
+  apply (metis list_eq2.intros(1) list_eq2.intros(5) list_eq2.intros(6))
+  apply (auto simp add: list_eq2_refl)
+  done
+
+lemma member_delete_list_eq2:
+  assumes a: "List.member r e"
+  shows "(e # removeAll e r) \<approx>2 r"
+  using a cons_delete_list_eq2[of e r]
+  by simp
+
+lemma list_eq2_equiv:
+  "(l \<approx> r) \<longleftrightarrow> (list_eq2 l r)"
+proof
+  show "list_eq2 l r \<Longrightarrow> l \<approx> r" by (induct rule: list_eq2.induct) auto
+next
+  {
+    fix n
+    assume a: "card_list l = n" and b: "l \<approx> r"
+    have "l \<approx>2 r"
+      using a b
+    proof (induct n arbitrary: l r)
+      case 0
+      have "card_list l = 0" by fact
+      then have "\<forall>x. \<not> List.member l x" by auto
+      then have z: "l = []" by auto
+      then have "r = []" using \<open>l \<approx> r\<close> by simp
+      then show ?case using z list_eq2_refl by simp
+    next
+      case (Suc m)
+      have b: "l \<approx> r" by fact
+      have d: "card_list l = Suc m" by fact
+      then have "\<exists>a. List.member l a" 
+        apply(simp)
+        apply(drule card_eq_SucD)
+        apply(blast)
+        done
+      then obtain a where e: "List.member l a" by auto
+      then have e': "List.member r a" using list_eq_def [simplified List.member_def [symmetric], of l r] b 
+        by auto
+      have f: "card_list (removeAll a l) = m" using e d by (simp)
+      have g: "removeAll a l \<approx> removeAll a r" using remove_fset.rsp b by simp
+      have "(removeAll a l) \<approx>2 (removeAll a r)" by (rule Suc.hyps[OF f g])
+      then have h: "(a # removeAll a l) \<approx>2 (a # removeAll a r)" by (rule list_eq2.intros(5))
+      have i: "l \<approx>2 (a # removeAll a l)"
+        by (rule list_eq2.intros(3)[OF member_delete_list_eq2[OF e]])
+      have "l \<approx>2 (a # removeAll a r)" by (rule list_eq2.intros(6)[OF i h])
+      then show ?case using list_eq2.intros(6)[OF _ member_delete_list_eq2[OF e']] by simp
+    qed
+    }
+  then show "l \<approx> r \<Longrightarrow> l \<approx>2 r" by blast
+qed
+
+
+(* We cannot write it as "assumes .. shows" since Isabelle changes
+   the quantifiers to schematic variables and reintroduces them in
+   a different order *)
+lemma fset_eq_cases:
+ "\<lbrakk>a1 = a2;
+   \<And>a b xs. \<lbrakk>a1 = insert_fset a (insert_fset b xs); a2 = insert_fset b (insert_fset a xs)\<rbrakk> \<Longrightarrow> P;
+   \<lbrakk>a1 = {||}; a2 = {||}\<rbrakk> \<Longrightarrow> P; \<And>xs ys. \<lbrakk>a1 = ys; a2 = xs; xs = ys\<rbrakk> \<Longrightarrow> P;
+   \<And>a xs. \<lbrakk>a1 = insert_fset a (insert_fset a xs); a2 = insert_fset a xs\<rbrakk> \<Longrightarrow> P;
+   \<And>xs ys a. \<lbrakk>a1 = insert_fset a xs; a2 = insert_fset a ys; xs = ys\<rbrakk> \<Longrightarrow> P;
+   \<And>xs1 xs2 xs3. \<lbrakk>a1 = xs1; a2 = xs3; xs1 = xs2; xs2 = xs3\<rbrakk> \<Longrightarrow> P\<rbrakk>
+  \<Longrightarrow> P"
+  by (lifting list_eq2.cases[simplified list_eq2_equiv[symmetric]])
+
+lemma fset_eq_induct:
+  assumes "x1 = x2"
+  and "\<And>a b xs. P (insert_fset a (insert_fset b xs)) (insert_fset b (insert_fset a xs))"
+  and "P {||} {||}"
+  and "\<And>xs ys. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P ys xs"
+  and "\<And>a xs. P (insert_fset a (insert_fset a xs)) (insert_fset a xs)"
+  and "\<And>xs ys a. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P (insert_fset a xs) (insert_fset a ys)"
+  and "\<And>xs1 xs2 xs3. \<lbrakk>xs1 = xs2; P xs1 xs2; xs2 = xs3; P xs2 xs3\<rbrakk> \<Longrightarrow> P xs1 xs3"
+  shows "P x1 x2"
+  using assms
+  by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]])
+
+ML \<open>
+fun dest_fsetT (Type (@{type_name fset}, [T])) = T
+  | dest_fsetT T = raise TYPE ("dest_fsetT: fset type expected", [T], []);
+\<close>
+
+no_notation
+  list_eq (infix "\<approx>" 50) and 
+  list_eq2 (infix "\<approx>2" 50)
+
+end
--- a/src/HOL/ROOT	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/ROOT	Tue Sep 20 11:35:10 2016 +0200
@@ -996,7 +996,7 @@
   options [document = false]
   theories
     DList
-    FSet
+    Quotient_FSet
     Quotient_Int
     Quotient_Message
     Lift_FSet
--- a/src/HOL/Set_Interval.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Set_Interval.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -1889,7 +1889,7 @@
   also have "... = (x - y) * (y * (\<Sum>p<Suc n. (x ^ p) * y ^ (n - p))) + (x - y) * (x * x^n)"
     by (simp only: mult.left_commute)
   also have "... = (x - y) * (\<Sum>p<Suc (Suc n). x ^ p * y ^ (Suc n - p))"
-    by (simp add: field_simps Suc_diff_le setsum_left_distrib setsum_right_distrib)
+    by (simp add: field_simps Suc_diff_le setsum_distrib_right setsum_distrib_left)
   finally show ?case .
 qed simp
 
@@ -1944,7 +1944,7 @@
   also from ngt1 have "\<dots> = ?n*a + (\<Sum>i\<in>{..<n}. ?I i*d)" by simp
   also from ngt1 have "\<dots> = (?n*a + d*(\<Sum>i\<in>{1..<n}. ?I i))"
     unfolding One_nat_def
-    by (simp add: setsum_right_distrib atLeast0LessThan[symmetric] setsum_shift_lb_Suc0_0_upt ac_simps)
+    by (simp add: setsum_distrib_left atLeast0LessThan[symmetric] setsum_shift_lb_Suc0_0_upt ac_simps)
   also have "2*\<dots> = 2*?n*a + d*2*(\<Sum>i\<in>{1..<n}. ?I i)"
     by (simp add: algebra_simps)
   also from ngt1 have "{1..<n} = {1..n - 1}"
--- a/src/HOL/Transcendental.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/Transcendental.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -639,11 +639,11 @@
   apply (subst lemma_realpow_rev_sumr)
   apply (subst sumr_diff_mult_const2)
   apply simp
-  apply (simp only: lemma_termdiff1 setsum_right_distrib)
+  apply (simp only: lemma_termdiff1 setsum_distrib_left)
   apply (rule setsum.cong [OF refl])
   apply (simp add: less_iff_Suc_add)
   apply clarify
-  apply (simp add: setsum_right_distrib diff_power_eq_setsum ac_simps
+  apply (simp add: setsum_distrib_left diff_power_eq_setsum ac_simps
       del: setsum_lessThan_Suc power_Suc)
   apply (subst mult.assoc [symmetric], subst power_add [symmetric])
   apply (simp add: ac_simps)
@@ -1448,7 +1448,7 @@
   also have "\<dots> = x * (\<Sum>i\<le>n. S x i * S y (n - i)) + y * (\<Sum>i\<le>n. S x i * S y (n - i))"
     by (rule distrib_right)
   also have "\<dots> = (\<Sum>i\<le>n. x * S x i * S y (n - i)) + (\<Sum>i\<le>n. S x i * y * S y (n - i))"
-    by (simp add: setsum_right_distrib ac_simps S_comm)
+    by (simp add: setsum_distrib_left ac_simps S_comm)
   also have "\<dots> = (\<Sum>i\<le>n. x * S x i * S y (n - i)) + (\<Sum>i\<le>n. S x i * (y * S y (n - i)))"
     by (simp add: ac_simps)
   also have "\<dots> = (\<Sum>i\<le>n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n - i))) +
@@ -3340,7 +3340,7 @@
        (if even p
         then of_real ((-1) ^ (p div 2) / (fact p)) * (\<Sum>n\<le>p. (p choose n) *\<^sub>R (x^n) * y^(p-n))
         else 0)"
-      by (auto simp: setsum_right_distrib field_simps scaleR_conv_of_real nonzero_of_real_divide)
+      by (auto simp: setsum_distrib_left field_simps scaleR_conv_of_real nonzero_of_real_divide)
     also have "\<dots> = cos_coeff p *\<^sub>R ((x + y) ^ p)"
       by (simp add: cos_coeff_def binomial_ring [of x y]  scaleR_conv_of_real atLeast0AtMost)
     finally show ?thesis .
@@ -5835,7 +5835,7 @@
     by (auto simp: pairs_le_eq_Sigma setsum.Sigma)
   also have "\<dots> = (\<Sum>r\<le>m + n. (\<Sum>k\<le>r. (a k) * (b (r - k))) * x ^ r)"
     apply (subst setsum_triangle_reindex_eq)
-    apply (auto simp: algebra_simps setsum_right_distrib intro!: setsum.cong)
+    apply (auto simp: algebra_simps setsum_distrib_left intro!: setsum.cong)
     apply (metis le_add_diff_inverse power_add)
     done
   finally show ?thesis .
@@ -5864,7 +5864,7 @@
   also have "\<dots> = (\<Sum>i\<le>n. a i * (x - y) * (\<Sum>j<i. y^(i - Suc j) * x^j))"
     by (simp add: power_diff_sumr2 mult.assoc)
   also have "\<dots> = (\<Sum>i\<le>n. \<Sum>j<i. a i * (x - y) * (y^(i - Suc j) * x^j))"
-    by (simp add: setsum_right_distrib)
+    by (simp add: setsum_distrib_left)
   also have "\<dots> = (\<Sum>(i,j) \<in> (SIGMA i : atMost n. lessThan i). a i * (x - y) * (y^(i - Suc j) * x^j))"
     by (simp add: setsum.Sigma)
   also have "\<dots> = (\<Sum>(j,i) \<in> (SIGMA j : lessThan n. {Suc j..n}). a i * (x - y) * (y^(i - Suc j) * x^j))"
@@ -5872,7 +5872,7 @@
   also have "\<dots> = (\<Sum>j<n. \<Sum>i=Suc j..n. a i * (x - y) * (y^(i - Suc j) * x^j))"
     by (simp add: setsum.Sigma)
   also have "\<dots> = (x - y) * (\<Sum>j<n. (\<Sum>i=Suc j..n. a i * y^(i - j - 1)) * x^j)"
-    by (simp add: setsum_right_distrib mult_ac)
+    by (simp add: setsum_distrib_left mult_ac)
   finally show ?thesis .
 qed
 
@@ -5894,7 +5894,7 @@
       by (auto simp add: setsum.reindex_bij_betw [OF h, symmetric] intro: setsum.strong_cong)
   qed
   then show ?thesis
-    by (simp add: polyfun_diff [OF assms] setsum_left_distrib)
+    by (simp add: polyfun_diff [OF assms] setsum_distrib_right)
 qed
 
 lemma polyfun_linear_factor:  (*COMPLEX_POLYFUN_LINEAR_FACTOR in HOL Light*)
@@ -5949,7 +5949,7 @@
       unfolding Set_Interval.setsum_atMost_Suc_shift
       by simp
     also have "\<dots> = w * (\<Sum>i\<le>n. c (Suc i) * w^i)"
-      by (simp add: setsum_right_distrib ac_simps)
+      by (simp add: setsum_distrib_left ac_simps)
     finally show ?thesis .
   qed
   then have w: "\<And>w. w \<noteq> 0 \<Longrightarrow> (\<Sum>i\<le>n. c (Suc i) * w^i) = 0"
--- a/src/HOL/ex/Sum_of_Powers.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/ex/Sum_of_Powers.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -110,7 +110,7 @@
     unfolding bernpoly_def by (rule DERIV_cong) (fast intro!: derivative_intros, simp)
   moreover have "(\<Sum>k\<le>n. real (Suc n - k) * x ^ (n - k) * (real (Suc n choose k) * bernoulli k)) = (n + 1) * bernpoly n x"
     unfolding bernpoly_def
-    by (auto intro: setsum.cong simp add: setsum_right_distrib real_binomial_eq_mult_binomial_Suc[of _ n] Suc_eq_plus1 of_nat_diff)
+    by (auto intro: setsum.cong simp add: setsum_distrib_left real_binomial_eq_mult_binomial_Suc[of _ n] Suc_eq_plus1 of_nat_diff)
   ultimately show ?thesis by auto
 qed
 
@@ -135,7 +135,7 @@
 lemma sum_of_powers: "(\<Sum>k\<le>n::nat. (real k) ^ m) = (bernpoly (Suc m) (n + 1) - bernpoly (Suc m) 0) / (m + 1)"
 proof -
   from diff_bernpoly[of "Suc m", simplified] have "(m + (1::real)) * (\<Sum>k\<le>n. (real k) ^ m) = (\<Sum>k\<le>n. bernpoly (Suc m) (real k + 1) - bernpoly (Suc m) (real k))"
-    by (auto simp add: setsum_right_distrib intro!: setsum.cong)
+    by (auto simp add: setsum_distrib_left intro!: setsum.cong)
   also have "... = (\<Sum>k\<le>n. bernpoly (Suc m) (real (k + 1)) - bernpoly (Suc m) (real k))"
     by simp
   also have "... = bernpoly (Suc m) (n + 1) - bernpoly (Suc m) 0"
--- a/src/HOL/ex/ThreeDivides.thy	Mon Sep 19 17:37:22 2016 +0200
+++ b/src/HOL/ex/ThreeDivides.thy	Tue Sep 20 11:35:10 2016 +0200
@@ -193,7 +193,7 @@
       "m = 10*(\<Sum>x<nd. m div 10 div 10^x mod 10 * 10^x) + c" by simp
     also have
       "\<dots> = (\<Sum>x<nd. m div 10 div 10^x mod 10 * 10^(x+1)) + c"
-      by (subst setsum_right_distrib) (simp add: ac_simps)
+      by (subst setsum_distrib_left) (simp add: ac_simps)
     also have
       "\<dots> = (\<Sum>x<nd. m div 10^(Suc x) mod 10 * 10^(Suc x)) + c"
       by (simp add: div_mult2_eq[symmetric])